--- a/CONTRIBUTORS Mon Aug 23 16:47:57 2010 +0200
+++ b/CONTRIBUTORS Tue Aug 24 08:22:17 2010 +0200
@@ -6,6 +6,9 @@
Contributions to this Isabelle version
--------------------------------------
+* August 2010: Johannes Hoelzl, Armin Heller, and Robert Himmelmann, TUM
+ Rewriting the Probability theory.
+
* July 2010: Florian Haftmann, TUM
Reworking and extension of the Isabelle/HOL framework.
--- a/NEWS Mon Aug 23 16:47:57 2010 +0200
+++ b/NEWS Tue Aug 24 08:22:17 2010 +0200
@@ -155,6 +155,13 @@
INCOMPATIBILITY.
+* Probability: Introduced pinfreal as real numbers with infinity. Use pinfreal
+as value for measures. Introduces Lebesgue Measure based on the integral in
+Multivariate Analysis. Proved Radon Nikodym for arbitrary sigma finite measure
+spaces.
+
+ INCOMPATIBILITY.
+
* Inductive package: offers new command "inductive_simps" to automatically
derive instantiated and simplified equations for inductive predicates,
similar to inductive_cases.
--- a/src/HOL/IsaMakefile Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/IsaMakefile Tue Aug 24 08:22:17 2010 +0200
@@ -1104,6 +1104,7 @@
Multivariate_Analysis/Finite_Cartesian_Product.thy \
Multivariate_Analysis/Integration.certs \
Multivariate_Analysis/Integration.thy \
+ Multivariate_Analysis/Gauge_Measure.thy \
Multivariate_Analysis/L2_Norm.thy \
Multivariate_Analysis/Multivariate_Analysis.thy \
Multivariate_Analysis/Operator_Norm.thy \
@@ -1121,20 +1122,19 @@
## HOL-Probability
-HOL-Probability: HOL $(OUT)/HOL-Probability
+HOL-Probability: HOL-Multivariate_Analysis $(OUT)/HOL-Probability
-$(OUT)/HOL-Probability: $(OUT)/HOL Probability/ROOT.ML \
+$(OUT)/HOL-Probability: $(OUT)/HOL-Multivariate_Analysis Probability/ROOT.ML \
Probability/Probability.thy Probability/Sigma_Algebra.thy \
- Probability/SeriesPlus.thy Probability/Caratheodory.thy \
+ Probability/Caratheodory.thy \
Probability/Borel.thy Probability/Measure.thy \
- Probability/Lebesgue.thy Probability/Product_Measure.thy \
+ Probability/Lebesgue_Integration.thy Probability/Lebesgue_Measure.thy \
+ Probability/Positive_Infinite_Real.thy Probability/Product_Measure.thy \
Probability/Probability_Space.thy Probability/Information.thy \
Probability/ex/Dining_Cryptographers.thy Library/FuncSet.thy \
- Library/Convex.thy Library/Product_Vector.thy \
- Library/Product_plus.thy Library/Inner_Product.thy \
- Library/Nat_Bijection.thy
- @cd Probability; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Probability
-
+ Probability/Lebesgue_Measure.thy \
+ Library/Nat_Bijection.thy Library/Countable.thy
+ @cd Probability; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Multivariate_Analysis HOL-Probability
## HOL-Nominal
--- a/src/HOL/Library/FuncSet.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Library/FuncSet.thy Tue Aug 24 08:22:17 2010 +0200
@@ -67,6 +67,10 @@
"f : Pi A B ==> (f x : B x ==> Q) ==> (x ~: A ==> Q) ==> Q"
by(auto simp: Pi_def)
+lemma Pi_cong:
+ "(\<And> w. w \<in> A \<Longrightarrow> f w = g w) \<Longrightarrow> f \<in> Pi A B \<longleftrightarrow> g \<in> Pi A B"
+ by (auto simp: Pi_def)
+
lemma funcset_id [simp]: "(\<lambda>x. x) \<in> A \<rightarrow> A"
by (auto intro: Pi_I)
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy Tue Aug 24 08:22:17 2010 +0200
@@ -5,149 +5,6 @@
imports Finite_Cartesian_Product Integration
begin
-instantiation prod :: (real_basis, real_basis) real_basis
-begin
-
-definition "basis i = (if i < DIM('a) then (basis i, 0) else (0, basis (i - DIM('a))))"
-
-instance
-proof
- let ?b = "basis :: nat \<Rightarrow> 'a \<times> 'b"
- let ?b_a = "basis :: nat \<Rightarrow> 'a"
- let ?b_b = "basis :: nat \<Rightarrow> 'b"
-
- note image_range =
- image_add_atLeastLessThan[symmetric, of 0 "DIM('a)" "DIM('b)", simplified]
-
- have split_range:
- "{..<DIM('b) + DIM('a)} = {..<DIM('a)} \<union> {DIM('a)..<DIM('b) + DIM('a)}"
- by auto
- have *: "?b ` {DIM('a)..<DIM('b) + DIM('a)} = {0} \<times> (?b_b ` {..<DIM('b)})"
- "?b ` {..<DIM('a)} = (?b_a ` {..<DIM('a)}) \<times> {0}"
- unfolding image_range image_image basis_prod_def_raw range_basis
- by (auto simp: zero_prod_def basis_eq_0_iff)
- hence b_split:
- "?b ` {..<DIM('b) + DIM('a)} = (?b_a ` {..<DIM('a)}) \<times> {0} \<union> {0} \<times> (?b_b ` {..<DIM('b)})" (is "_ = ?prod")
- by (subst split_range) (simp add: image_Un)
-
- have b_0: "?b ` {DIM('b) + DIM('a)..} = {0}" unfolding basis_prod_def_raw
- by (auto simp: zero_prod_def image_iff basis_eq_0_iff elim!: ballE[of _ _ "DIM('a) + DIM('b)"])
-
- have split_UNIV:
- "UNIV = {..<DIM('b) + DIM('a)} \<union> {DIM('b)+DIM('a)..}"
- by auto
-
- have range_b: "range ?b = ?prod \<union> {0}"
- by (subst split_UNIV) (simp add: image_Un b_split b_0)
-
- have prod: "\<And>f A B. setsum f (A \<times> B) = (\<Sum>a\<in>A. \<Sum>b\<in>B. f (a, b))"
- by (simp add: setsum_cartesian_product)
-
- show "span (range ?b) = UNIV"
- unfolding span_explicit range_b
- proof safe
- fix a::'a and b::'b
- from in_span_basis[of a] in_span_basis[of b]
- obtain Sa ua Sb ub where span:
- "finite Sa" "Sa \<subseteq> basis ` {..<DIM('a)}" "a = (\<Sum>v\<in>Sa. ua v *\<^sub>R v)"
- "finite Sb" "Sb \<subseteq> basis ` {..<DIM('b)}" "b = (\<Sum>v\<in>Sb. ub v *\<^sub>R v)"
- unfolding span_explicit by auto
-
- let ?S = "((Sa - {0}) \<times> {0} \<union> {0} \<times> (Sb - {0}))"
- have *:
- "?S \<inter> {v. fst v = 0} \<inter> {v. snd v = 0} = {}"
- "?S \<inter> - {v. fst v = 0} \<inter> {v. snd v = 0} = (Sa - {0}) \<times> {0}"
- "?S \<inter> {v. fst v = 0} \<inter> - {v. snd v = 0} = {0} \<times> (Sb - {0})"
- by (auto simp: zero_prod_def)
- show "\<exists>S u. finite S \<and> S \<subseteq> ?prod \<union> {0} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = (a, b)"
- apply (rule exI[of _ ?S])
- apply (rule exI[of _ "\<lambda>(v, w). (if w = 0 then ua v else 0) + (if v = 0 then ub w else 0)"])
- using span
- apply (simp add: prod_case_unfold setsum_addf if_distrib cond_application_beta setsum_cases prod *)
- by (auto simp add: setsum_prod intro!: setsum_mono_zero_cong_left)
- qed simp
-
- show "\<exists>d>0. ?b ` {d..} = {0} \<and> independent (?b ` {..<d}) \<and> inj_on ?b {..<d}"
- apply (rule exI[of _ "DIM('b) + DIM('a)"]) unfolding b_0
- proof (safe intro!: DIM_positive del: notI)
- show inj_on: "inj_on ?b {..<DIM('b) + DIM('a)}" unfolding split_range
- using inj_on_iff[OF basis_inj[where 'a='a]] inj_on_iff[OF basis_inj[where 'a='b]]
- by (auto intro!: inj_onI simp: basis_prod_def basis_eq_0_iff)
-
- show "independent (?b ` {..<DIM('b) + DIM('a)})"
- unfolding independent_eq_inj_on[OF inj_on]
- proof safe
- fix i u assume i_upper: "i < DIM('b) + DIM('a)" and
- "(\<Sum>j\<in>{..<DIM('b) + DIM('a)} - {i}. u (?b j) *\<^sub>R ?b j) = ?b i" (is "?SUM = _")
- let ?left = "{..<DIM('a)}" and ?right = "{DIM('a)..<DIM('b) + DIM('a)}"
- show False
- proof cases
- assume "i < DIM('a)"
- hence "(basis i, 0) = ?SUM" unfolding `?SUM = ?b i` unfolding basis_prod_def by auto
- also have "\<dots> = (\<Sum>j\<in>?left - {i}. u (?b j) *\<^sub>R ?b j) +
- (\<Sum>j\<in>?right. u (?b j) *\<^sub>R ?b j)"
- using `i < DIM('a)` by (subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong)
- also have "\<dots> = (\<Sum>j\<in>?left - {i}. u (?b_a j, 0) *\<^sub>R (?b_a j, 0)) +
- (\<Sum>j\<in>?right. u (0, ?b_b (j-DIM('a))) *\<^sub>R (0, ?b_b (j-DIM('a))))"
- unfolding basis_prod_def by auto
- finally have "basis i = (\<Sum>j\<in>?left - {i}. u (?b_a j, 0) *\<^sub>R ?b_a j)"
- by (simp add: setsum_prod)
- moreover
- note independent_basis[where 'a='a, unfolded independent_eq_inj_on[OF basis_inj]]
- note this[rule_format, of i "\<lambda>v. u (v, 0)"]
- ultimately show False using `i < DIM('a)` by auto
- next
- let ?i = "i - DIM('a)"
- assume not: "\<not> i < DIM('a)" hence "DIM('a) \<le> i" by auto
- hence "?i < DIM('b)" using `i < DIM('b) + DIM('a)` by auto
-
- have inj_on: "inj_on (\<lambda>j. j - DIM('a)) {DIM('a)..<DIM('b) + DIM('a)}"
- by (auto intro!: inj_onI)
- with i_upper not have *: "{..<DIM('b)} - {?i} = (\<lambda>j. j-DIM('a))`(?right - {i})"
- by (auto simp: inj_on_image_set_diff image_minus_const_atLeastLessThan_nat)
-
- have "(0, basis ?i) = ?SUM" unfolding `?SUM = ?b i`
- unfolding basis_prod_def using not `?i < DIM('b)` by auto
- also have "\<dots> = (\<Sum>j\<in>?left. u (?b j) *\<^sub>R ?b j) +
- (\<Sum>j\<in>?right - {i}. u (?b j) *\<^sub>R ?b j)"
- using not by (subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong)
- also have "\<dots> = (\<Sum>j\<in>?left. u (?b_a j, 0) *\<^sub>R (?b_a j, 0)) +
- (\<Sum>j\<in>?right - {i}. u (0, ?b_b (j-DIM('a))) *\<^sub>R (0, ?b_b (j-DIM('a))))"
- unfolding basis_prod_def by auto
- finally have "basis ?i = (\<Sum>j\<in>{..<DIM('b)} - {?i}. u (0, ?b_b j) *\<^sub>R ?b_b j)"
- unfolding *
- by (subst setsum_reindex[OF inj_on[THEN subset_inj_on]])
- (auto simp: setsum_prod)
- moreover
- note independent_basis[where 'a='b, unfolded independent_eq_inj_on[OF basis_inj]]
- note this[rule_format, of ?i "\<lambda>v. u (0, v)"]
- ultimately show False using `?i < DIM('b)` by auto
- qed
- qed
- qed
-qed
-end
-
-lemma DIM_prod[simp]: "DIM('a \<times> 'b) = DIM('b::real_basis) + DIM('a::real_basis)"
- by (rule dimension_eq) (auto simp: basis_prod_def zero_prod_def basis_eq_0_iff)
-
-instance prod :: (euclidean_space, euclidean_space) euclidean_space
-proof (default, safe)
- let ?b = "basis :: nat \<Rightarrow> 'a \<times> 'b"
- fix i j assume "i < DIM('a \<times> 'b)" "j < DIM('a \<times> 'b)"
- thus "?b i \<bullet> ?b j = (if i = j then 1 else 0)"
- unfolding basis_prod_def by (auto simp: dot_basis)
-qed
-
-instantiation prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space
-begin
-
-definition "x \<le> (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i<DIM('a\<times>'b). x $$ i \<le> y $$ i)"
-definition "x < (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i<DIM('a\<times>'b). x $$ i < y $$ i)"
-
-instance proof qed (auto simp: less_prod_def less_eq_prod_def)
-end
-
lemma delta_mult_idempotent:
"(if k=a then 1 else (0::'a::semiring_1)) * (if k=a then 1 else 0) = (if k=a then 1 else 0)" by (cases "k=a", auto)
@@ -1450,10 +1307,6 @@
unfolding nth_conv_component
using component_le_infnorm[of x] .
-lemma dist_nth_le_cart: "dist (x $ i) (y $ i) \<le> dist x y"
- unfolding dist_vector_def
- by (rule member_le_setL2) simp_all
-
instance cart :: (perfect_space, finite) perfect_space
proof
fix x :: "'a ^ 'b"
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1934,6 +1934,16 @@
assumes eucl_le: "x \<le> y \<longleftrightarrow> (\<forall>i < DIM('a). x $$ i \<le> y $$ i)"
and eucl_less: "x < y \<longleftrightarrow> (\<forall>i < DIM('a). x $$ i < y $$ i)"
+lemma eucl_less_not_refl[simp, intro!]: "\<not> x < (x::'a::ordered_euclidean_space)"
+ unfolding eucl_less[where 'a='a] by auto
+
+lemma euclidean_trans[trans]:
+ fixes x y z :: "'a::ordered_euclidean_space"
+ shows "x < y \<Longrightarrow> y < z \<Longrightarrow> x < z"
+ and "x \<le> y \<Longrightarrow> y < z \<Longrightarrow> x < z"
+ and "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z"
+ by (force simp: eucl_less[where 'a='a] eucl_le[where 'a='a])+
+
subsection {* Linearity and Bilinearity continued *}
lemma linear_bounded:
@@ -3388,4 +3398,150 @@
instance complex :: euclidean_space
proof qed (auto simp add: basis_complex_def inner_complex_def)
+section {* Products Spaces *}
+
+instantiation prod :: (real_basis, real_basis) real_basis
+begin
+
+definition "basis i = (if i < DIM('a) then (basis i, 0) else (0, basis (i - DIM('a))))"
+
+instance
+proof
+ let ?b = "basis :: nat \<Rightarrow> 'a \<times> 'b"
+ let ?b_a = "basis :: nat \<Rightarrow> 'a"
+ let ?b_b = "basis :: nat \<Rightarrow> 'b"
+
+ note image_range =
+ image_add_atLeastLessThan[symmetric, of 0 "DIM('a)" "DIM('b)", simplified]
+
+ have split_range:
+ "{..<DIM('b) + DIM('a)} = {..<DIM('a)} \<union> {DIM('a)..<DIM('b) + DIM('a)}"
+ by auto
+ have *: "?b ` {DIM('a)..<DIM('b) + DIM('a)} = {0} \<times> (?b_b ` {..<DIM('b)})"
+ "?b ` {..<DIM('a)} = (?b_a ` {..<DIM('a)}) \<times> {0}"
+ unfolding image_range image_image basis_prod_def_raw range_basis
+ by (auto simp: zero_prod_def basis_eq_0_iff)
+ hence b_split:
+ "?b ` {..<DIM('b) + DIM('a)} = (?b_a ` {..<DIM('a)}) \<times> {0} \<union> {0} \<times> (?b_b ` {..<DIM('b)})" (is "_ = ?prod")
+ by (subst split_range) (simp add: image_Un)
+
+ have b_0: "?b ` {DIM('b) + DIM('a)..} = {0}" unfolding basis_prod_def_raw
+ by (auto simp: zero_prod_def image_iff basis_eq_0_iff elim!: ballE[of _ _ "DIM('a) + DIM('b)"])
+
+ have split_UNIV:
+ "UNIV = {..<DIM('b) + DIM('a)} \<union> {DIM('b)+DIM('a)..}"
+ by auto
+
+ have range_b: "range ?b = ?prod \<union> {0}"
+ by (subst split_UNIV) (simp add: image_Un b_split b_0)
+
+ have prod: "\<And>f A B. setsum f (A \<times> B) = (\<Sum>a\<in>A. \<Sum>b\<in>B. f (a, b))"
+ by (simp add: setsum_cartesian_product)
+
+ show "span (range ?b) = UNIV"
+ unfolding span_explicit range_b
+ proof safe
+ fix a::'a and b::'b
+ from in_span_basis[of a] in_span_basis[of b]
+ obtain Sa ua Sb ub where span:
+ "finite Sa" "Sa \<subseteq> basis ` {..<DIM('a)}" "a = (\<Sum>v\<in>Sa. ua v *\<^sub>R v)"
+ "finite Sb" "Sb \<subseteq> basis ` {..<DIM('b)}" "b = (\<Sum>v\<in>Sb. ub v *\<^sub>R v)"
+ unfolding span_explicit by auto
+
+ let ?S = "((Sa - {0}) \<times> {0} \<union> {0} \<times> (Sb - {0}))"
+ have *:
+ "?S \<inter> {v. fst v = 0} \<inter> {v. snd v = 0} = {}"
+ "?S \<inter> - {v. fst v = 0} \<inter> {v. snd v = 0} = (Sa - {0}) \<times> {0}"
+ "?S \<inter> {v. fst v = 0} \<inter> - {v. snd v = 0} = {0} \<times> (Sb - {0})"
+ by (auto simp: zero_prod_def)
+ show "\<exists>S u. finite S \<and> S \<subseteq> ?prod \<union> {0} \<and> (\<Sum>v\<in>S. u v *\<^sub>R v) = (a, b)"
+ apply (rule exI[of _ ?S])
+ apply (rule exI[of _ "\<lambda>(v, w). (if w = 0 then ua v else 0) + (if v = 0 then ub w else 0)"])
+ using span
+ apply (simp add: prod_case_unfold setsum_addf if_distrib cond_application_beta setsum_cases prod *)
+ by (auto simp add: setsum_prod intro!: setsum_mono_zero_cong_left)
+ qed simp
+
+ show "\<exists>d>0. ?b ` {d..} = {0} \<and> independent (?b ` {..<d}) \<and> inj_on ?b {..<d}"
+ apply (rule exI[of _ "DIM('b) + DIM('a)"]) unfolding b_0
+ proof (safe intro!: DIM_positive del: notI)
+ show inj_on: "inj_on ?b {..<DIM('b) + DIM('a)}" unfolding split_range
+ using inj_on_iff[OF basis_inj[where 'a='a]] inj_on_iff[OF basis_inj[where 'a='b]]
+ by (auto intro!: inj_onI simp: basis_prod_def basis_eq_0_iff)
+
+ show "independent (?b ` {..<DIM('b) + DIM('a)})"
+ unfolding independent_eq_inj_on[OF inj_on]
+ proof safe
+ fix i u assume i_upper: "i < DIM('b) + DIM('a)" and
+ "(\<Sum>j\<in>{..<DIM('b) + DIM('a)} - {i}. u (?b j) *\<^sub>R ?b j) = ?b i" (is "?SUM = _")
+ let ?left = "{..<DIM('a)}" and ?right = "{DIM('a)..<DIM('b) + DIM('a)}"
+ show False
+ proof cases
+ assume "i < DIM('a)"
+ hence "(basis i, 0) = ?SUM" unfolding `?SUM = ?b i` unfolding basis_prod_def by auto
+ also have "\<dots> = (\<Sum>j\<in>?left - {i}. u (?b j) *\<^sub>R ?b j) +
+ (\<Sum>j\<in>?right. u (?b j) *\<^sub>R ?b j)"
+ using `i < DIM('a)` by (subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong)
+ also have "\<dots> = (\<Sum>j\<in>?left - {i}. u (?b_a j, 0) *\<^sub>R (?b_a j, 0)) +
+ (\<Sum>j\<in>?right. u (0, ?b_b (j-DIM('a))) *\<^sub>R (0, ?b_b (j-DIM('a))))"
+ unfolding basis_prod_def by auto
+ finally have "basis i = (\<Sum>j\<in>?left - {i}. u (?b_a j, 0) *\<^sub>R ?b_a j)"
+ by (simp add: setsum_prod)
+ moreover
+ note independent_basis[where 'a='a, unfolded independent_eq_inj_on[OF basis_inj]]
+ note this[rule_format, of i "\<lambda>v. u (v, 0)"]
+ ultimately show False using `i < DIM('a)` by auto
+ next
+ let ?i = "i - DIM('a)"
+ assume not: "\<not> i < DIM('a)" hence "DIM('a) \<le> i" by auto
+ hence "?i < DIM('b)" using `i < DIM('b) + DIM('a)` by auto
+
+ have inj_on: "inj_on (\<lambda>j. j - DIM('a)) {DIM('a)..<DIM('b) + DIM('a)}"
+ by (auto intro!: inj_onI)
+ with i_upper not have *: "{..<DIM('b)} - {?i} = (\<lambda>j. j-DIM('a))`(?right - {i})"
+ by (auto simp: inj_on_image_set_diff image_minus_const_atLeastLessThan_nat)
+
+ have "(0, basis ?i) = ?SUM" unfolding `?SUM = ?b i`
+ unfolding basis_prod_def using not `?i < DIM('b)` by auto
+ also have "\<dots> = (\<Sum>j\<in>?left. u (?b j) *\<^sub>R ?b j) +
+ (\<Sum>j\<in>?right - {i}. u (?b j) *\<^sub>R ?b j)"
+ using not by (subst setsum_Un_disjoint[symmetric]) (auto intro!: setsum_cong)
+ also have "\<dots> = (\<Sum>j\<in>?left. u (?b_a j, 0) *\<^sub>R (?b_a j, 0)) +
+ (\<Sum>j\<in>?right - {i}. u (0, ?b_b (j-DIM('a))) *\<^sub>R (0, ?b_b (j-DIM('a))))"
+ unfolding basis_prod_def by auto
+ finally have "basis ?i = (\<Sum>j\<in>{..<DIM('b)} - {?i}. u (0, ?b_b j) *\<^sub>R ?b_b j)"
+ unfolding *
+ by (subst setsum_reindex[OF inj_on[THEN subset_inj_on]])
+ (auto simp: setsum_prod)
+ moreover
+ note independent_basis[where 'a='b, unfolded independent_eq_inj_on[OF basis_inj]]
+ note this[rule_format, of ?i "\<lambda>v. u (0, v)"]
+ ultimately show False using `?i < DIM('b)` by auto
+ qed
+ qed
+ qed
+qed
end
+
+lemma DIM_prod[simp]: "DIM('a \<times> 'b) = DIM('b::real_basis) + DIM('a::real_basis)"
+ by (rule dimension_eq) (auto simp: basis_prod_def zero_prod_def basis_eq_0_iff)
+
+instance prod :: (euclidean_space, euclidean_space) euclidean_space
+proof (default, safe)
+ let ?b = "basis :: nat \<Rightarrow> 'a \<times> 'b"
+ fix i j assume "i < DIM('a \<times> 'b)" "j < DIM('a \<times> 'b)"
+ thus "?b i \<bullet> ?b j = (if i = j then 1 else 0)"
+ unfolding basis_prod_def by (auto simp: dot_basis)
+qed
+
+instantiation prod :: (ordered_euclidean_space, ordered_euclidean_space) ordered_euclidean_space
+begin
+
+definition "x \<le> (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i<DIM('a\<times>'b). x $$ i \<le> y $$ i)"
+definition "x < (y::('a\<times>'b)) \<longleftrightarrow> (\<forall>i<DIM('a\<times>'b). x $$ i < y $$ i)"
+
+instance proof qed (auto simp: less_prod_def less_eq_prod_def)
+end
+
+
+end
--- a/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Finite_Cartesian_Product.thy Tue Aug 24 08:22:17 2010 +0200
@@ -254,7 +254,7 @@
definition dist_vector_def:
"dist x y = setL2 (\<lambda>i. dist (x$i) (y$i)) UNIV"
-lemma dist_nth_le: "dist (x $ i) (y $ i) \<le> dist x y"
+lemma dist_nth_le_cart: "dist (x $ i) (y $ i) \<le> dist x y"
unfolding dist_vector_def
by (rule member_le_setL2) simp_all
@@ -283,7 +283,7 @@
apply (rule_tac x=e in exI, clarify)
apply (drule spec, erule mp, clarify)
apply (drule spec, drule spec, erule mp)
- apply (erule le_less_trans [OF dist_nth_le])
+ apply (erule le_less_trans [OF dist_nth_le_cart])
apply (subgoal_tac "\<forall>i\<in>UNIV. \<exists>e>0. \<forall>y. dist y (x$i) < e \<longrightarrow> y \<in> A i")
apply (drule finite_choice [OF finite], clarify)
apply (rule_tac x="Min (range f)" in exI, simp)
@@ -315,7 +315,7 @@
lemma Cauchy_Cart_nth:
"Cauchy (\<lambda>n. X n) \<Longrightarrow> Cauchy (\<lambda>n. X n $ i)"
-unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_nth_le])
+unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_nth_le_cart])
lemma Cauchy_vector:
fixes X :: "nat \<Rightarrow> 'a::metric_space ^ 'n"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Gauge_Measure.thy Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,3447 @@
+
+header {* Lebsegue measure (defined via the gauge integral). *}
+(* Author: John Harrison
+ Translation from HOL light: Robert Himmelmann, TU Muenchen *)
+
+theory Gauge_Measure
+ imports Integration
+begin
+
+(* ------------------------------------------------------------------------- *)
+(* Lebesgue measure (in the case where the measure is finite). *)
+(* For the non-finite version, please see Probability/Lebesgue_Measure.thy *)
+(* ------------------------------------------------------------------------- *)
+
+definition has_gmeasure (infixr "has'_gmeasure" 80) where
+ "s has_gmeasure m \<equiv> ((\<lambda>x. 1::real) has_integral m) s"
+
+definition gmeasurable :: "('n::ordered_euclidean_space) set \<Rightarrow> bool" where
+ "gmeasurable s \<equiv> (\<exists>m. s has_gmeasure m)"
+
+lemma gmeasurableI[dest]:"s has_gmeasure m \<Longrightarrow> gmeasurable s"
+ unfolding gmeasurable_def by auto
+
+definition gmeasure where
+ "gmeasure s \<equiv> (if gmeasurable s then (SOME m. s has_gmeasure m) else 0)"
+
+lemma has_gmeasure_measure: "gmeasurable s \<longleftrightarrow> s has_gmeasure (gmeasure s)"
+ unfolding gmeasure_def gmeasurable_def
+ apply meson apply(subst if_P) defer apply(rule someI) by auto
+
+lemma has_gmeasure_measureI[intro]:"gmeasurable s \<Longrightarrow> s has_gmeasure (gmeasure s)"
+ using has_gmeasure_measure by auto
+
+lemma has_gmeasure_unique: "s has_gmeasure m1 \<Longrightarrow> s has_gmeasure m2 \<Longrightarrow> m1 = m2"
+ unfolding has_gmeasure_def apply(rule has_integral_unique) by auto
+
+lemma measure_unique[intro]: assumes "s has_gmeasure m" shows "gmeasure s = m"
+ apply(rule has_gmeasure_unique[OF _ assms]) using assms
+ unfolding has_gmeasure_measure[THEN sym] by auto
+
+lemma has_gmeasure_measurable_measure:
+ "s has_gmeasure m \<longleftrightarrow> gmeasurable s \<and> gmeasure s = m"
+ by(auto intro!:measure_unique simp:has_gmeasure_measure[THEN sym])
+
+lemmas has_gmeasure_imp_measurable = gmeasurableI
+
+lemma has_gmeasure:
+ "s has_gmeasure m \<longleftrightarrow> ((\<lambda>x. if x \<in> s then 1 else 0) has_integral m) UNIV"
+ unfolding has_integral_restrict_univ has_gmeasure_def ..
+
+lemma gmeasurable: "gmeasurable s \<longleftrightarrow> (\<lambda>x. 1::real) integrable_on s"
+ unfolding gmeasurable_def integrable_on_def has_gmeasure_def by auto
+
+lemma gmeasurable_integrable:
+ "gmeasurable s \<longleftrightarrow> (\<lambda>x. if x \<in> s then 1 else (0::real)) integrable_on UNIV"
+ unfolding gmeasurable_def integrable_on_def has_gmeasure ..
+
+lemma measure_integral:
+ assumes "gmeasurable s" shows "gmeasure s = (integral s (\<lambda>x. 1))"
+ apply(rule integral_unique[THEN sym])
+ unfolding has_gmeasure_def[symmetric] using assms by auto
+
+lemma measure_integral_univ: assumes "gmeasurable s"
+ shows "gmeasure s = (integral UNIV (\<lambda>x. if x \<in> s then 1 else 0))"
+ apply(rule integral_unique[THEN sym])
+ using assms by(auto simp:has_gmeasure[THEN sym])
+
+lemmas integral_measure = measure_integral[THEN sym]
+
+lemmas integral_measure_univ = measure_integral_univ[THEN sym]
+
+lemma has_gmeasure_interval[intro]:
+ "{a..b} has_gmeasure content{a..b}" (is ?t1)
+ "{a<..<b} has_gmeasure content{a..b}" (is ?t2)
+proof- show ?t1 unfolding has_gmeasure_def using has_integral_const[where c="1::real"] by auto
+ show ?t2 unfolding has_gmeasure apply(rule has_integral_spike[of "{a..b} - {a<..<b}",
+ where f="\<lambda>x. (if x \<in> {a..b} then 1 else 0)"]) apply(rule negligible_frontier_interval)
+ using interval_open_subset_closed[of a b]
+ using `?t1` unfolding has_gmeasure by auto
+qed
+
+lemma gmeasurable_interval[intro]: "gmeasurable {a..b}" "gmeasurable {a<..<b}"
+ by(auto intro:gmeasurableI)
+
+lemma measure_interval[simp]: "gmeasure{a..b} = content{a..b}" "gmeasure({a<..<b}) = content{a..b}"
+ by(auto intro:measure_unique)
+
+lemma nonnegative_absolutely_integrable: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
+ assumes "\<forall>x\<in>s. \<forall>i<DIM('m). 0 \<le> f(x)$$i" "f integrable_on s"
+ shows "f absolutely_integrable_on s"
+ unfolding absolutely_integrable_abs_eq apply rule defer
+ apply(rule integrable_eq[of _ f]) using assms apply-apply(subst euclidean_eq) by auto
+
+lemma gmeasurable_inter[dest]: assumes "gmeasurable s" "gmeasurable t" shows "gmeasurable (s \<inter> t)"
+proof- have *:"(\<lambda>x. if x \<in> s \<inter> t then 1 else (0::real)) =
+ (\<lambda>x. \<chi>\<chi> i. min (((if x \<in> s then 1 else 0)::real)$$i) (((if x \<in> t then 1 else 0)::real)$$i))"
+ apply(rule ext) by auto
+ show ?thesis unfolding gmeasurable_integrable apply(rule absolutely_integrable_onD)
+ unfolding * apply(rule absolutely_integrable_min)
+ apply(rule_tac[!] nonnegative_absolutely_integrable)
+ using assms unfolding gmeasurable_integrable by auto
+qed
+
+lemma gmeasurable_union: assumes "gmeasurable s" "gmeasurable t"
+ shows "gmeasurable (s \<union> t)"
+proof- have *:"(\<lambda>x. if x \<in> s \<union> t then 1 else (0::real)) =
+ (\<lambda>x. \<chi>\<chi> i. max (((if x \<in> s then 1 else 0)::real)$$i) (((if x \<in> t then 1 else 0)::real)$$i)) "
+ by(rule ext,auto)
+ show ?thesis unfolding gmeasurable_integrable apply(rule absolutely_integrable_onD)
+ unfolding * apply(rule absolutely_integrable_max)
+ apply(rule_tac[!]nonnegative_absolutely_integrable)
+ using assms unfolding gmeasurable_integrable by auto
+qed
+
+lemma has_gmeasure_disjoint_union:
+ assumes "s1 has_gmeasure m1" "s2 has_gmeasure m2" "s1 \<inter> s2 = {}"
+ shows "(s1 \<union> s2) has_gmeasure (m1 + m2)"
+proof- have *:"\<And>x. (if x \<in> s1 then 1 else 0) + (if x \<in> s2 then 1 else 0) =
+ (if x \<in> s1 \<union> s2 then 1 else (0::real))" using assms(3) by auto
+ show ?thesis using has_integral_add[OF assms(1-2)[unfolded has_gmeasure]]
+ unfolding has_gmeasure * .
+qed
+
+lemma measure_disjoint_union: assumes "gmeasurable s" "gmeasurable t" "s \<inter> t = {}"
+ shows "gmeasure(s \<union> t) = gmeasure s + gmeasure t"
+ apply rule apply(rule has_gmeasure_disjoint_union) using assms by auto
+
+lemma has_gmeasure_pos_le[dest]: assumes "s has_gmeasure m" shows "0 \<le> m"
+ apply(rule has_integral_nonneg) using assms unfolding has_gmeasure by auto
+
+lemma not_measurable_measure:"\<not> gmeasurable s \<Longrightarrow> gmeasure s = 0"
+ unfolding gmeasure_def if_not_P ..
+
+lemma measure_pos_le[intro]: "0 <= gmeasure s"
+ apply(cases "gmeasurable s") unfolding not_measurable_measure
+ unfolding has_gmeasure_measure by auto
+
+lemma has_gmeasure_subset[dest]:
+ assumes "s1 has_gmeasure m1" "s2 has_gmeasure m2" "s1 \<subseteq> s2"
+ shows "m1 <= m2"
+ using has_integral_subset_le[OF assms(3,1,2)[unfolded has_gmeasure_def]] by auto
+
+lemma measure_subset[dest]: assumes "gmeasurable s" "gmeasurable t" "s \<subseteq> t"
+ shows "gmeasure s \<le> gmeasure t"
+ using assms unfolding has_gmeasure_measure by auto
+
+lemma has_gmeasure_0:"s has_gmeasure 0 \<longleftrightarrow> negligible s" (is "?l = ?r")
+proof assume ?r thus ?l unfolding indicator_def_raw negligible apply(erule_tac x="UNIV" in allE)
+ unfolding has_integral_restrict_univ has_gmeasure_def .
+next assume ?l note this[unfolded has_gmeasure_def has_integral_alt']
+ note * = conjunctD2[OF this,rule_format]
+ show ?r unfolding negligible_def
+ proof safe case goal1
+ from *(1)[of a b,unfolded integrable_on_def] guess y apply-
+ apply(subst (asm) has_integral_restrict_univ[THEN sym]) by (erule exE) note y=this
+ have "0 \<le> y" apply(rule has_integral_nonneg[OF y]) by auto
+ moreover have "y \<le> 0" apply(rule has_integral_le[OF y])
+ apply(rule `?l`[unfolded has_gmeasure_def has_integral_restrict_univ[THEN sym,of"\<lambda>x. 1"]]) by auto
+ ultimately have "y = 0" by auto
+ thus ?case using y unfolding has_integral_restrict_univ indicator_def_raw by auto
+ qed
+qed
+
+lemma measure_eq_0: "negligible s ==> gmeasure s = 0"
+ apply(rule measure_unique) unfolding has_gmeasure_0 by auto
+
+lemma has_gmeasure_empty[intro]: "{} has_gmeasure 0"
+ unfolding has_gmeasure_0 by auto
+
+lemma measure_empty[simp]: "gmeasure {} = 0"
+ apply(rule measure_eq_0) by auto
+
+lemma gmeasurable_empty[intro]: "gmeasurable {}" by(auto intro:gmeasurableI)
+
+lemma gmeasurable_measure_eq_0:
+ "gmeasurable s ==> (gmeasure s = 0 \<longleftrightarrow> negligible s)"
+ unfolding has_gmeasure_measure has_gmeasure_0[THEN sym] by(auto intro:measure_unique)
+
+lemma gmeasurable_measure_pos_lt:
+ "gmeasurable s ==> (0 < gmeasure s \<longleftrightarrow> ~negligible s)"
+ unfolding gmeasurable_measure_eq_0[THEN sym]
+ using measure_pos_le[of s] unfolding le_less by fastsimp
+
+lemma negligible_interval:True .. (*
+ "(!a b. negligible{a..b} \<longleftrightarrow> {a<..<b} = {}) \<and>
+ (!a b. negligible({a<..<b}) \<longleftrightarrow> {a<..<b} = {})"
+qed REWRITE_TAC[GSYM HAS_GMEASURE_0] THEN
+ MESON_TAC[HAS_GMEASURE_INTERVAL; CONTENT_EQ_0_INTERIOR;
+ INTERIOR_CLOSED_INTERVAL; HAS_GMEASURE_UNIQUE]);;*)
+
+lemma gmeasurable_finite_unions:
+ assumes "finite f" "\<And>s. s \<in> f \<Longrightarrow> gmeasurable s"
+ shows "gmeasurable (\<Union> f)" using assms(1,2)
+proof induct case (insert s F)
+ show ?case unfolding Union_insert apply(rule gmeasurable_union)
+ using insert by auto
+qed auto
+
+lemma has_gmeasure_diff_subset: assumes "s1 has_gmeasure m1" "s2 has_gmeasure m2" "s2 \<subseteq> s1"
+ shows "(s1 - s2) has_gmeasure (m1 - m2)"
+proof- have *:"(\<lambda>x. (if x \<in> s1 then 1 else 0) - (if x \<in> s2 then 1 else (0::real))) =
+ (\<lambda>x. if x \<in> s1 - s2 then 1 else 0)" apply(rule ext) using assms(3) by auto
+ show ?thesis using has_integral_sub[OF assms(1-2)[unfolded has_gmeasure]]
+ unfolding has_gmeasure * .
+qed
+
+lemma gmeasurable_diff: assumes "gmeasurable s" "gmeasurable t"
+ shows "gmeasurable (s - t)"
+proof- have *:"\<And>s t. gmeasurable s \<Longrightarrow> gmeasurable t \<Longrightarrow> t \<subseteq> s ==> gmeasurable (s - t)"
+ unfolding gmeasurable_def apply(erule exE)+ apply(rule,rule has_gmeasure_diff_subset)
+ by assumption+
+ have **:"s - t = s - (s \<inter> t)" by auto
+ show ?thesis unfolding ** apply(rule *) using assms by auto
+qed
+
+lemma measure_diff_subset: True .. (*
+ "!s t. gmeasurable s \<and> gmeasurable t \<and> t \<subseteq> s
+ ==> measure(s DIFF t) = gmeasure s - gmeasure t"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC MEASURE_UNIQUE THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_DIFF_SUBSET; GSYM HAS_GMEASURE_MEASURE]);; *)
+
+lemma has_gmeasure_union_negligible[dest]:
+ assumes "s has_gmeasure m" "negligible t"
+ shows "(s \<union> t) has_gmeasure m" unfolding has_gmeasure
+ apply(rule has_integral_spike[OF assms(2) _ assms(1)[unfolded has_gmeasure]]) by auto
+
+lemma has_gmeasure_diff_negligible[dest]:
+ assumes "s has_gmeasure m" "negligible t"
+ shows "(s - t) has_gmeasure m" unfolding has_gmeasure
+ apply(rule has_integral_spike[OF assms(2) _ assms(1)[unfolded has_gmeasure]]) by auto
+
+lemma has_gmeasure_union_negligible_eq: True .. (*
+ "!s t:real^N->bool m.
+ negligible t ==> ((s \<union> t) has_gmeasure m \<longleftrightarrow> s has_gmeasure m)"
+qed REPEAT STRIP_TAC THEN EQ_TAC THEN DISCH_TAC THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_UNION_NEGLIGIBLE] THEN
+ SUBST1_TAC(SET_RULE `s:real^N->bool = (s \<union> t) DIFF (t DIFF s)`) THEN
+ MATCH_MP_TAC HAS_GMEASURE_DIFF_NEGLIGIBLE THEN ASM_REWRITE_TAC[] THEN
+ MATCH_MP_TAC NEGLIGIBLE_DIFF THEN ASM_REWRITE_TAC[]);; *)
+
+lemma has_gmeasure_diff_negligible_eq: True .. (*
+ "!s t:real^N->bool m.
+ negligible t ==> ((s DIFF t) has_gmeasure m \<longleftrightarrow> s has_gmeasure m)"
+qed REPEAT STRIP_TAC THEN EQ_TAC THEN DISCH_TAC THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_DIFF_NEGLIGIBLE] THEN
+ SUBST1_TAC(SET_RULE `s:real^N->bool = (s DIFF t) \<union> (t \<inter> s)`) THEN
+ MATCH_MP_TAC HAS_GMEASURE_UNION_NEGLIGIBLE THEN
+ ASM_SIMP_TAC[NEGLIGIBLE_INTER]);; *)
+
+lemma has_gmeasure_almost: assumes "s has_gmeasure m" "negligible t" "s \<union> t = s' \<union> t"
+ shows "s' has_gmeasure m"
+proof- have *:"s' \<union> t - (t - s') = s'" by blast
+ show ?thesis using has_gmeasure_union_negligible[OF assms(1-2)] unfolding assms(3)
+ apply-apply(drule has_gmeasure_diff_negligible[where t="t - s'"])
+ apply(rule negligible_diff) using assms(2) unfolding * by auto
+qed
+
+lemma has_gmeasure_almost_eq: True .. (*
+ "!s s' t. negligible t \<and> s \<union> t = s' \<union> t
+ ==> (s has_gmeasure m \<longleftrightarrow> s' has_gmeasure m)"
+qed MESON_TAC[HAS_GMEASURE_ALMOST]);; *)
+
+lemma gmeasurable_almost: True .. (*
+ "!s s' t. gmeasurable s \<and> negligible t \<and> s \<union> t = s' \<union> t
+ ==> gmeasurable s'"
+qed REWRITE_TAC[measurable] THEN MESON_TAC[HAS_GMEASURE_ALMOST]);; *)
+
+lemma has_gmeasure_negligible_union:
+ assumes "s1 has_gmeasure m1" "s2 has_gmeasure m2" "negligible(s1 \<inter> s2)"
+ shows "(s1 \<union> s2) has_gmeasure (m1 + m2)"
+ apply(rule has_gmeasure_almost[of "(s1 - (s1 \<inter> s2)) \<union> (s2 - (s1 \<inter> s2))" _ "s1 \<inter> s2"])
+ apply(rule has_gmeasure_disjoint_union)
+ apply(rule has_gmeasure_almost[of s1,OF _ assms(3)]) prefer 3
+ apply(rule has_gmeasure_almost[of s2,OF _ assms(3)])
+ using assms by auto
+
+lemma measure_negligible_union: True .. (*
+ "!s t. gmeasurable s \<and> gmeasurable t \<and> negligible(s \<inter> t)
+ ==> measure(s \<union> t) = gmeasure s + gmeasure t"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC MEASURE_UNIQUE THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_NEGLIGIBLE_UNION; GSYM HAS_GMEASURE_MEASURE]);; *)
+
+lemma has_gmeasure_negligible_symdiff: True .. (*
+ "!s t:real^N->bool m.
+ s has_gmeasure m \<and>
+ negligible((s DIFF t) \<union> (t DIFF s))
+ ==> t has_gmeasure m"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_GMEASURE_ALMOST THEN
+ MAP_EVERY EXISTS_TAC
+ [`s:real^N->bool`; `(s DIFF t) \<union> (t DIFF s):real^N->bool`] THEN
+ ASM_REWRITE_TAC[] THEN SET_TAC[]);; *)
+
+lemma gmeasurable_negligible_symdiff: True .. (*
+ "!s t:real^N->bool.
+ gmeasurable s \<and> negligible((s DIFF t) \<union> (t DIFF s))
+ ==> gmeasurable t"
+qed REWRITE_TAC[measurable] THEN
+ MESON_TAC[HAS_GMEASURE_NEGLIGIBLE_SYMDIFF]);; *)
+
+lemma measure_negligible_symdiff: True .. (*
+ "!s t:real^N->bool.
+ (measurable s \/ gmeasurable t) \<and>
+ negligible((s DIFF t) \<union> (t DIFF s))
+ ==> gmeasure s = gmeasure t"
+qed MESON_TAC[HAS_GMEASURE_NEGLIGIBLE_SYMDIFF; MEASURE_UNIQUE; UNION_COMM;
+ HAS_GMEASURE_MEASURE]);; *)
+
+lemma has_gmeasure_negligible_unions: assumes "finite f"
+ "\<And>s. s \<in> f ==> s has_gmeasure (m s)"
+ "\<And>s t. s \<in> f \<Longrightarrow> t \<in> f \<Longrightarrow> ~(s = t) ==> negligible(s \<inter> t)"
+ shows "(\<Union> f) has_gmeasure (setsum m f)" using assms
+proof induct case (insert x s)
+ have *:"(x \<inter> \<Union>s) = \<Union>{x \<inter> y| y. y\<in>s}"by auto
+ show ?case unfolding Union_insert ring_class.setsum.insert[OF insert(1-2)]
+ apply(rule has_gmeasure_negligible_union) unfolding *
+ apply(rule insert) defer apply(rule insert) apply(rule insert) defer
+ apply(rule insert) prefer 4 apply(rule negligible_unions)
+ defer apply safe apply(rule insert) using insert by auto
+qed auto
+
+lemma measure_negligible_unions:
+ assumes "finite f" "\<And>s. s \<in> f ==> s has_gmeasure (m s)"
+ "\<And>s t. s \<in> f \<Longrightarrow> t \<in> f \<Longrightarrow> s \<noteq> t ==> negligible(s \<inter> t)"
+ shows "gmeasure(\<Union> f) = setsum m f"
+ apply rule apply(rule has_gmeasure_negligible_unions)
+ using assms by auto
+
+lemma has_gmeasure_disjoint_unions:
+ assumes"finite f" "\<And>s. s \<in> f ==> s has_gmeasure (m s)"
+ "\<And>s t. s \<in> f \<Longrightarrow> t \<in> f \<Longrightarrow> s \<noteq> t ==> s \<inter> t = {}"
+ shows "(\<Union> f) has_gmeasure (setsum m f)"
+ apply(rule has_gmeasure_negligible_unions[OF assms(1-2)]) using assms(3) by auto
+
+lemma measure_disjoint_unions:
+ assumes "finite f" "\<And>s. s \<in> f ==> s has_gmeasure (m s)"
+ "\<And>s t. s \<in> f \<Longrightarrow> t \<in> f \<Longrightarrow> s \<noteq> t ==> s \<inter> t = {}"
+ shows "gmeasure(\<Union> f) = setsum m f"
+ apply rule apply(rule has_gmeasure_disjoint_unions[OF assms]) by auto
+
+lemma has_gmeasure_negligible_unions_image:
+ assumes "finite s" "\<And>x. x \<in> s ==> gmeasurable(f x)"
+ "\<And>x y. x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x \<noteq> y \<Longrightarrow> negligible((f x) \<inter> (f y))"
+ shows "(\<Union> (f ` s)) has_gmeasure (setsum (\<lambda>x. gmeasure(f x)) s)"
+proof- have *:"setsum (\<lambda>x. gmeasure(f x)) s = setsum gmeasure (f ` s)"
+ apply(subst setsum_reindex_nonzero) defer
+ apply(subst gmeasurable_measure_eq_0)
+ proof- case goal2 thus ?case using assms(3)[of x y] by auto
+ qed(insert assms, auto)
+ show ?thesis unfolding * apply(rule has_gmeasure_negligible_unions) using assms by auto
+qed
+
+lemma measure_negligible_unions_image: True .. (*
+ "!f:A->real^N->bool s.
+ FINITE s \<and>
+ (!x. x \<in> s ==> gmeasurable(f x)) \<and>
+ (!x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) ==> negligible((f x) \<inter> (f y)))
+ ==> measure(UNIONS (IMAGE f s)) = sum s (\<lambda>x. measure(f x))"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC MEASURE_UNIQUE THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_NEGLIGIBLE_UNIONS_IMAGE]);;*)
+
+lemma has_gmeasure_disjoint_unions_image: True .. (*
+ "!f:A->real^N->bool s.
+ FINITE s \<and>
+ (!x. x \<in> s ==> gmeasurable(f x)) \<and>
+ (!x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) ==> DISJOINT (f x) (f y))
+ ==> (UNIONS (IMAGE f s)) has_gmeasure (sum s (\<lambda>x. measure(f x)))"
+qed REWRITE_TAC[DISJOINT] THEN REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC HAS_GMEASURE_NEGLIGIBLE_UNIONS_IMAGE THEN
+ ASM_SIMP_TAC[NEGLIGIBLE_EMPTY]);;*)
+
+lemma measure_disjoint_unions_image: True .. (*
+ "!f:A->real^N->bool s.
+ FINITE s \<and>
+ (!x. x \<in> s ==> gmeasurable(f x)) \<and>
+ (!x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) ==> DISJOINT (f x) (f y))
+ ==> measure(UNIONS (IMAGE f s)) = sum s (\<lambda>x. measure(f x))"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC MEASURE_UNIQUE THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_DISJOINT_UNIONS_IMAGE]);;*)
+
+lemma has_gmeasure_negligible_unions_image_strong: True .. (*
+ "!f:A->real^N->bool s.
+ FINITE {x | x \<in> s \<and> ~(f x = {})} \<and>
+ (!x. x \<in> s ==> gmeasurable(f x)) \<and>
+ (!x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) ==> negligible((f x) \<inter> (f y)))
+ ==> (UNIONS (IMAGE f s)) has_gmeasure (sum s (\<lambda>x. measure(f x)))"
+qed REPEAT STRIP_TAC THEN
+ MP_TAC(ISPECL [`f:A->real^N->bool`;
+ `{x | x \<in> s \<and> ~((f:A->real^N->bool) x = {})}`]
+ HAS_GMEASURE_NEGLIGIBLE_UNIONS_IMAGE) THEN
+ ASM_SIMP_TAC[IN_ELIM_THM; FINITE_RESTRICT] THEN
+ MATCH_MP_TAC EQ_IMP THEN BINOP_TAC THENL
+ [GEN_REWRITE_TAC I [EXTENSION] THEN
+ REWRITE_TAC[IN_UNIONS; IN_IMAGE; IN_ELIM_THM] THEN
+ MESON_TAC[NOT_IN_EMPTY];
+ CONV_TAC SYM_CONV THEN MATCH_MP_TAC SUM_SUPERSET THEN
+ SIMP_TAC[SUBSET; IN_ELIM_THM; TAUT `a \<and> ~(a \<and> b) \<longleftrightarrow> a \<and> ~b`] THEN
+ REWRITE_TAC[MEASURE_EMPTY]]);; *)
+
+lemma measure_negligible_unions_image_strong: True .. (*
+ "!f:A->real^N->bool s.
+ FINITE {x | x \<in> s \<and> ~(f x = {})} \<and>
+ (!x. x \<in> s ==> gmeasurable(f x)) \<and>
+ (!x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) ==> negligible((f x) \<inter> (f y)))
+ ==> measure(UNIONS (IMAGE f s)) = sum s (\<lambda>x. measure(f x))"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC MEASURE_UNIQUE THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_NEGLIGIBLE_UNIONS_IMAGE_STRONG]);; *)
+
+lemma has_gmeasure_disjoint_unions_image_strong: True .. (*
+ "!f:A->real^N->bool s.
+ FINITE {x | x \<in> s \<and> ~(f x = {})} \<and>
+ (!x. x \<in> s ==> gmeasurable(f x)) \<and>
+ (!x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) ==> DISJOINT (f x) (f y))
+ ==> (UNIONS (IMAGE f s)) has_gmeasure (sum s (\<lambda>x. measure(f x)))"
+qed REWRITE_TAC[DISJOINT] THEN REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC HAS_GMEASURE_NEGLIGIBLE_UNIONS_IMAGE_STRONG THEN
+ ASM_SIMP_TAC[NEGLIGIBLE_EMPTY]);; *)
+
+lemma measure_disjoint_unions_image_strong: True .. (*
+ "!f:A->real^N->bool s.
+ FINITE {x | x \<in> s \<and> ~(f x = {})} \<and>
+ (!x. x \<in> s ==> gmeasurable(f x)) \<and>
+ (!x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) ==> DISJOINT (f x) (f y))
+ ==> measure(UNIONS (IMAGE f s)) = sum s (\<lambda>x. measure(f x))"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC MEASURE_UNIQUE THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_DISJOINT_UNIONS_IMAGE_STRONG]);; *)
+
+lemma measure_union: True .. (*
+ "!s t:real^N->bool.
+ gmeasurable s \<and> gmeasurable t
+ ==> measure(s \<union> t) = measure(s) + measure(t) - measure(s \<inter> t)"
+qed REPEAT STRIP_TAC THEN
+ ONCE_REWRITE_TAC[SET_RULE
+ `s \<union> t = (s \<inter> t) \<union> (s DIFF t) \<union> (t DIFF s)`] THEN
+ ONCE_REWRITE_TAC[REAL_ARITH `a + b - c = c + (a - c) + (b - c)`] THEN
+ MP_TAC(ISPECL [`s DIFF t:real^N->bool`; `t DIFF s:real^N->bool`]
+ MEASURE_DISJOINT_UNION) THEN
+ ASM_SIMP_TAC[MEASURABLE_DIFF] THEN
+ ANTS_TAC THENL [SET_TAC[]; ALL_TAC] THEN
+ MP_TAC(ISPECL [`s \<inter> t:real^N->bool`;
+ `(s DIFF t) \<union> (t DIFF s):real^N->bool`]
+ MEASURE_DISJOINT_UNION) THEN
+ ASM_SIMP_TAC[MEASURABLE_DIFF; GMEASURABLE_UNION; GMEASURABLE_INTER] THEN
+ ANTS_TAC THENL [SET_TAC[]; ALL_TAC] THEN
+ REPEAT(DISCH_THEN SUBST1_TAC) THEN AP_TERM_TAC THEN BINOP_TAC THEN
+ REWRITE_TAC[REAL_EQ_SUB_LADD] THEN MATCH_MP_TAC EQ_TRANS THENL
+ [EXISTS_TAC `measure((s DIFF t) \<union> (s \<inter> t):real^N->bool)`;
+ EXISTS_TAC `measure((t DIFF s) \<union> (s \<inter> t):real^N->bool)`] THEN
+ (CONJ_TAC THENL
+ [CONV_TAC SYM_CONV THEN MATCH_MP_TAC MEASURE_DISJOINT_UNION THEN
+ ASM_SIMP_TAC[MEASURABLE_DIFF; GMEASURABLE_INTER];
+ AP_TERM_TAC] THEN
+ SET_TAC[]));; *)
+
+lemma measure_union_le: True .. (*
+ "!s t:real^N->bool.
+ gmeasurable s \<and> gmeasurable t
+ ==> measure(s \<union> t) <= gmeasure s + gmeasure t"
+qed REPEAT STRIP_TAC THEN ASM_SIMP_TAC[MEASURE_UNION] THEN
+ REWRITE_TAC[REAL_ARITH `a + b - c <= a + b \<longleftrightarrow> 0 <= c`] THEN
+ MATCH_MP_TAC MEASURE_POS_LE THEN ASM_SIMP_TAC[MEASURABLE_INTER]);; *)
+
+lemma measure_unions_le: True .. (*
+ "!f:(real^N->bool)->bool.
+ FINITE f \<and> (!s. s \<in> f ==> gmeasurable s)
+ ==> measure(UNIONS f) <= sum f (\<lambda>s. gmeasure s)"
+qed REWRITE_TAC[IMP_CONJ] THEN
+ MATCH_MP_TAC FINITE_INDUCT_STRONG THEN
+ SIMP_TAC[UNIONS_0; UNIONS_INSERT; SUM_CLAUSES] THEN
+ REWRITE_TAC[MEASURE_EMPTY; REAL_LE_REFL] THEN
+ MAP_EVERY X_GEN_TAC [`s:real^N->bool`; `f:(real^N->bool)->bool`] THEN
+ REWRITE_TAC[IN_INSERT] THEN REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `measure(s:real^N->bool) + measure(UNIONS f:real^N->bool)` THEN
+ ASM_SIMP_TAC[MEASURE_UNION_LE; GMEASURABLE_UNIONS] THEN
+ REWRITE_TAC[REAL_LE_LADD] THEN FIRST_X_ASSUM MATCH_MP_TAC THEN
+ ASM_SIMP_TAC[]);; *)
+
+lemma measure_unions_le_image: True .. (*
+ "!f:A->bool s:A->(real^N->bool).
+ FINITE f \<and> (!a. a \<in> f ==> gmeasurable(s a))
+ ==> measure(UNIONS (IMAGE s f)) <= sum f (\<lambda>a. measure(s a))"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `sum (IMAGE s (f:A->bool)) (\<lambda>k:real^N->bool. gmeasure k)` THEN
+ ASM_SIMP_TAC[MEASURE_UNIONS_LE; FORALL_IN_IMAGE; FINITE_IMAGE] THEN
+ GEN_REWRITE_TAC (RAND_CONV o RAND_CONV) [GSYM o_DEF] THEN
+ REWRITE_TAC[ETA_AX] THEN MATCH_MP_TAC SUM_IMAGE_LE THEN
+ ASM_SIMP_TAC[MEASURE_POS_LE]);; *)
+
+lemma gmeasurable_inner_outer: True .. (*
+ "!s:real^N->bool.
+ gmeasurable s \<longleftrightarrow>
+ !e. 0 < e
+ ==> ?t u. t \<subseteq> s \<and> s \<subseteq> u \<and>
+ gmeasurable t \<and> gmeasurable u \<and>
+ abs(measure t - gmeasure u) < e"
+qed GEN_TAC THEN EQ_TAC THEN DISCH_TAC THENL
+ [GEN_TAC THEN DISCH_TAC THEN REPEAT(EXISTS_TAC `s:real^N->bool`) THEN
+ ASM_REWRITE_TAC[SUBSET_REFL; REAL_SUB_REFL; REAL_ABS_NUM];
+ ALL_TAC] THEN
+ REWRITE_TAC[MEASURABLE_INTEGRABLE] THEN MATCH_MP_TAC INTEGRABLE_STRADDLE THEN
+ X_GEN_TAC `e:real` THEN DISCH_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `e:real`) THEN
+ ASM_REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
+ MAP_EVERY X_GEN_TAC [`t:real^N->bool`; `u:real^N->bool`] THEN STRIP_TAC THEN
+ MAP_EVERY EXISTS_TAC
+ [`(\<lambda>x. if x \<in> t then 1 else 0):real^N->real^1`;
+ `(\<lambda>x. if x \<in> u then 1 else 0):real^N->real^1`;
+ `lift(measure(t:real^N->bool))`;
+ `lift(measure(u:real^N->bool))`] THEN
+ ASM_REWRITE_TAC[GSYM HAS_GMEASURE; GSYM HAS_GMEASURE_MEASURE] THEN
+ ASM_REWRITE_TAC[GSYM LIFT_SUB; NORM_LIFT] THEN REPEAT STRIP_TAC THEN
+ REPEAT(COND_CASES_TAC THEN
+ ASM_REWRITE_TAC[_VEC; REAL_POS; REAL_LE_REFL]) THEN
+ ASM SET_TAC[]);; *)
+
+lemma has_gmeasure_inner_outer: True .. (*
+ "!s:real^N->bool m.
+ s has_gmeasure m \<longleftrightarrow>
+ (!e. 0 < e ==> ?t. t \<subseteq> s \<and> gmeasurable t \<and>
+ m - e < gmeasure t) \<and>
+ (!e. 0 < e ==> ?u. s \<subseteq> u \<and> gmeasurable u \<and>
+ gmeasure u < m + e)"
+qed REPEAT GEN_TAC THEN
+ GEN_REWRITE_TAC LAND_CONV [HAS_GMEASURE_MEASURABLE_MEASURE] THEN EQ_TAC THENL
+ [REPEAT STRIP_TAC THEN EXISTS_TAC `s:real^N->bool` THEN
+ ASM_REWRITE_TAC[SUBSET_REFL] THEN ASM_REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ DISCH_THEN(CONJUNCTS_THEN2 (LABEL_TAC "t") (LABEL_TAC "u")) THEN
+ MATCH_MP_TAC(TAUT `a \<and> (a ==> b) ==> a \<and> b`) THEN CONJ_TAC THENL
+ [GEN_REWRITE_TAC I [MEASURABLE_INNER_OUTER] THEN
+ X_GEN_TAC `e:real` THEN DISCH_TAC THEN
+ REMOVE_THEN "u" (MP_TAC o SPEC `e / 2`) THEN
+ REMOVE_THEN "t" (MP_TAC o SPEC `e / 2`) THEN
+ ASM_SIMP_TAC[REAL_LT_DIV; REAL_OF_NUM_LT; ARITH] THEN
+ REWRITE_TAC[IMP_IMP; LEFT_AND_EXISTS_THM] THEN
+ REWRITE_TAC[RIGHT_AND_EXISTS_THM] THEN
+ REPEAT(MATCH_MP_TAC MONO_EXISTS THEN GEN_TAC) THEN
+ STRIP_TAC THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC(REAL_ARITH
+ `0 < e \<and> t <= u \<and> m - e / 2 < t \<and> u < m + e / 2
+ ==> abs(t - u) < e`) THEN
+ ASM_REWRITE_TAC[] THEN MATCH_MP_TAC MEASURE_SUBSET THEN
+ ASM_REWRITE_TAC[] THEN ASM SET_TAC[];
+ DISCH_TAC THEN MATCH_MP_TAC(REAL_ARITH
+ `~(0 < x - y) \<and> ~(0 < y - x) ==> x = y`) THEN
+ CONJ_TAC THEN DISCH_TAC THENL
+ [REMOVE_THEN "u" (MP_TAC o SPEC `measure(s:real^N->bool) - m`) THEN
+ ASM_REWRITE_TAC[REAL_SUB_ADD2; GSYM REAL_NOT_LE];
+ REMOVE_THEN "t" (MP_TAC o SPEC `m - measure(s:real^N->bool)`) THEN
+ ASM_REWRITE_TAC[REAL_SUB_SUB2; GSYM REAL_NOT_LE]] THEN
+ ASM_MESON_TAC[MEASURE_SUBSET]]);; *)
+
+lemma has_gmeasure_inner_outer_le: True .. (*
+ "!s:real^N->bool m.
+ s has_gmeasure m \<longleftrightarrow>
+ (!e. 0 < e ==> ?t. t \<subseteq> s \<and> gmeasurable t \<and>
+ m - e <= gmeasure t) \<and>
+ (!e. 0 < e ==> ?u. s \<subseteq> u \<and> gmeasurable u \<and>
+ gmeasure u <= m + e)"
+qed REWRITE_TAC[HAS_GMEASURE_INNER_OUTER] THEN
+ MESON_TAC[REAL_ARITH `0 < e \<and> m - e / 2 <= t ==> m - e < t`;
+ REAL_ARITH `0 < e \<and> u <= m + e / 2 ==> u < m + e`;
+ REAL_ARITH `0 < e \<longleftrightarrow> 0 < e / 2`; REAL_LT_IMP_LE]);; *)
+
+lemma has_gmeasure_limit: True .. (*
+ "!s. s has_gmeasure m \<longleftrightarrow>
+ !e. 0 < e
+ ==> ?B. 0 < B \<and>
+ !a b. ball(0,B) \<subseteq> {a..b}
+ ==> ?z. (s \<inter> {a..b}) has_gmeasure z \<and>
+ abs(z - m) < e"
+qed GEN_TAC THEN REWRITE_TAC[HAS_GMEASURE] THEN
+ GEN_REWRITE_TAC LAND_CONV [HAS_INTEGRAL] THEN
+ REWRITE_TAC[IN_UNIV] THEN
+ GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV)
+ [GSYM HAS_INTEGRAL_RESTRICT_UNIV] THEN
+ REWRITE_TAC[MESON[IN_INTER]
+ `(if x \<in> k \<inter> s then a else b) =
+ (if x \<in> s then if x \<in> k then a else b else b)`] THEN
+ REWRITE_TAC[EXISTS_LIFT; GSYM LIFT_SUB; NORM_LIFT]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* properties of gmeasure under simple affine transformations. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma has_gmeasure_affinity: True .. (*
+ "!s m c y. s has_gmeasure y
+ ==> (IMAGE (\<lambda>x:real^N. m % x + c) s)
+ has_gmeasure abs(m) pow (dimindex(:N)) * y"
+qed REPEAT GEN_TAC THEN ASM_CASES_TAC `m = 0` THENL
+ [ASM_REWRITE_TAC[REAL_ABS_NUM; VECTOR_ADD_LID; VECTOR_MUL_LZERO] THEN
+ ONCE_REWRITE_TAC[MATCH_MP (ARITH_RULE `~(x = 0) ==> x = SUC(x - 1)`)
+ (SPEC_ALL DIMINDEX_NONZERO)] THEN DISCH_TAC THEN
+ REWRITE_TAC[real_pow; REAL_MUL_LZERO; HAS_GMEASURE_0] THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN EXISTS_TAC `{c:real^N}` THEN
+ SIMP_TAC[NEGLIGIBLE_FINITE; FINITE_RULES] THEN SET_TAC[];
+ ALL_TAC] THEN
+ REWRITE_TAC[HAS_GMEASURE] THEN
+ ONCE_REWRITE_TAC[HAS_INTEGRAL] THEN REWRITE_TAC[IN_UNIV] THEN
+ DISCH_TAC THEN X_GEN_TAC `e:real` THEN DISCH_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `e:real / abs(m) pow dimindex(:N)`) THEN
+ ASM_SIMP_TAC[REAL_LT_DIV; REAL_POW_LT; GSYM REAL_ABS_NZ; REAL_POW_LT] THEN
+ DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN
+ EXISTS_TAC `abs(m) * B + norm(c:real^N)` THEN
+ ASM_SIMP_TAC[REAL_ARITH `0 < B \<and> 0 <= x ==> 0 < B + x`;
+ NORM_POS_LE; REAL_LT_MUL; GSYM REAL_ABS_NZ; REAL_POW_LT] THEN
+ MAP_EVERY X_GEN_TAC [`u:real^N`; `v:real^N`] THEN DISCH_TAC THEN
+ REWRITE_TAC[IN_IMAGE] THEN
+ ASM_SIMP_TAC[VECTOR_EQ_AFFINITY; UNWIND_THM1] THEN
+ FIRST_X_ASSUM(MP_TAC o SPECL
+ [`if 0 <= m then inv m % u + --(inv m % c):real^N
+ else inv m % v + --(inv m % c)`;
+ `if 0 <= m then inv m % v + --(inv m % c):real^N
+ else inv m % u + --(inv m % c)`]) THEN
+ MATCH_MP_TAC(TAUT `a \<and> (a ==> b ==> c) ==> (a ==> b) ==> c`) THEN
+ CONJ_TAC THENL
+ [REWRITE_TAC[SUBSET] THEN X_GEN_TAC `x:real^N` THEN
+ FIRST_X_ASSUM(MP_TAC o GEN_REWRITE_RULE I [SUBSET]) THEN
+ DISCH_THEN(MP_TAC o SPEC `m % x + c:real^N`) THEN
+ MATCH_MP_TAC MONO_IMP THEN REWRITE_TAC[IN_BALL; IN_INTERVAL] THEN
+ CONJ_TAC THENL
+ [REWRITE_TAC[NORM_ARITH `dist(0,x) = norm(x:real^N)`] THEN
+ DISCH_TAC THEN MATCH_MP_TAC(NORM_ARITH
+ `norm(x:real^N) < a ==> norm(x + y) < a + norm(y)`) THEN
+ ASM_SIMP_TAC[NORM_MUL; REAL_LT_LMUL; GSYM REAL_ABS_NZ];
+ ALL_TAC] THEN
+ SIMP_TAC[VECTOR_ADD_COMPONENT; VECTOR_MUL_COMPONENT; VECTOR_NEG_COMPONENT;
+ COND_COMPONENT] THEN
+ MATCH_MP_TAC MONO_FORALL THEN GEN_TAC THEN MATCH_MP_TAC MONO_IMP THEN
+ REWRITE_TAC[REAL_ARITH `m * u + --(m * c):real = (u - c) * m`] THEN
+ SUBST1_TAC(REAL_ARITH
+ `inv(m) = if 0 <= inv(m) then abs(inv m) else --(abs(inv m))`) THEN
+ SIMP_TAC[REAL_LE_INV_EQ] THEN
+ REWRITE_TAC[REAL_ARITH `(x - y:real) * --z = (y - x) * z`] THEN
+ REWRITE_TAC[REAL_ABS_INV; GSYM real_div] THEN COND_CASES_TAC THEN
+ ASM_SIMP_TAC[REAL_LE_LDIV_EQ; REAL_LE_RDIV_EQ; GSYM REAL_ABS_NZ] THEN
+ ASM_REWRITE_TAC[real_abs] THEN REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ REWRITE_TAC[SUBSET] THEN DISCH_THEN(MP_TAC o SPEC `0:real^N`) THEN
+ ASM_REWRITE_TAC[CENTRE_IN_BALL] THEN DISCH_TAC THEN
+ DISCH_THEN(X_CHOOSE_THEN `z:real^1`
+ (fun th -> EXISTS_TAC `(abs m pow dimindex (:N)) % z:real^1` THEN
+ MP_TAC th)) THEN
+ DISCH_THEN(CONJUNCTS_THEN2 MP_TAC ASSUME_TAC) THEN
+ FIRST_ASSUM(MP_TAC o MATCH_MP(REAL_FIELD `~(x = 0) ==> ~(inv x = 0)`)) THEN
+ REWRITE_TAC[TAUT `a ==> b ==> c \<longleftrightarrow> b \<and> a ==> c`] THEN
+ DISCH_THEN(MP_TAC o SPEC `--(inv m % c):real^N` o
+ MATCH_MP HAS_INTEGRAL_AFFINITY) THEN
+ ASM_REWRITE_TAC[IMAGE_AFFINITY_INTERVAL; REAL_INV_INV] THEN
+ SIMP_TAC[COND_ID] THEN COND_CASES_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
+ REWRITE_TAC[VECTOR_ADD_LDISTRIB; VECTOR_MUL_ASSOC;
+ VECTOR_MUL_LNEG; VECTOR_MUL_RNEG] THEN
+ ASM_SIMP_TAC[REAL_MUL_RINV; VECTOR_MUL_LID; VECTOR_NEG_NEG] THEN
+ REWRITE_TAC[VECTOR_ARITH `(u + --c) + c:real^N = u`] THEN
+ REWRITE_TAC[REAL_ABS_INV; REAL_INV_INV; GSYM REAL_POW_INV] THEN
+ DISCH_THEN(fun th -> REWRITE_TAC[th]) THEN
+ REWRITE_TAC[LIFT_CMUL; GSYM VECTOR_SUB_LDISTRIB] THEN
+ REWRITE_TAC[NORM_MUL; REAL_ABS_POW; REAL_ABS_ABS] THEN
+ ONCE_REWRITE_TAC[REAL_MUL_SYM] THEN
+ ASM_SIMP_TAC[GSYM REAL_LT_RDIV_EQ; REAL_POW_LT; GSYM REAL_ABS_NZ]);; *)
+
+lemma stretch_galois: True .. (*
+ "!x:real^N y:real^N m.
+ (!k. 1 <= k \<and> k <= dimindex(:N) ==> ~(m k = 0))
+ ==> ((y = (lambda k. m k * x$k)) \<longleftrightarrow> (lambda k. inv(m k) * y$k) = x)"
+qed REPEAT GEN_TAC THEN SIMP_TAC[CART_EQ; LAMBDA_BETA] THEN
+ MATCH_MP_TAC(MESON[]
+ `(!x. p x ==> (q x \<longleftrightarrow> r x))
+ ==> (!x. p x) ==> ((!x. q x) \<longleftrightarrow> (!x. r x))`) THEN
+ GEN_TAC THEN ASM_CASES_TAC `1 <= k \<and> k <= dimindex(:N)` THEN
+ ASM_REWRITE_TAC[] THEN CONV_TAC REAL_FIELD);; *)
+
+lemma has_gmeasure_stretch: True .. (*
+ "!s m y. s has_gmeasure y
+ ==> (IMAGE (\<lambda>x:real^N. lambda k. m k * x$k) s :real^N->bool)
+ has_gmeasure abs(product (1..dimindex(:N)) m) * y"
+qed REPEAT STRIP_TAC THEN ASM_CASES_TAC
+ `!k. 1 <= k \<and> k <= dimindex(:N) ==> ~(m k = 0)`
+ THENL
+ [ALL_TAC;
+ FIRST_X_ASSUM(MP_TAC o GEN_REWRITE_RULE I [NOT_FORALL_THM]) THEN
+ REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; LEFT_IMP_EXISTS_THM] THEN
+ X_GEN_TAC `k:num` THEN STRIP_TAC THEN
+ SUBGOAL_THEN `product(1..dimindex (:N)) m = 0` SUBST1_TAC THENL
+ [ASM_MESON_TAC[PRODUCT_EQ_0_NUMSEG]; ALL_TAC] THEN
+ REWRITE_TAC[REAL_ABS_NUM; REAL_MUL_LZERO; HAS_GMEASURE_0] THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `{x:real^N | x$k = 0}` THEN
+ ASM_SIMP_TAC[NEGLIGIBLE_STANDARD_HYPERPLANE; SUBSET; FORALL_IN_IMAGE] THEN
+ ASM_SIMP_TAC[IN_ELIM_THM; LAMBDA_BETA; REAL_MUL_LZERO]] THEN
+ UNDISCH_TAC `(s:real^N->bool) has_gmeasure y` THEN
+ REWRITE_TAC[HAS_GMEASURE] THEN
+ ONCE_REWRITE_TAC[HAS_INTEGRAL] THEN REWRITE_TAC[IN_UNIV] THEN
+ DISCH_TAC THEN X_GEN_TAC `e:real` THEN DISCH_TAC THEN
+ SUBGOAL_THEN `0 < abs(product(1..dimindex(:N)) m)` ASSUME_TAC THENL
+ [ASM_MESON_TAC[REAL_ABS_NZ; REAL_LT_DIV; PRODUCT_EQ_0_NUMSEG];
+ ALL_TAC] THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `e:real / abs(product(1..dimindex(:N)) m)`) THEN
+ ASM_SIMP_TAC[REAL_LT_DIV] THEN
+ DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN
+ EXISTS_TAC `sup(IMAGE (\<lambda>k. abs(m k) * B) (1..dimindex(:N)))` THEN
+ MATCH_MP_TAC(TAUT `a \<and> (a ==> b) ==> a \<and> b`) THEN CONJ_TAC THENL
+ [ASM_SIMP_TAC[REAL_LT_SUP_FINITE; FINITE_IMAGE; NUMSEG_EMPTY; FINITE_NUMSEG;
+ IN_NUMSEG; GSYM NOT_LE; DIMINDEX_GE_1; IMAGE_EQ_EMPTY;
+ EXISTS_IN_IMAGE] THEN
+ ASM_MESON_TAC[IN_NUMSEG; DIMINDEX_GE_1; LE_REFL; REAL_LT_MUL; REAL_ABS_NZ];
+ DISCH_TAC] THEN
+ MAP_EVERY X_GEN_TAC [`u:real^N`; `v:real^N`] THEN DISCH_TAC THEN
+ ASM_SIMP_TAC[IN_IMAGE; STRETCH_GALOIS; UNWIND_THM1] THEN
+ FIRST_X_ASSUM(MP_TAC o SPECL
+ [`(lambda k. min (inv(m k) * (u:real^N)$k)
+ (inv(m k) * (v:real^N)$k)):real^N`;
+ `(lambda k. max (inv(m k) * (u:real^N)$k)
+ (inv(m k) * (v:real^N)$k)):real^N`]) THEN
+ MATCH_MP_TAC(TAUT `a \<and> (b ==> a ==> c) ==> (a ==> b) ==> c`) THEN
+ CONJ_TAC THENL
+ [ALL_TAC;
+ REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN X_GEN_TAC `z:real^1` THEN
+ DISCH_THEN(CONJUNCTS_THEN2 MP_TAC ASSUME_TAC) THEN
+ SUBGOAL_THEN `!k. 1 <= k \<and> k <= dimindex (:N) ==> ~(inv(m k) = 0)`
+ MP_TAC THENL [ASM_SIMP_TAC[REAL_INV_EQ_0]; ALL_TAC] THEN
+ ONCE_REWRITE_TAC[GSYM IMP_CONJ_ALT] THEN
+ DISCH_THEN(MP_TAC o MATCH_MP HAS_INTEGRAL_STRETCH)] THEN
+ (MP_TAC(ISPECL [`u:real^N`; `v:real^N`; `\i:num. inv(m i)`]
+ IMAGE_STRETCH_INTERVAL) THEN
+ SUBGOAL_THEN `~(interval[u:real^N,v] = {})` ASSUME_TAC THENL
+ [FIRST_ASSUM(MATCH_MP_TAC o MATCH_MP (SET_RULE
+ `s \<subseteq> t ==> ~(s = {}) ==> ~(t = {})`)) THEN
+ ASM_REWRITE_TAC[BALL_EQ_EMPTY; GSYM REAL_NOT_LT];
+ ALL_TAC] THEN
+ ASM_REWRITE_TAC[] THEN DISCH_THEN(SUBST1_TAC o SYM))
+ THENL
+ [FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (SET_RULE
+ `b \<subseteq> s ==> b' \<subseteq> IMAGE f b ==> b' \<subseteq> IMAGE f s`)) THEN
+ REWRITE_TAC[IN_BALL; SUBSET; NORM_ARITH `dist(0,x) = norm x`;
+ IN_IMAGE] THEN
+ ASM_SIMP_TAC[STRETCH_GALOIS; REAL_INV_EQ_0; UNWIND_THM1; REAL_INV_INV] THEN
+ X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+ MATCH_MP_TAC REAL_LET_TRANS THEN
+ EXISTS_TAC
+ `norm(sup(IMAGE(\<lambda>k. abs(m k)) (1..dimindex(:N))) % x:real^N)` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC NORM_LE_COMPONENTWISE THEN
+ SIMP_TAC[LAMBDA_BETA; VECTOR_MUL_COMPONENT; REAL_ABS_MUL] THEN
+ REPEAT STRIP_TAC THEN MATCH_MP_TAC REAL_LE_RMUL THEN
+ REWRITE_TAC[REAL_ABS_POS] THEN
+ MATCH_MP_TAC(REAL_ARITH `x <= y ==> x <= abs y`) THEN
+ ASM_SIMP_TAC[REAL_LE_SUP_FINITE; FINITE_IMAGE; IMAGE_EQ_EMPTY;
+ NUMSEG_EMPTY; FINITE_NUMSEG; GSYM NOT_LE; DIMINDEX_GE_1] THEN
+ REWRITE_TAC[EXISTS_IN_IMAGE; IN_NUMSEG] THEN ASM_MESON_TAC[REAL_LE_REFL];
+ ALL_TAC] THEN
+ REWRITE_TAC[NORM_MUL] THEN MATCH_MP_TAC REAL_LTE_TRANS THEN
+ EXISTS_TAC `abs(sup(IMAGE(\<lambda>k. abs(m k)) (1..dimindex(:N)))) * B` THEN
+ SUBGOAL_THEN `0 < sup(IMAGE(\<lambda>k. abs(m k)) (1..dimindex(:N)))`
+ ASSUME_TAC THENL
+ [ASM_SIMP_TAC[REAL_LT_SUP_FINITE; FINITE_IMAGE; IMAGE_EQ_EMPTY;
+ NUMSEG_EMPTY; FINITE_NUMSEG; GSYM NOT_LE; DIMINDEX_GE_1] THEN
+ REWRITE_TAC[EXISTS_IN_IMAGE; GSYM REAL_ABS_NZ; IN_NUMSEG] THEN
+ ASM_MESON_TAC[DIMINDEX_GE_1; LE_REFL];
+ ALL_TAC] THEN
+ ASM_SIMP_TAC[REAL_LT_LMUL_EQ; REAL_ARITH `0 < x ==> 0 < abs x`] THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `sup(IMAGE(\<lambda>k. abs(m k)) (1..dimindex(:N))) * B` THEN
+ ASM_SIMP_TAC[REAL_LE_RMUL_EQ; REAL_ARITH `0 < x ==> abs x <= x`] THEN
+ ASM_SIMP_TAC[REAL_LE_SUP_FINITE; FINITE_IMAGE; IMAGE_EQ_EMPTY;
+ NUMSEG_EMPTY; FINITE_NUMSEG; GSYM NOT_LE; DIMINDEX_GE_1] THEN
+ ASM_SIMP_TAC[EXISTS_IN_IMAGE; REAL_LE_RMUL_EQ] THEN
+ ASM_SIMP_TAC[REAL_SUP_LE_FINITE; FINITE_IMAGE; IMAGE_EQ_EMPTY;
+ NUMSEG_EMPTY; FINITE_NUMSEG; GSYM NOT_LE; DIMINDEX_GE_1] THEN
+ MP_TAC(ISPEC `IMAGE (\<lambda>k. abs (m k)) (1..dimindex(:N))` SUP_FINITE) THEN
+ REWRITE_TAC[FORALL_IN_IMAGE] THEN
+ ASM_SIMP_TAC[FINITE_IMAGE; FINITE_NUMSEG; IMAGE_EQ_EMPTY; NUMSEG_EMPTY;
+ GSYM NOT_LE; DIMINDEX_GE_1] THEN
+ REWRITE_TAC[IN_IMAGE] THEN MESON_TAC[];
+
+ MATCH_MP_TAC(MESON[]
+ `s = t \<and> P z ==> (f has_integral z) s ==> Q
+ ==> ?w. (f has_integral w) t \<and> P w`) THEN
+ SIMP_TAC[GSYM PRODUCT_INV; FINITE_NUMSEG; GSYM REAL_ABS_INV] THEN
+ REWRITE_TAC[REAL_INV_INV] THEN CONJ_TAC THENL
+ [REWRITE_TAC[GSYM IMAGE_o] THEN MATCH_MP_TAC(SET_RULE
+ `(!x. f x = x) ==> IMAGE f s = s`) THEN
+ SIMP_TAC[o_THM; LAMBDA_BETA; CART_EQ] THEN
+ ASM_SIMP_TAC[REAL_MUL_ASSOC; REAL_MUL_RINV; REAL_MUL_LID];
+ REWRITE_TAC[ABS_; _SUB; LIFT_; _CMUL] THEN
+ REWRITE_TAC[GSYM REAL_SUB_LDISTRIB; ETA_AX] THEN
+ REWRITE_TAC[REAL_ABS_MUL; REAL_ABS_ABS] THEN
+ ONCE_REWRITE_TAC[REAL_MUL_SYM] THEN
+ ASM_SIMP_TAC[GSYM REAL_LT_RDIV_EQ] THEN
+ ASM_MESON_TAC[ABS_; _SUB; LIFT_]]]);; *)
+
+lemma has_gmeasure_translation: True .. (*
+ "!s m a. s has_gmeasure m ==> (IMAGE (\<lambda>x:real^N. a + x) s) has_gmeasure m"
+qed REPEAT GEN_TAC THEN
+ MP_TAC(ISPECL [`s:real^N->bool`; `1`; `a:real^N`; `m:real`]
+ HAS_GMEASURE_AFFINITY) THEN
+ REWRITE_TAC[VECTOR_MUL_LID; REAL_ABS_NUM; REAL_POW_ONE; REAL_MUL_LID] THEN
+ REWRITE_TAC[VECTOR_ADD_SYM]);; *)
+
+lemma negligible_translation: True .. (*
+ "!s a. negligible s ==> negligible (IMAGE (\<lambda>x:real^N. a + x) s)"
+qed SIMP_TAC[GSYM HAS_GMEASURE_0; HAS_GMEASURE_TRANSLATION]);; *)
+
+lemma has_gmeasure_translation_eq: True .. (*
+ "!s m. (IMAGE (\<lambda>x:real^N. a + x) s) has_gmeasure m \<longleftrightarrow> s has_gmeasure m"
+qed REPEAT GEN_TAC THEN EQ_TAC THEN REWRITE_TAC[HAS_GMEASURE_TRANSLATION] THEN
+ DISCH_THEN(MP_TAC o SPEC `--a:real^N` o
+ MATCH_MP HAS_GMEASURE_TRANSLATION) THEN
+ MATCH_MP_TAC EQ_IMP THEN AP_THM_TAC THEN AP_TERM_TAC THEN
+ REWRITE_TAC[GSYM IMAGE_o; o_DEF; VECTOR_ARITH `--a + a + b:real^N = b`] THEN
+ SET_TAC[]);; *)
+
+lemma negligible_translation_rev: True .. (*
+ "!s a. negligible (IMAGE (\<lambda>x:real^N. a + x) s) ==> negligible s"
+qed SIMP_TAC[GSYM HAS_GMEASURE_0; HAS_GMEASURE_TRANSLATION_EQ]);; *)
+
+lemma negligible_translation_eq: True .. (*
+ "!s a. negligible (IMAGE (\<lambda>x:real^N. a + x) s) \<longleftrightarrow> negligible s"
+qed SIMP_TAC[GSYM HAS_GMEASURE_0; HAS_GMEASURE_TRANSLATION_EQ]);; *)
+
+lemma gmeasurable_translation: True .. (*
+ "!s. gmeasurable (IMAGE (\<lambda>x. a + x) s) \<longleftrightarrow> gmeasurable s"
+qed REWRITE_TAC[measurable; HAS_GMEASURE_TRANSLATION_EQ]);; *)
+
+lemma measure_translation: True .. (*
+ "!s. gmeasurable s ==> measure(IMAGE (\<lambda>x. a + x) s) = gmeasure s"
+qed REWRITE_TAC[HAS_GMEASURE_MEASURE] THEN REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC MEASURE_UNIQUE THEN
+ ASM_REWRITE_TAC[HAS_GMEASURE_TRANSLATION_EQ]);; *)
+
+lemma has_gmeasure_scaling: True .. (*
+ "!s m c. s has_gmeasure m
+ ==> (IMAGE (\<lambda>x:real^N. c % x) s) has_gmeasure
+ (abs(c) pow dimindex(:N)) * m"
+qed REPEAT GEN_TAC THEN
+ MP_TAC(ISPECL [`s:real^N->bool`; `c:real`; `0:real^N`; `m:real`]
+ HAS_GMEASURE_AFFINITY) THEN
+ REWRITE_TAC[VECTOR_ADD_RID]);; *)
+
+lemma has_gmeasure_scaling_eq: True .. (*
+ "!s m c. ~(c = 0)
+ ==> (IMAGE (\<lambda>x:real^N. c % x) s
+ has_gmeasure (abs(c) pow dimindex(:N)) * m \<longleftrightarrow>
+ s has_gmeasure m)"
+qed REPEAT STRIP_TAC THEN EQ_TAC THEN REWRITE_TAC[HAS_GMEASURE_SCALING] THEN
+ DISCH_THEN(MP_TAC o SPEC `inv(c)` o MATCH_MP HAS_GMEASURE_SCALING) THEN
+ REWRITE_TAC[GSYM IMAGE_o; o_DEF; GSYM REAL_ABS_MUL] THEN
+ REWRITE_TAC[GSYM REAL_POW_MUL; VECTOR_MUL_ASSOC; REAL_MUL_ASSOC] THEN
+ ASM_SIMP_TAC[GSYM REAL_ABS_MUL; REAL_MUL_LINV] THEN
+ REWRITE_TAC[REAL_POW_ONE; REAL_ABS_NUM; REAL_MUL_LID; VECTOR_MUL_LID] THEN
+ MATCH_MP_TAC EQ_IMP THEN AP_THM_TAC THEN AP_TERM_TAC THEN SET_TAC[]);; *)
+
+lemma gmeasurable_scaling: True .. (*
+ "!s c. gmeasurable s ==> gmeasurable (IMAGE (\<lambda>x. c % x) s)"
+qed REWRITE_TAC[measurable] THEN MESON_TAC[HAS_GMEASURE_SCALING]);; *)
+
+lemma gmeasurable_scaling_eq: True .. (*
+ "!s c. ~(c = 0) ==> (measurable (IMAGE (\<lambda>x. c % x) s) \<longleftrightarrow> gmeasurable s)"
+qed REPEAT STRIP_TAC THEN EQ_TAC THEN REWRITE_TAC[MEASURABLE_SCALING] THEN
+ DISCH_THEN(MP_TAC o SPEC `inv c` o MATCH_MP GMEASURABLE_SCALING) THEN
+ REWRITE_TAC[GSYM IMAGE_o; o_DEF; GSYM REAL_ABS_MUL] THEN
+ MATCH_MP_TAC EQ_IMP THEN AP_TERM_TAC THEN
+ ASM_SIMP_TAC[VECTOR_MUL_ASSOC; REAL_MUL_LINV; VECTOR_MUL_LID] THEN
+ SET_TAC[]);; *)
+
+lemma measure_scaling: True .. (*
+ "!s. gmeasurable s
+ ==> measure(IMAGE (\<lambda>x:real^N. c % x) s) =
+ (abs(c) pow dimindex(:N)) * gmeasure s"
+qed REWRITE_TAC[HAS_GMEASURE_MEASURE] THEN REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC MEASURE_UNIQUE THEN ASM_SIMP_TAC[HAS_GMEASURE_SCALING]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Measurability of countable unions and intersections of various kinds. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma has_gmeasure_nested_unions:
+ assumes "\<And>n. gmeasurable(s n)" "\<And>n. gmeasure(s n) \<le> B" "\<And>n. s(n) \<subseteq> s(Suc n)"
+ shows "gmeasurable(\<Union> { s n | n. n \<in> UNIV }) \<and>
+ (\<lambda>n. gmeasure(s n)) ----> gmeasure(\<Union> { s(n) | n. n \<in> UNIV })"
+proof- let ?g = "\<lambda>x. if x \<in> \<Union>{s n |n. n \<in> UNIV} then 1 else (0::real)"
+ have "?g integrable_on UNIV \<and> (\<lambda>k. integral UNIV (\<lambda>x. if x \<in> s k then 1 else 0)) ----> integral UNIV ?g"
+ proof(rule monotone_convergence_increasing)
+ case goal1 show ?case using assms(1) unfolding gmeasurable_integrable by auto
+ case goal2 show ?case using assms(3) by auto
+ have "\<forall>m n. m\<le>n \<longrightarrow> s m \<subseteq> s n" apply(subst transitive_stepwise_le_eq)
+ using assms(3) by auto note * = this[rule_format]
+ have **:"\<And>x e n. \<lbrakk>x \<in> s n; 0 < e\<rbrakk> \<Longrightarrow> \<exists>N. \<forall>n. x \<notin> s n \<longrightarrow> N \<le> n \<longrightarrow> dist 0 1 < e"
+ apply(rule_tac x=n in exI) using * by auto
+ case goal3 show ?case unfolding Lim_sequentially by(auto intro!: **)
+ case goal4 show ?case unfolding bounded_def apply(rule_tac x=0 in exI)
+ apply(rule_tac x=B in exI) unfolding dist_real_def apply safe
+ unfolding measure_integral_univ[OF assms(1),THEN sym]
+ apply(subst abs_of_nonpos) using assms(1,2) by auto
+ qed note conjunctD2[OF this]
+ thus ?thesis unfolding gmeasurable_integrable[THEN sym] measure_integral_univ[OF assms(1)]
+ apply- unfolding measure_integral_univ by auto
+qed
+
+lemmas gmeasurable_nested_unions = has_gmeasure_nested_unions(1)
+
+lemma sums_alt:"f sums s = (\<lambda>n. setsum f {0..n}) ----> s"
+proof- have *:"\<And>n. {0..<Suc n} = {0..n}" by auto
+ show ?thesis unfolding sums_def apply(subst LIMSEQ_Suc_iff[THEN sym]) unfolding * ..
+qed
+
+lemma has_gmeasure_countable_negligible_unions:
+ assumes "\<And>n. gmeasurable(s n)" "\<And>m n. m \<noteq> n \<Longrightarrow> negligible(s m \<inter> s n)"
+ "\<And>n. setsum (\<lambda>k. gmeasure(s k)) {0..n} <= B"
+ shows "gmeasurable(\<Union> { s(n) |n. n \<in> UNIV })" (is ?m)
+ "((\<lambda>n. gmeasure(s n)) sums (gmeasure(\<Union> { s(n) |n. n \<in> UNIV })))" (is ?s)
+proof- have *:"\<And>n. (\<Union> (s ` {0..n})) has_gmeasure (setsum (\<lambda>k. gmeasure(s k)) {0..n})"
+ apply(rule has_gmeasure_negligible_unions_image) using assms by auto
+ have **:"(\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV}) = (\<Union>{s n |n. n \<in> UNIV})" unfolding simple_image by fastsimp
+ have "gmeasurable (\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV}) \<and>
+ (\<lambda>n. gmeasure (\<Union>(s ` {0..n}))) ----> gmeasure (\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV})"
+ apply(rule has_gmeasure_nested_unions) apply(rule gmeasurableI,rule *)
+ unfolding measure_unique[OF *] defer apply(rule Union_mono,rule image_mono) using assms(3) by auto
+ note lem = conjunctD2[OF this,unfolded **]
+ show ?m using lem(1) .
+ show ?s using lem(2) unfolding sums_alt measure_unique[OF *] .
+qed
+
+lemma negligible_countable_unions: True .. (*
+ "!s:num->real^N->bool.
+ (!n. negligible(s n)) ==> negligible(UNIONS {s(n) | n \<in> (:num)})"
+qed REPEAT STRIP_TAC THEN
+ MP_TAC(ISPECL [`s:num->real^N->bool`; `0`]
+ HAS_GMEASURE_COUNTABLE_NEGLIGIBLE_UNIONS) THEN
+ ASM_SIMP_TAC[MEASURE_EQ_0; SUM_0; REAL_LE_REFL; LIFT_NUM] THEN ANTS_TAC THENL
+ [ASM_MESON_TAC[HAS_GMEASURE_0; gmeasurable; INTER_SUBSET; NEGLIGIBLE_SUBSET];
+ ALL_TAC] THEN
+ SIMP_TAC[GSYM GMEASURABLE_MEASURE_EQ_0] THEN
+ STRIP_TAC THEN REWRITE_TAC[GSYM LIFT_EQ] THEN
+ MATCH_MP_TAC SERIES_UNIQUE THEN REWRITE_TAC[LIFT_NUM] THEN
+ MAP_EVERY EXISTS_TAC [`(\<lambda>k. 0):num->real^1`; `from 0`] THEN
+ ASM_REWRITE_TAC[SERIES_0]);; *)
+
+lemma gmeasurable_countable_unions_strong:
+ assumes "\<And>n. gmeasurable(s n)" "\<And>n::nat. gmeasure(\<Union>{s k |k. k \<le> n}) \<le> B"
+ shows "gmeasurable(\<Union>{ s(n) |n. n \<in> UNIV })"
+proof- have *:"\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV} = \<Union>range s" unfolding simple_image by fastsimp
+ show ?thesis unfolding simple_image
+ apply(rule gmeasurable_nested_unions[of "\<lambda>n. \<Union>(s ` {0..n})", THEN conjunct1,unfolded *])
+ proof- fix n::nat show "gmeasurable (\<Union>s ` {0..n})"
+ apply(rule gmeasurable_finite_unions) using assms(1) by auto
+ show "gmeasure (\<Union>s ` {0..n}) \<le> B"
+ using assms(2)[of n] unfolding simple_image[THEN sym] by auto
+ show "\<Union>s ` {0..n} \<subseteq> \<Union>s ` {0..Suc n}" apply(rule Union_mono) by auto
+ qed
+qed
+
+lemma has_gmeasure_countable_negligible_unions_bounded: True .. (*
+ "!s:num->real^N->bool.
+ (!n. gmeasurable(s n)) \<and>
+ (!m n. ~(m = n) ==> negligible(s m \<inter> s n)) \<and>
+ bounded(\<Union>{ s(n) | n \<in> (:num) })
+ ==> gmeasurable(\<Union>{ s(n) | n \<in> (:num) }) \<and>
+ ((\<lambda>n. lift(measure(s n))) sums
+ lift(measure(\<Union>{ s(n) | n \<in> (:num) }))) (from 0)"
+qed REPEAT GEN_TAC THEN STRIP_TAC THEN
+ FIRST_ASSUM(MP_TAC o MATCH_MP BOUNDED_SUBSET_CLOSED_INTERVAL) THEN
+ REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
+ MAP_EVERY X_GEN_TAC [`a:real^N`; `b:real^N`] THEN DISCH_TAC THEN
+ MATCH_MP_TAC HAS_GMEASURE_COUNTABLE_NEGLIGIBLE_UNIONS THEN
+ EXISTS_TAC `measure(interval[a:real^N,b])` THEN
+ ASM_REWRITE_TAC[] THEN X_GEN_TAC `n:num` THEN MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `measure(UNIONS (IMAGE (s:num->real^N->bool) (0..n)))` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC REAL_EQ_IMP_LE THEN CONV_TAC SYM_CONV THEN
+ MATCH_MP_TAC MEASURE_NEGLIGIBLE_UNIONS_IMAGE THEN
+ ASM_SIMP_TAC[FINITE_NUMSEG];
+ MATCH_MP_TAC MEASURE_SUBSET THEN REWRITE_TAC[MEASURABLE_INTERVAL] THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC GMEASURABLE_UNIONS THEN
+ ASM_SIMP_TAC[FINITE_IMAGE; FINITE_NUMSEG; FORALL_IN_IMAGE];
+ ASM SET_TAC[]]]);; *)
+
+lemma gmeasurable_countable_negligible_unions_bounded: True .. (*
+ "!s:num->real^N->bool.
+ (!n. gmeasurable(s n)) \<and>
+ (!m n. ~(m = n) ==> negligible(s m \<inter> s n)) \<and>
+ bounded(\<Union>{ s(n) | n \<in> (:num) })
+ ==> gmeasurable(\<Union>{ s(n) | n \<in> (:num) })"
+qed SIMP_TAC[HAS_GMEASURE_COUNTABLE_NEGLIGIBLE_UNIONS_BOUNDED]);; *)
+
+lemma gmeasurable_countable_unions: True .. (*
+ "!s:num->real^N->bool B.
+ (!n. gmeasurable(s n)) \<and>
+ (!n. sum (0..n) (\<lambda>k. measure(s k)) \<le> B)
+ ==> gmeasurable(\<Union>{ s(n) | n \<in> (:num) })"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC GMEASURABLE_COUNTABLE_UNIONS_STRONG THEN
+ EXISTS_TAC `B:real` THEN ASM_REWRITE_TAC[] THEN
+ X_GEN_TAC `n:num` THEN MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `sum(0..n) (\<lambda>k. measure(s k:real^N->bool))` THEN
+ ASM_REWRITE_TAC[] THEN
+ W(MP_TAC o PART_MATCH (rand o rand) MEASURE_UNIONS_LE_IMAGE o
+ rand o snd) THEN
+ ASM_REWRITE_TAC[FINITE_NUMSEG] THEN
+ ONCE_REWRITE_TAC[GSYM SIMPLE_IMAGE] THEN
+ REWRITE_TAC[IN_NUMSEG; LE_0]);; *)
+
+lemma gmeasurable_countable_inters: True .. (*
+ "!s:num->real^N->bool.
+ (!n. gmeasurable(s n))
+ ==> gmeasurable(INTERS { s(n) | n \<in> (:num) })"
+qed REPEAT STRIP_TAC THEN
+ SUBGOAL_THEN `INTERS { s(n):real^N->bool | n \<in> (:num) } =
+ s 0 DIFF (\<Union>{s 0 DIFF s n | n \<in> (:num)})`
+ SUBST1_TAC THENL
+ [GEN_REWRITE_TAC I [EXTENSION] THEN
+ REWRITE_TAC[IN_INTERS; IN_DIFF; IN_UNIONS] THEN
+ REWRITE_TAC[SIMPLE_IMAGE; FORALL_IN_IMAGE; EXISTS_IN_IMAGE] THEN
+ ASM SET_TAC[];
+ ALL_TAC] THEN
+ MATCH_MP_TAC GMEASURABLE_DIFF THEN ASM_REWRITE_TAC[] THEN
+ MATCH_MP_TAC GMEASURABLE_COUNTABLE_UNIONS_STRONG THEN
+ EXISTS_TAC `measure(s 0:real^N->bool)` THEN
+ ASM_SIMP_TAC[MEASURABLE_DIFF; LE_0] THEN
+ GEN_TAC THEN MATCH_MP_TAC MEASURE_SUBSET THEN
+ ASM_REWRITE_TAC[] THEN CONJ_TAC THENL
+ [ALL_TAC;
+ REWRITE_TAC[SUBSET; FORALL_IN_UNIONS; IN_ELIM_THM; IN_DIFF] THEN
+ MESON_TAC[IN_DIFF]] THEN
+ ONCE_REWRITE_TAC[GSYM IN_NUMSEG_0] THEN
+ ONCE_REWRITE_TAC[SIMPLE_IMAGE] THEN
+ ASM_SIMP_TAC[FORALL_IN_IMAGE; FINITE_IMAGE; FINITE_NUMSEG;
+ GMEASURABLE_DIFF; GMEASURABLE_UNIONS]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* measurability of compact and bounded open sets. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma gmeasurable_compact: True .. (*
+ "!s:real^N->bool. compact s ==> gmeasurable s"
+qed lemma lemma = prove
+ (`!f s:real^N->bool.
+ (!n. FINITE(f n)) \<and>
+ (!n. s \<subseteq> UNIONS(f n)) \<and>
+ (!x. ~(x \<in> s) ==> ?n. ~(x \<in> UNIONS(f n))) \<and>
+ (!n a. a \<in> f(SUC n) ==> ?b. b \<in> f(n) \<and> a \<subseteq> b) \<and>
+ (!n a. a \<in> f(n) ==> gmeasurable a)
+ ==> gmeasurable s"
+qed REPEAT STRIP_TAC THEN
+ SUBGOAL_THEN `!n. UNIONS(f(SUC n):(real^N->bool)->bool) \<subseteq> UNIONS(f n)`
+ ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
+ SUBGOAL_THEN `s = INTERS { UNIONS(f n) | n \<in> (:num) }:real^N->bool`
+ SUBST1_TAC THENL
+ [ONCE_REWRITE_TAC[SIMPLE_IMAGE] THEN
+ MATCH_MP_TAC SUBSET_ANTISYM THEN CONJ_TAC THEN
+ REWRITE_TAC[SUBSET; IN_INTERS; FORALL_IN_IMAGE; IN_UNIV] THEN
+ REWRITE_TAC[IN_IMAGE] THEN ASM SET_TAC[];
+ MATCH_MP_TAC GMEASURABLE_COUNTABLE_INTERS THEN
+ ASM_REWRITE_TAC[] THEN GEN_TAC THEN
+ MATCH_MP_TAC GMEASURABLE_UNIONS THEN
+ ASM_MESON_TAC[]]) in
+ REPEAT STRIP_TAC THEN MATCH_MP_TAC lemma THEN
+ EXISTS_TAC
+ `\n. { k | ?u:real^N. (!i. 1 \<le> i \<and> i \<le> dimindex(:N)
+ ==> integer(u$i)) \<and>
+ k = { x:real^N | !i. 1 \<le> i \<and> i \<le> dimindex(:N)
+ ==> u$i / 2 pow n \<le> x$i \<and>
+ x$i < (u$i + 1) / 2 pow n } \<and>
+ ~(s \<inter> k = {})}` THEN
+ REWRITE_TAC[IN_ELIM_THM] THEN REPEAT CONJ_TAC THENL
+ [X_GEN_TAC `n:num` THEN
+ SIMP_TAC[REAL_LE_LDIV_EQ; REAL_LT_RDIV_EQ; REAL_LT_POW2] THEN
+ SUBGOAL_THEN
+ `?N. !x:real^N i. x \<in> s \<and> 1 \<le> i \<and> i \<le> dimindex(:N)
+ ==> abs(x$i * 2 pow n) < N`
+ STRIP_ASSUME_TAC THENL
+ [FIRST_ASSUM(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN
+ REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN
+ X_GEN_TAC `B:real` THEN STRIP_TAC THEN
+ MP_TAC(SPEC `B * 2 pow n` (MATCH_MP REAL_ARCH REAL_LT_01)) THEN
+ MATCH_MP_TAC MONO_EXISTS THEN REWRITE_TAC[REAL_MUL_RID] THEN
+ X_GEN_TAC `N:num` THEN
+ REWRITE_TAC[REAL_ABS_MUL; REAL_ABS_POW; REAL_ABS_NUM] THEN
+ SIMP_TAC[GSYM REAL_LT_RDIV_EQ; REAL_LT_POW2] THEN
+ ASM_MESON_TAC[COMPONENT_LE_NORM; REAL_LE_TRANS; REAL_LET_TRANS];
+ ALL_TAC] THEN
+ MATCH_MP_TAC FINITE_SUBSET THEN
+ EXISTS_TAC
+ `IMAGE (\<lambda>u. {x | !i. 1 \<le> i \<and> i \<le> dimindex(:N)
+ ==> (u:real^N)$i \<le> (x:real^N)$i * 2 pow n \<and>
+ x$i * 2 pow n < u$i + 1})
+ {u | !i. 1 \<le> i \<and> i \<le> dimindex(:N) ==> integer (u$i) \<and>
+ abs(u$i) \<le> N}` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC FINITE_IMAGE THEN MATCH_MP_TAC FINITE_CART THEN
+ REWRITE_TAC[GSYM REAL_BOUNDS_LE; FINITE_INTSEG];
+ REWRITE_TAC[SUBSET; IN_ELIM_THM; IN_IMAGE] THEN
+ X_GEN_TAC `l:real^N->bool` THEN
+ MATCH_MP_TAC MONO_EXISTS THEN X_GEN_TAC `u:real^N` THEN
+ STRIP_TAC THEN FIRST_X_ASSUM SUBST_ALL_TAC THEN ASM_SIMP_TAC[] THEN
+ X_GEN_TAC `k:num` THEN STRIP_TAC THEN
+ MATCH_MP_TAC REAL_LE_REVERSE_INTEGERS THEN
+ ASM_SIMP_TAC[INTEGER_CLOSED] THEN
+ FIRST_X_ASSUM(MP_TAC o GEN_REWRITE_RULE I [GSYM MEMBER_NOT_EMPTY]) THEN
+ DISCH_THEN(X_CHOOSE_THEN `x:real^N` MP_TAC) THEN
+ REWRITE_TAC[IN_INTER; IN_ELIM_THM] THEN
+ DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC (MP_TAC o SPEC `k:num`)) THEN
+ ASM_REWRITE_TAC[] THEN
+ FIRST_X_ASSUM(MP_TAC o SPECL [`x:real^N`; `k:num`]) THEN
+ ASM_REWRITE_TAC[] THEN REAL_ARITH_TAC];
+ X_GEN_TAC `n:num` THEN REWRITE_TAC[SUBSET; IN_UNIONS; IN_ELIM_THM] THEN
+ X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+ REWRITE_TAC[LEFT_AND_EXISTS_THM] THEN
+ ONCE_REWRITE_TAC[SWAP_EXISTS_THM] THEN
+ EXISTS_TAC `(lambda i. floor(2 pow n * (x:real^N)$i)):real^N` THEN
+ ONCE_REWRITE_TAC[TAUT `(a \<and> b \<and> c) \<and> d \<longleftrightarrow> b \<and> a \<and> c \<and> d`] THEN
+ REWRITE_TAC[UNWIND_THM2] THEN SIMP_TAC[LAMBDA_BETA; FLOOR] THEN
+ REWRITE_TAC[GSYM MEMBER_NOT_EMPTY; IN_INTER] THEN
+ REWRITE_TAC[LEFT_AND_EXISTS_THM] THEN EXISTS_TAC `x:real^N` THEN
+ ASM_REWRITE_TAC[IN_ELIM_THM] THEN
+ SIMP_TAC[REAL_LE_LDIV_EQ; REAL_LT_RDIV_EQ; REAL_LT_POW2] THEN
+ REWRITE_TAC[REAL_MUL_SYM; FLOOR];
+ X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+ FIRST_ASSUM(MP_TAC o MATCH_MP COMPACT_IMP_CLOSED) THEN
+ REWRITE_TAC[closed; open_def] THEN
+ DISCH_THEN(MP_TAC o SPEC `x:real^N`) THEN
+ ASM_REWRITE_TAC[IN_DIFF; IN_UNIV] THEN
+ DISCH_THEN(X_CHOOSE_THEN `e:real` STRIP_ASSUME_TAC) THEN
+ MP_TAC(SPECL [`inv(2)`; `e / (dimindex(:N))`] REAL_ARCH_POW_INV) THEN
+ ASM_SIMP_TAC[REAL_LT_DIV; REAL_LT_RDIV_EQ; REAL_OF_NUM_LT;
+ DIMINDEX_GE_1; ARITH_RULE `0 < x \<longleftrightarrow> 1 \<le> x`] THEN
+ CONV_TAC REAL_RAT_REDUCE_CONV THEN
+ MATCH_MP_TAC MONO_EXISTS THEN X_GEN_TAC `n:num` THEN DISCH_TAC THEN
+ REWRITE_TAC[IN_UNIONS; IN_ELIM_THM] THEN
+ REWRITE_TAC[LEFT_AND_EXISTS_THM] THEN
+ ONCE_REWRITE_TAC[SWAP_EXISTS_THM] THEN
+ ONCE_REWRITE_TAC[TAUT `(a \<and> b \<and> c) \<and> d \<longleftrightarrow> b \<and> a \<and> c \<and> d`] THEN
+ REWRITE_TAC[UNWIND_THM2] THEN REWRITE_TAC[NOT_EXISTS_THM] THEN
+ X_GEN_TAC `u:real^N` THEN REWRITE_TAC[GSYM MEMBER_NOT_EMPTY] THEN
+ REWRITE_TAC[IN_INTER; IN_ELIM_THM] THEN
+ DISCH_THEN(CONJUNCTS_THEN2 MP_TAC ASSUME_TAC o CONJUNCT2) THEN
+ DISCH_THEN(X_CHOOSE_THEN `y:real^N`
+ (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN
+ REWRITE_TAC[] THEN FIRST_X_ASSUM MATCH_MP_TAC THEN
+ FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REAL_ARITH
+ `d < e ==> x \<le> d ==> x < e`)) THEN
+ REWRITE_TAC[dist] THEN
+ W(MP_TAC o PART_MATCH lhand NORM_LE_L1 o lhand o snd) THEN
+ MATCH_MP_TAC(REAL_ARITH `a \<le> b ==> x \<le> a ==> x \<le> b`) THEN
+ GEN_REWRITE_TAC (funpow 3 RAND_CONV) [GSYM CARD_NUMSEG_1] THEN
+ ONCE_REWRITE_TAC[REAL_MUL_SYM] THEN MATCH_MP_TAC SUM_BOUND THEN
+ SIMP_TAC[FINITE_NUMSEG; IN_NUMSEG; VECTOR_SUB_COMPONENT] THEN
+ X_GEN_TAC `k:num` THEN STRIP_TAC THEN
+ REPEAT(FIRST_X_ASSUM(MP_TAC o SPEC `k:num`)) THEN
+ ASM_REWRITE_TAC[real_div; REAL_ADD_RDISTRIB] THEN
+ REWRITE_TAC[REAL_MUL_LID; GSYM REAL_POW_INV] THEN REAL_ARITH_TAC;
+ MAP_EVERY X_GEN_TAC [`n:num`; `a:real^N->bool`] THEN
+ DISCH_THEN(X_CHOOSE_THEN `u:real^N`
+ (CONJUNCTS_THEN2 ASSUME_TAC MP_TAC)) THEN
+ DISCH_THEN(CONJUNCTS_THEN2 (ASSUME_TAC o SYM) ASSUME_TAC) THEN
+ REWRITE_TAC[LEFT_AND_EXISTS_THM] THEN
+ ONCE_REWRITE_TAC[SWAP_EXISTS_THM] THEN
+ ONCE_REWRITE_TAC[TAUT `(a \<and> b \<and> c) \<and> d \<longleftrightarrow> b \<and> a \<and> c \<and> d`] THEN
+ REWRITE_TAC[UNWIND_THM2] THEN
+ EXISTS_TAC `(lambda i. floor((u:real^N)$i / 2)):real^N` THEN
+ ASM_SIMP_TAC[VECTOR_MUL_COMPONENT; LAMBDA_BETA; FLOOR] THEN
+ MATCH_MP_TAC(SET_RULE `~(s \<inter> a = {}) \<and> a \<subseteq> b
+ ==> ~(s \<inter> b = {}) \<and> a \<subseteq> b`) THEN
+ ASM_REWRITE_TAC[] THEN EXPAND_TAC "a" THEN REWRITE_TAC[SUBSET] THEN
+ X_GEN_TAC `x:real^N` THEN REWRITE_TAC[IN_ELIM_THM] THEN
+ MATCH_MP_TAC MONO_FORALL THEN X_GEN_TAC `k:num` THEN
+ DISCH_THEN(fun th -> STRIP_TAC THEN MP_TAC th) THEN ASM_REWRITE_TAC[] THEN
+ REWRITE_TAC[real_pow; real_div; REAL_INV_MUL; REAL_MUL_ASSOC] THEN
+ REWRITE_TAC[GSYM real_div] THEN
+ SIMP_TAC[REAL_LE_LDIV_EQ; REAL_LT_RDIV_EQ; REAL_LT_POW2] THEN
+ MP_TAC(SPEC `(u:real^N)$k / 2` FLOOR) THEN
+ REWRITE_TAC[REAL_ARITH `u / 2 < floor(u / 2) + 1 \<longleftrightarrow>
+ u < 2 * floor(u / 2) + 2`] THEN
+ ASM_SIMP_TAC[REAL_LT_INTEGERS; INTEGER_CLOSED; FLOOR_FRAC] THEN
+ REAL_ARITH_TAC;
+ REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
+ MAP_EVERY X_GEN_TAC [`n:num`; `a:real^N->bool`; `u:real^N`] THEN
+ DISCH_THEN(SUBST1_TAC o CONJUNCT1 o CONJUNCT2) THEN
+ ONCE_REWRITE_TAC[MEASURABLE_INNER_OUTER] THEN
+ GEN_TAC THEN DISCH_TAC THEN
+ EXISTS_TAC `interval(inv(2 pow n) % u:real^N,
+ inv(2 pow n) % (u + 1))` THEN
+ EXISTS_TAC `interval[inv(2 pow n) % u:real^N,
+ inv(2 pow n) % (u + 1)]` THEN
+ REWRITE_TAC[MEASURABLE_INTERVAL; MEASURE_INTERVAL] THEN
+ ASM_REWRITE_TAC[REAL_SUB_REFL; REAL_ABS_0] THEN
+ REWRITE_TAC[SUBSET; IN_INTERVAL; IN_ELIM_THM] THEN
+ CONJ_TAC THEN X_GEN_TAC `y:real^N` THEN
+ MATCH_MP_TAC MONO_FORALL THEN X_GEN_TAC `k:num` THEN
+ DISCH_THEN(fun th -> STRIP_TAC THEN MP_TAC th) THEN
+ ASM_SIMP_TAC[VECTOR_MUL_COMPONENT; VECTOR_ADD_COMPONENT;
+ VEC_COMPONENT] THEN
+ REAL_ARITH_TAC]);; *)
+
+lemma gmeasurable_open: True .. (*
+ "!s:real^N->bool. bounded s \<and> open s ==> gmeasurable s"
+qed REPEAT STRIP_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o MATCH_MP BOUNDED_SUBSET_CLOSED_INTERVAL) THEN
+ REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
+ MAP_EVERY X_GEN_TAC [`a:real^N`; `b:real^N`] THEN DISCH_TAC THEN
+ FIRST_ASSUM(SUBST1_TAC o MATCH_MP (SET_RULE
+ `s \<subseteq> t ==> s = t DIFF (t DIFF s)`)) THEN
+ MATCH_MP_TAC GMEASURABLE_DIFF THEN
+ REWRITE_TAC[MEASURABLE_INTERVAL] THEN
+ MATCH_MP_TAC GMEASURABLE_COMPACT THEN
+ SIMP_TAC[COMPACT_EQ_BOUNDED_CLOSED; BOUNDED_DIFF; BOUNDED_INTERVAL] THEN
+ MATCH_MP_TAC CLOSED_DIFF THEN ASM_REWRITE_TAC[CLOSED_INTERVAL]);; *)
+
+lemma gmeasurable_closure: True .. (*
+ "!s. bounded s ==> gmeasurable(closure s)"
+qed SIMP_TAC[MEASURABLE_COMPACT; COMPACT_EQ_BOUNDED_CLOSED; CLOSED_CLOSURE;
+ BOUNDED_CLOSURE]);; *)
+
+lemma gmeasurable_interior: True .. (*
+ "!s. bounded s ==> gmeasurable(interior s)"
+qed SIMP_TAC[MEASURABLE_OPEN; OPEN_INTERIOR; BOUNDED_INTERIOR]);; *)
+
+lemma gmeasurable_frontier: True .. (*
+ "!s:real^N->bool. bounded s ==> gmeasurable(frontier s)"
+qed REPEAT STRIP_TAC THEN REWRITE_TAC[frontier] THEN
+ MATCH_MP_TAC GMEASURABLE_DIFF THEN
+ ASM_SIMP_TAC[MEASURABLE_CLOSURE; GMEASURABLE_INTERIOR] THEN
+ MATCH_MP_TAC SUBSET_TRANS THEN EXISTS_TAC `s:real^N->bool` THEN
+ REWRITE_TAC[INTERIOR_SUBSET; CLOSURE_SUBSET]);; *)
+
+lemma measure_frontier: True .. (*
+ "!s:real^N->bool.
+ bounded s
+ ==> measure(frontier s) = measure(closure s) - measure(interior s)"
+qed REPEAT STRIP_TAC THEN REWRITE_TAC[frontier] THEN
+ MATCH_MP_TAC MEASURE_DIFF_SUBSET THEN
+ ASM_SIMP_TAC[MEASURABLE_CLOSURE; GMEASURABLE_INTERIOR] THEN
+ MATCH_MP_TAC SUBSET_TRANS THEN EXISTS_TAC `s:real^N->bool` THEN
+ REWRITE_TAC[INTERIOR_SUBSET; CLOSURE_SUBSET]);; *)
+
+lemma gmeasurable_jordan: True .. (*
+ "!s:real^N->bool. bounded s \<and> negligible(frontier s) ==> gmeasurable s"
+qed REPEAT STRIP_TAC THEN ONCE_REWRITE_TAC[MEASURABLE_INNER_OUTER] THEN
+ GEN_TAC THEN DISCH_TAC THEN
+ EXISTS_TAC `interior(s):real^N->bool` THEN
+ EXISTS_TAC `closure(s):real^N->bool` THEN
+ ASM_SIMP_TAC[MEASURABLE_INTERIOR; GMEASURABLE_CLOSURE] THEN
+ REWRITE_TAC[INTERIOR_SUBSET; CLOSURE_SUBSET] THEN
+ ONCE_REWRITE_TAC[REAL_ABS_SUB] THEN
+ ASM_SIMP_TAC[GSYM MEASURE_FRONTIER; REAL_ABS_NUM; MEASURE_EQ_0]);; *)
+
+lemma has_gmeasure_elementary: True .. (*
+ "!d s. d division_of s ==> s has_gmeasure (sum d content)"
+qed REPEAT STRIP_TAC THEN REWRITE_TAC[has_gmeasure] THEN
+ FIRST_ASSUM(ASSUME_TAC o MATCH_MP DIVISION_OF_FINITE) THEN
+ ASM_SIMP_TAC[LIFT_SUM] THEN
+ MATCH_MP_TAC HAS_INTEGRAL_COMBINE_DIVISION THEN
+ ASM_REWRITE_TAC[o_THM] THEN REWRITE_TAC[GSYM has_gmeasure] THEN
+ ASM_MESON_TAC[HAS_GMEASURE_INTERVAL; division_of]);; *)
+
+lemma gmeasurable_elementary: True .. (*
+ "!d s. d division_of s ==> gmeasurable s"
+qed REWRITE_TAC[measurable] THEN MESON_TAC[HAS_GMEASURE_ELEMENTARY]);; *)
+
+lemma measure_elementary: True .. (*
+ "!d s. d division_of s ==> gmeasure s = sum d content"
+qed MESON_TAC[HAS_GMEASURE_ELEMENTARY; MEASURE_UNIQUE]);; *)
+
+lemma gmeasurable_inter_interval: True .. (*
+ "!s a b:real^N. gmeasurable s ==> gmeasurable (s \<inter> {a..b})"
+qed SIMP_TAC[MEASURABLE_INTER; GMEASURABLE_INTERVAL]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* A nice lemma for negligibility proofs. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma STARLIKE_NEGLIGIBLE_BOUNDED_MEASURABLE: True .. (*
+ "!s. gmeasurable s \<and> bounded s \<and>
+ (!c x:real^N. 0 \<le> c \<and> x \<in> s \<and> (c % x) \<in> s ==> c = 1)
+ ==> negligible s"
+qed REPEAT STRIP_TAC THEN
+ SUBGOAL_THEN `~(0 < measure(s:real^N->bool))`
+ (fun th -> ASM_MESON_TAC[th; GMEASURABLE_MEASURE_POS_LT]) THEN
+ DISCH_TAC THEN
+ MP_TAC(SPEC `(0:real^N) INSERT s`
+ BOUNDED_SUBSET_CLOSED_INTERVAL_SYMMETRIC) THEN
+ ASM_SIMP_TAC[BOUNDED_INSERT; COMPACT_IMP_BOUNDED; NOT_EXISTS_THM] THEN
+ X_GEN_TAC `a:real^N` THEN REWRITE_TAC[INSERT_SUBSET] THEN STRIP_TAC THEN
+ SUBGOAL_THEN
+ `?N. EVEN N \<and> 0 < N \<and>
+ measure(interval[--a:real^N,a])
+ < (N * measure(s:real^N->bool)) / 4 pow dimindex (:N)`
+ STRIP_ASSUME_TAC THENL
+ [FIRST_ASSUM(MP_TAC o SPEC
+ `measure(interval[--a:real^N,a]) * 4 pow (dimindex(:N))` o
+ MATCH_MP REAL_ARCH) THEN
+ SIMP_TAC[REAL_LT_RDIV_EQ; REAL_POW_LT; REAL_OF_NUM_LT; ARITH] THEN
+ SIMP_TAC[GSYM REAL_LT_LDIV_EQ; ASSUME `0 < measure(s:real^N->bool)`] THEN
+ DISCH_THEN(X_CHOOSE_THEN `N:num` STRIP_ASSUME_TAC) THEN
+ EXISTS_TAC `2 * (N DIV 2 + 1)` THEN REWRITE_TAC[EVEN_MULT; ARITH] THEN
+ CONJ_TAC THENL [ARITH_TAC; ALL_TAC] THEN
+ FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REAL_ARITH
+ `x < a ==> a \<le> b ==> x < b`)) THEN
+ REWRITE_TAC[REAL_OF_NUM_LE] THEN ARITH_TAC;
+ ALL_TAC] THEN
+ MP_TAC(ISPECL [`\<Union>(IMAGE (\<lambda>m. IMAGE (\<lambda>x:real^N. (m / N) % x) s)
+ (1..N))`;
+ `interval[--a:real^N,a]`] MEASURE_SUBSET) THEN
+ MP_TAC(ISPECL [`measure:(real^N->bool)->real`;
+ `IMAGE (\<lambda>m. IMAGE (\<lambda>x:real^N. (m / N) % x) s) (1..N)`]
+ HAS_GMEASURE_DISJOINT_UNIONS) THEN
+ SIMP_TAC[FINITE_IMAGE; FINITE_NUMSEG; IMP_CONJ] THEN
+ REWRITE_TAC[RIGHT_FORALL_IMP_THM; FORALL_IN_IMAGE] THEN ANTS_TAC THENL
+ [REPEAT STRIP_TAC THEN REWRITE_TAC[GSYM HAS_GMEASURE_MEASURE] THEN
+ MATCH_MP_TAC GMEASURABLE_SCALING THEN ASM_REWRITE_TAC[];
+ ALL_TAC] THEN
+ REWRITE_TAC[RIGHT_IMP_FORALL_THM; IMP_IMP] THEN
+ ONCE_REWRITE_TAC[TAUT `(a \<and> b) \<and> ~c ==> d \<longleftrightarrow> a \<and> b \<and> ~d ==> c`] THEN
+ SUBGOAL_THEN
+ `!m n. m \<in> 1..N \<and> n \<in> 1..N \<and>
+ ~(DISJOINT (IMAGE (\<lambda>x:real^N. m / N % x) s)
+ (IMAGE (\<lambda>x. n / N % x) s))
+ ==> m = n`
+ ASSUME_TAC THENL
+ [MAP_EVERY X_GEN_TAC [`m:num`; `n:num`] THEN
+ REPEAT(DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC MP_TAC)) THEN
+ REWRITE_TAC[DISJOINT; GSYM MEMBER_NOT_EMPTY] THEN
+ REWRITE_TAC[EXISTS_IN_IMAGE; IN_INTER] THEN
+ DISCH_THEN(X_CHOOSE_THEN `x:real^N`
+ (CONJUNCTS_THEN2 ASSUME_TAC MP_TAC)) THEN
+ REWRITE_TAC[IN_IMAGE] THEN
+ DISCH_THEN(X_CHOOSE_THEN `y:real^N`
+ (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN
+ DISCH_THEN(MP_TAC o AP_TERM `(%) (N / m) :real^N->real^N`) THEN
+ SUBGOAL_THEN `~(N = 0) \<and> ~(m = 0)` STRIP_ASSUME_TAC THENL
+ [REWRITE_TAC[REAL_OF_NUM_EQ] THEN
+ REPEAT(FIRST_X_ASSUM(MP_TAC o GEN_REWRITE_RULE I [IN_NUMSEG])) THEN
+ ARITH_TAC;
+ ALL_TAC] THEN
+ FIRST_X_ASSUM(ASSUME_TAC o GEN_REWRITE_RULE (BINDER_CONV o BINDER_CONV)
+ [GSYM CONTRAPOS_THM]) THEN
+ ASM_SIMP_TAC[VECTOR_MUL_ASSOC; REAL_FIELD
+ `~(x = 0) \<and> ~(y = 0) ==> x / y * y / x = 1`] THEN
+ ASM_SIMP_TAC[REAL_FIELD
+ `~(x = 0) \<and> ~(y = 0) ==> x / y * z / x = z / y`] THEN
+ REWRITE_TAC[VECTOR_MUL_LID] THEN DISCH_THEN SUBST_ALL_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPECL [`n / m`; `y:real^N`]) THEN
+ ASM_SIMP_TAC[REAL_LE_DIV; REAL_POS; REAL_FIELD
+ `~(y = 0) ==> (x / y = 1 \<longleftrightarrow> x = y)`] THEN
+ REWRITE_TAC[REAL_OF_NUM_EQ; EQ_SYM_EQ];
+ ALL_TAC] THEN
+ ANTS_TAC THENL [ASM_MESON_TAC[]; DISCH_TAC] THEN
+ REWRITE_TAC[NOT_IMP] THEN REPEAT CONJ_TAC THENL
+ [REWRITE_TAC[measurable] THEN ASM_MESON_TAC[];
+ REWRITE_TAC[MEASURABLE_INTERVAL];
+ REWRITE_TAC[UNIONS_SUBSET; FORALL_IN_IMAGE] THEN
+ REWRITE_TAC[SUBSET; FORALL_IN_IMAGE] THEN
+ X_GEN_TAC `n:num` THEN DISCH_TAC THEN X_GEN_TAC `x:real^N` THEN
+ DISCH_TAC THEN
+ MP_TAC(ISPECL [`--a:real^N`; `a:real^N`] CONVEX_INTERVAL) THEN
+ DISCH_THEN(MP_TAC o REWRITE_RULE[CONVEX_ALT] o CONJUNCT1) THEN
+ DISCH_THEN(MP_TAC o SPECL [`0:real^N`; `x:real^N`; `n / N`]) THEN
+ ASM_REWRITE_TAC[VECTOR_MUL_RZERO; VECTOR_ADD_LID] THEN
+ DISCH_THEN MATCH_MP_TAC THEN SIMP_TAC[REAL_LE_DIV; REAL_POS] THEN
+ CONJ_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
+ FIRST_X_ASSUM(MP_TAC o GEN_REWRITE_RULE I [IN_NUMSEG]) THEN
+ DISCH_THEN(MP_TAC o MATCH_MP (ARITH_RULE
+ `1 \<le> n \<and> n \<le> N ==> 0 < N \<and> n \<le> N`)) THEN
+ SIMP_TAC[GSYM REAL_OF_NUM_LE; GSYM REAL_OF_NUM_LT; REAL_LE_LDIV_EQ] THEN
+ SIMP_TAC[REAL_MUL_LID];
+ ALL_TAC] THEN
+ FIRST_X_ASSUM(SUBST1_TAC o MATCH_MP MEASURE_UNIQUE) THEN
+ ASM_SIMP_TAC[MEASURE_SCALING; REAL_NOT_LE] THEN
+ FIRST_X_ASSUM(K ALL_TAC o SPEC `0`) THEN
+ MATCH_MP_TAC REAL_LTE_TRANS THEN EXISTS_TAC
+ `sum (1..N) (measure o (\<lambda>m. IMAGE (\<lambda>x:real^N. m / N % x) s))` THEN
+ CONJ_TAC THENL
+ [ALL_TAC;
+ MATCH_MP_TAC REAL_EQ_IMP_LE THEN CONV_TAC SYM_CONV THEN
+ MATCH_MP_TAC SUM_IMAGE THEN REWRITE_TAC[] THEN
+ REPEAT STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN
+ ASM_REWRITE_TAC[SET_RULE `DISJOINT s s \<longleftrightarrow> s = {}`; IMAGE_EQ_EMPTY] THEN
+ DISCH_THEN SUBST_ALL_TAC THEN
+ ASM_MESON_TAC[REAL_LT_REFL; MEASURE_EMPTY]] THEN
+ FIRST_X_ASSUM(K ALL_TAC o SPEC `0`) THEN
+ ASM_SIMP_TAC[o_DEF; MEASURE_SCALING; SUM_RMUL] THEN
+ FIRST_ASSUM(MATCH_MP_TAC o MATCH_MP (REAL_ARITH
+ `x < a ==> a \<le> b ==> x < b`)) THEN
+ ASM_SIMP_TAC[REAL_LE_LDIV_EQ; REAL_POW_LT; REAL_OF_NUM_LT; ARITH] THEN
+ ONCE_REWRITE_TAC[REAL_ARITH `(a * b) * c:real = (a * c) * b`] THEN
+ ASM_SIMP_TAC[REAL_LE_RMUL_EQ] THEN REWRITE_TAC[GSYM SUM_RMUL] THEN
+ REWRITE_TAC[GSYM REAL_POW_MUL] THEN
+ REWRITE_TAC[REAL_ABS_DIV; REAL_ABS_NUM] THEN
+ FIRST_X_ASSUM(X_CHOOSE_THEN `M:num` SUBST_ALL_TAC o
+ GEN_REWRITE_RULE I [EVEN_EXISTS]) THEN
+ REWRITE_TAC[GSYM REAL_OF_NUM_MUL] THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[GSYM REAL_OF_NUM_MUL]) THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[REAL_ARITH `0 < 2 * x \<longleftrightarrow> 0 < x`]) THEN
+ ASM_SIMP_TAC[REAL_FIELD `0 < y ==> x / (2 * y) * 4 = x * 2 / y`] THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `sum(M..(2*M)) (\<lambda>i. (i * 2 / M) pow dimindex (:N))` THEN
+ CONJ_TAC THENL
+ [ALL_TAC;
+ MATCH_MP_TAC SUM_SUBSET_SIMPLE THEN
+ SIMP_TAC[REAL_POW_LE; REAL_LE_MUL; REAL_LE_DIV; REAL_POS] THEN
+ REWRITE_TAC[IN_NUMSEG; FINITE_NUMSEG; SUBSET] THEN
+ FIRST_X_ASSUM(MP_TAC o GEN_REWRITE_RULE I [REAL_OF_NUM_LT]) THEN
+ ARITH_TAC] THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `sum(M..(2*M)) (\<lambda>i. 2)` THEN CONJ_TAC THENL
+ [REWRITE_TAC[SUM_CONST_NUMSEG] THEN
+ REWRITE_TAC[ARITH_RULE `(2 * M + 1) - M = M + 1`] THEN
+ REWRITE_TAC[GSYM REAL_OF_NUM_ADD] THEN REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ MATCH_MP_TAC SUM_LE THEN REWRITE_TAC[FINITE_NUMSEG; IN_NUMSEG] THEN
+ X_GEN_TAC `n:num` THEN STRIP_TAC THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `2 pow (dimindex(:N))` THEN
+ CONJ_TAC THENL
+ [GEN_REWRITE_TAC LAND_CONV [GSYM REAL_POW_1] THEN
+ MATCH_MP_TAC REAL_POW_MONO THEN REWRITE_TAC[DIMINDEX_GE_1] THEN
+ ARITH_TAC;
+ ALL_TAC] THEN
+ MATCH_MP_TAC REAL_POW_LE2 THEN
+ REWRITE_TAC[REAL_POS; ARITH; real_div; REAL_MUL_ASSOC] THEN
+ ASM_SIMP_TAC[GSYM real_div; REAL_LE_RDIV_EQ] THEN
+ REWRITE_TAC[REAL_OF_NUM_MUL; REAL_OF_NUM_LE] THEN
+ UNDISCH_TAC `M:num \<le> n` THEN ARITH_TAC);; *)
+
+lemma STARLIKE_NEGLIGIBLE_LEMMA: True .. (*
+ "!s. compact s \<and>
+ (!c x:real^N. 0 \<le> c \<and> x \<in> s \<and> (c % x) \<in> s ==> c = 1)
+ ==> negligible s"
+qed REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC STARLIKE_NEGLIGIBLE_BOUNDED_MEASURABLE THEN
+ ASM_MESON_TAC[MEASURABLE_COMPACT; COMPACT_IMP_BOUNDED]);; *)
+
+lemma STARLIKE_NEGLIGIBLE: True .. (*
+ "!s a. closed s \<and>
+ (!c x:real^N. 0 \<le> c \<and> (a + x) \<in> s \<and> (a + c % x) \<in> s ==> c = 1)
+ ==> negligible s"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC NEGLIGIBLE_TRANSLATION_REV THEN
+ EXISTS_TAC `--a:real^N` THEN ONCE_REWRITE_TAC[NEGLIGIBLE_ON_INTERVALS] THEN
+ MAP_EVERY X_GEN_TAC [`u:real^N`; `v:real^N`] THEN
+ MATCH_MP_TAC STARLIKE_NEGLIGIBLE_LEMMA THEN CONJ_TAC THENL
+ [MATCH_MP_TAC CLOSED_INTER_COMPACT THEN REWRITE_TAC[COMPACT_INTERVAL] THEN
+ ASM_SIMP_TAC[CLOSED_TRANSLATION];
+ REWRITE_TAC[IN_IMAGE; IN_INTER] THEN
+ ONCE_REWRITE_TAC[VECTOR_ARITH `x:real^N = --a + y \<longleftrightarrow> y = a + x`] THEN
+ REWRITE_TAC[UNWIND_THM2] THEN ASM MESON_TAC[]]);; *)
+
+lemma STARLIKE_NEGLIGIBLE_STRONG: True .. (*
+ "!s a. closed s \<and>
+ (!c x:real^N. 0 \<le> c \<and> c < 1 \<and> (a + x) \<in> s
+ ==> ~((a + c % x) \<in> s))
+ ==> negligible s"
+qed REPEAT GEN_TAC THEN STRIP_TAC THEN MATCH_MP_TAC STARLIKE_NEGLIGIBLE THEN
+ EXISTS_TAC `a:real^N` THEN ASM_REWRITE_TAC[] THEN
+ MAP_EVERY X_GEN_TAC [`c:real`; `x:real^N`] THEN STRIP_TAC THEN
+ MATCH_MP_TAC(REAL_ARITH `~(x < y) \<and> ~(y < x) ==> x = y`) THEN
+ STRIP_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN DISCH_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPECL [`inv c`; `c % x:real^N`]) THEN
+ ASM_REWRITE_TAC[REAL_LE_INV_EQ; VECTOR_MUL_ASSOC] THEN
+ ASM_SIMP_TAC[REAL_MUL_LINV; REAL_ARITH `1 < c ==> ~(c = 0)`] THEN
+ ASM_REWRITE_TAC[VECTOR_MUL_LID] THEN
+ GEN_REWRITE_TAC RAND_CONV [GSYM REAL_INV_1] THEN
+ MATCH_MP_TAC REAL_LT_INV2 THEN ASM_REWRITE_TAC[REAL_LT_01]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* In particular. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma NEGLIGIBLE_HYPERPLANE: True .. (*
+ "!a b. ~(a = 0 \<and> b = 0) ==> negligible {x:real^N | a dot x = b}"
+qed REPEAT GEN_TAC THEN ASM_CASES_TAC `a:real^N = 0` THEN
+ ASM_SIMP_TAC[DOT_LZERO; SET_RULE `{x | F} = {}`; NEGLIGIBLE_EMPTY] THEN
+ MATCH_MP_TAC STARLIKE_NEGLIGIBLE THEN
+ SUBGOAL_THEN `?x:real^N. ~(a dot x = b)` MP_TAC THENL
+ [MATCH_MP_TAC(MESON[] `!a:real^N. P a \/ P(--a) ==> ?x. P x`) THEN
+ EXISTS_TAC `a:real^N` THEN REWRITE_TAC[DOT_RNEG] THEN
+ MATCH_MP_TAC(REAL_ARITH `~(a = 0) ==> ~(a = b) \/ ~(--a = b)`) THEN
+ ASM_REWRITE_TAC[DOT_EQ_0];
+ ALL_TAC] THEN
+ MATCH_MP_TAC MONO_EXISTS THEN X_GEN_TAC `c:real^N` THEN DISCH_TAC THEN
+ REWRITE_TAC[CLOSED_HYPERPLANE; IN_ELIM_THM; DOT_RADD; DOT_RMUL] THEN
+ MAP_EVERY X_GEN_TAC [`t:real`; `y:real^N`] THEN
+ DISCH_THEN(MP_TAC o MATCH_MP (REAL_ARITH
+ `0 \<le> t \<and> ac + ay = b \<and> ac + t * ay = b
+ ==> ((ay = 0 ==> ac = b) \<and> (t - 1) * ay = 0)`)) THEN
+ ASM_SIMP_TAC[REAL_ENTIRE; REAL_SUB_0] THEN CONV_TAC TAUT);; *)
+
+lemma NEGLIGIBLE_LOWDIM: True .. (*
+ "!s:real^N->bool. dim(s) < dimindex(:N) ==> negligible s"
+qed GEN_TAC THEN DISCH_THEN(MP_TAC o MATCH_MP LOWDIM_SUBSET_HYPERPLANE) THEN
+ DISCH_THEN(X_CHOOSE_THEN `a:real^N` STRIP_ASSUME_TAC) THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `span(s):real^N->bool` THEN REWRITE_TAC[SPAN_INC] THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `{x:real^N | a dot x = 0}` THEN
+ ASM_SIMP_TAC[NEGLIGIBLE_HYPERPLANE]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Measurability of bounded convex sets. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma NEGLIGIBLE_CONVEX_FRONTIER: True .. (*
+ "!s:real^N->bool. convex s ==> negligible(frontier s)"
+qed SUBGOAL_THEN
+ `!s:real^N->bool. convex s \<and> (0) \<in> s ==> negligible(frontier s)`
+ ASSUME_TAC THENL
+ [ALL_TAC;
+ X_GEN_TAC `s:real^N->bool` THEN DISCH_TAC THEN
+ ASM_CASES_TAC `s:real^N->bool = {}` THEN
+ ASM_REWRITE_TAC[FRONTIER_EMPTY; NEGLIGIBLE_EMPTY] THEN
+ FIRST_X_ASSUM(MP_TAC o GEN_REWRITE_RULE I [GSYM MEMBER_NOT_EMPTY]) THEN
+ DISCH_THEN(X_CHOOSE_TAC `a:real^N`) THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `IMAGE (\<lambda>x:real^N. --a + x) s`) THEN
+ ASM_SIMP_TAC[CONVEX_TRANSLATION; IN_IMAGE] THEN
+ ASM_REWRITE_TAC[UNWIND_THM2;
+ VECTOR_ARITH `0:real^N = --a + x \<longleftrightarrow> x = a`] THEN
+ REWRITE_TAC[FRONTIER_TRANSLATION; NEGLIGIBLE_TRANSLATION_EQ]] THEN
+ REPEAT STRIP_TAC THEN MP_TAC(ISPEC `s:real^N->bool` DIM_SUBSET_UNIV) THEN
+ REWRITE_TAC[ARITH_RULE `d:num \<le> e \<longleftrightarrow> d < e \/ d = e`] THEN STRIP_TAC THENL
+ [MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `closure s:real^N->bool` THEN
+ REWRITE_TAC[frontier; SUBSET_DIFF] THEN
+ MATCH_MP_TAC NEGLIGIBLE_LOWDIM THEN ASM_REWRITE_TAC[DIM_CLOSURE];
+ ALL_TAC] THEN
+ SUBGOAL_THEN `?a:real^N. a \<in> interior s` CHOOSE_TAC THENL
+ [X_CHOOSE_THEN `b:real^N->bool` STRIP_ASSUME_TAC
+ (ISPEC `s:real^N->bool` BASIS_EXISTS) THEN
+ FIRST_X_ASSUM SUBST_ALL_TAC THEN
+ MP_TAC(ISPEC `b:real^N->bool` INTERIOR_SIMPLEX_NONEMPTY) THEN
+ ASM_REWRITE_TAC[] THEN
+ MATCH_MP_TAC MONO_EXISTS THEN REWRITE_TAC[GSYM SUBSET] THEN
+ MATCH_MP_TAC SUBSET_INTERIOR THEN MATCH_MP_TAC HULL_MINIMAL THEN
+ ASM_REWRITE_TAC[INSERT_SUBSET];
+ ALL_TAC] THEN
+ MATCH_MP_TAC STARLIKE_NEGLIGIBLE_STRONG THEN
+ EXISTS_TAC `a:real^N` THEN REWRITE_TAC[FRONTIER_CLOSED] THEN
+ REPEAT GEN_TAC THEN STRIP_TAC THEN
+ REWRITE_TAC[frontier; IN_DIFF; DE_MORGAN_THM] THEN DISJ2_TAC THEN
+ SIMP_TAC[VECTOR_ARITH
+ `a + c % x:real^N = (a + x) - (1 - c) % ((a + x) - a)`] THEN
+ MATCH_MP_TAC IN_INTERIOR_CLOSURE_CONVEX_SHRINK THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[frontier; IN_DIFF]) THEN
+ ASM_REWRITE_TAC[] THEN ASM_REAL_ARITH_TAC);; *)
+
+lemma GMEASURABLE_CONVEX: True .. (*
+ "!s:real^N->bool. convex s \<and> bounded s ==> gmeasurable s"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC GMEASURABLE_JORDAN THEN
+ ASM_SIMP_TAC[NEGLIGIBLE_CONVEX_FRONTIER]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Various special cases. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma NEGLIGIBLE_SPHERE: True .. (*
+ "!a r. negligible {x:real^N | dist(a,x) = r}"
+qed REWRITE_TAC[GSYM FRONTIER_CBALL] THEN
+ SIMP_TAC[NEGLIGIBLE_CONVEX_FRONTIER; CONVEX_CBALL]);; *)
+
+lemma GMEASURABLE_BALL: True .. (*
+ "!a r. gmeasurable(ball(a,r))"
+qed SIMP_TAC[MEASURABLE_OPEN; BOUNDED_BALL; OPEN_BALL]);; *)
+
+lemma GMEASURABLE_CBALL: True .. (*
+ "!a r. gmeasurable(cball(a,r))"
+qed SIMP_TAC[MEASURABLE_COMPACT; COMPACT_CBALL]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Negligibility of image under non-injective linear map. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma NEGLIGIBLE_LINEAR_SINGULAR_IMAGE: True .. (*
+ "!f:real^N->real^N s.
+ linear f \<and> ~(!x y. f(x) = f(y) ==> x = y)
+ ==> negligible(IMAGE f s)"
+qed REPEAT GEN_TAC THEN
+ DISCH_THEN(MP_TAC o MATCH_MP LINEAR_SINGULAR_IMAGE_HYPERPLANE) THEN
+ DISCH_THEN(X_CHOOSE_THEN `a:real^N` STRIP_ASSUME_TAC) THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `{x:real^N | a dot x = 0}` THEN
+ ASM_SIMP_TAC[NEGLIGIBLE_HYPERPLANE]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Approximation of gmeasurable set by union of intervals. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma COVERING_LEMMA: True .. (*
+ "!a b:real^N s g.
+ s \<subseteq> {a..b} \<and> ~({a<..<b} = {}) \<and> gauge g
+ ==> ?d. COUNTABLE d \<and>
+ (!k. k \<in> d ==> k \<subseteq> {a..b} \<and> ~(k = {}) \<and>
+ (\<exists>c d. k = {c..d})) \<and>
+ (!k1 k2. k1 \<in> d \<and> k2 \<in> d \<and> ~(k1 = k2)
+ ==> interior k1 \<inter> interior k2 = {}) \<and>
+ (!k. k \<in> d ==> ?x. x \<in> (s \<inter> k) \<and> k \<subseteq> g(x)) \<and>
+ s \<subseteq> \<Union>d"
+qed REPEAT STRIP_TAC THEN
+ SUBGOAL_THEN
+ `?d. COUNTABLE d \<and>
+ (!k. k \<in> d ==> k \<subseteq> {a..b} \<and> ~(k = {}) \<and>
+ (\<exists>c d:real^N. k = {c..d})) \<and>
+ (!k1 k2. k1 \<in> d \<and> k2 \<in> d
+ ==> k1 \<subseteq> k2 \/ k2 \<subseteq> k1 \/
+ interior k1 \<inter> interior k2 = {}) \<and>
+ (!x. x \<in> s ==> ?k. k \<in> d \<and> x \<in> k \<and> k \<subseteq> g(x)) \<and>
+ (!k. k \<in> d ==> FINITE {l | l \<in> d \<and> k \<subseteq> l})`
+ ASSUME_TAC THENL
+ [EXISTS_TAC
+ `IMAGE (\<lambda>(n,v).
+ interval[(lambda i. a$i + (v$i) / 2 pow n *
+ ((b:real^N)$i - (a:real^N)$i)):real^N,
+ (lambda i. a$i + ((v$i) + 1) / 2 pow n * (b$i - a$i))])
+ {n,v | n \<in> (:num) \<and>
+ v \<in> {v:num^N | !i. 1 \<le> i \<and> i \<le> dimindex(:N)
+ ==> v$i < 2 EXP n}}` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC COUNTABLE_IMAGE THEN
+ MATCH_MP_TAC COUNTABLE_PRODUCT_DEPENDENT THEN
+ REWRITE_TAC[NUM_COUNTABLE; IN_UNIV] THEN
+ GEN_TAC THEN MATCH_MP_TAC FINITE_IMP_COUNTABLE THEN
+ MATCH_MP_TAC FINITE_CART THEN REWRITE_TAC[FINITE_NUMSEG_LT];
+ ALL_TAC] THEN
+ CONJ_TAC THENL
+ [REWRITE_TAC[FORALL_IN_IMAGE; FORALL_PAIR_THM] THEN
+ MAP_EVERY X_GEN_TAC [`n:num`; `v:num^N`] THEN
+ REWRITE_TAC[IN_ELIM_PAIR_THM] THEN
+ REWRITE_TAC[IN_ELIM_THM; IN_UNIV] THEN DISCH_TAC THEN
+ REWRITE_TAC[CONJ_ASSOC] THEN CONJ_TAC THENL [ALL_TAC; MESON_TAC[]] THEN
+ SIMP_TAC[INTERVAL_NE_EMPTY; SUBSET_INTERVAL; LAMBDA_BETA] THEN
+ REWRITE_TAC[REAL_LE_LADD; REAL_LE_ADDR; REAL_ARITH
+ `a + x * (b - a) \<le> b \<longleftrightarrow> 0 \<le> (1 - x) * (b - a)`] THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[INTERVAL_NE_EMPTY]) THEN
+ REPEAT STRIP_TAC THEN
+ (MATCH_MP_TAC REAL_LE_MUL ORELSE MATCH_MP_TAC REAL_LE_RMUL) THEN
+ ASM_SIMP_TAC[REAL_SUB_LE; REAL_LE_DIV2_EQ; REAL_LT_POW2] THEN
+ ASM_SIMP_TAC[REAL_LE_LDIV_EQ; REAL_LE_RDIV_EQ; REAL_LT_POW2] THEN
+ REWRITE_TAC[REAL_MUL_LZERO; REAL_POS; REAL_MUL_LID; REAL_LE_ADDR] THEN
+ SIMP_TAC[REAL_OF_NUM_ADD; REAL_OF_NUM_POW; REAL_OF_NUM_LE] THEN
+ ASM_SIMP_TAC[ARITH_RULE `x + 1 \<le> y \<longleftrightarrow> x < y`; REAL_LT_IMP_LE];
+ ALL_TAC] THEN
+ CONJ_TAC THENL
+ [ONCE_REWRITE_TAC[IMP_CONJ] THEN
+ REWRITE_TAC[FORALL_IN_IMAGE; FORALL_PAIR_THM; RIGHT_FORALL_IMP_THM] THEN
+ REWRITE_TAC[IN_ELIM_PAIR_THM; IN_UNIV] THEN REWRITE_TAC[IN_ELIM_THM] THEN
+ REWRITE_TAC[RIGHT_IMP_FORALL_THM] THEN
+ GEN_REWRITE_TAC BINDER_CONV [SWAP_FORALL_THM] THEN
+ MATCH_MP_TAC WLOG_LE THEN CONJ_TAC THENL
+ [REPEAT GEN_TAC THEN
+ GEN_REWRITE_TAC RAND_CONV [SWAP_FORALL_THM] THEN
+ REPEAT(AP_TERM_TAC THEN ABS_TAC) THEN SET_TAC[];
+ ALL_TAC] THEN
+ MAP_EVERY X_GEN_TAC [`m:num`; `n:num`] THEN DISCH_TAC THEN
+ MAP_EVERY X_GEN_TAC [`v:num^N`; `w:num^N`] THEN REPEAT DISCH_TAC THEN
+ REWRITE_TAC[INTERIOR_CLOSED_INTERVAL; SUBSET_INTERVAL] THEN
+ SIMP_TAC[DISJOINT_INTERVAL; LAMBDA_BETA] THEN
+ MATCH_MP_TAC(TAUT `p \/ q \/ r ==> (a ==> p) \/ (b ==> q) \/ r`) THEN
+ ONCE_REWRITE_TAC[TAUT `a \<and> b \<and> c \<longleftrightarrow> ~(a \<and> b ==> ~c)`] THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[INTERVAL_NE_EMPTY]) THEN
+ ASM_SIMP_TAC[REAL_LE_LADD; REAL_LE_RMUL_EQ; REAL_SUB_LT; LAMBDA_BETA] THEN
+ REWRITE_TAC[NOT_IMP; REAL_LE_LADD] THEN
+ ASM_SIMP_TAC[REAL_LE_DIV2_EQ; REAL_LT_POW2] THEN
+ REWRITE_TAC[REAL_ARITH `~(x + 1 \<le> x)`] THEN DISJ2_TAC THEN
+ MATCH_MP_TAC(MESON[]
+ `(!i. ~P i ==> Q i) ==> (!i. Q i) \/ (\<exists>i. P i)`) THEN
+ X_GEN_TAC `i:num` THEN
+ DISCH_THEN(fun th -> STRIP_TAC THEN MP_TAC th) THEN
+ ASM_REWRITE_TAC[DE_MORGAN_THM; REAL_NOT_LE] THEN
+ UNDISCH_TAC `m:num \<le> n` THEN REWRITE_TAC[LE_EXISTS] THEN
+ DISCH_THEN(X_CHOOSE_THEN `p:num` SUBST1_TAC) THEN
+ ONCE_REWRITE_TAC[ADD_SYM] THEN
+ REWRITE_TAC[REAL_POW_ADD; real_div; REAL_INV_MUL] THEN
+ REWRITE_TAC[REAL_MUL_ASSOC] THEN REWRITE_TAC[GSYM real_div] THEN
+ ASM_SIMP_TAC[REAL_LE_DIV2_EQ; REAL_LT_POW2; REAL_LT_DIV2_EQ] THEN
+ ASM_SIMP_TAC[REAL_LE_LDIV_EQ; REAL_LE_RDIV_EQ; REAL_LT_POW2;
+ REAL_LT_LDIV_EQ; REAL_LT_RDIV_EQ] THEN
+ SIMP_TAC[REAL_LT_INTEGERS; INTEGER_CLOSED] THEN REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ CONJ_TAC THENL
+ [X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+ SUBGOAL_THEN
+ `?e. 0 < e \<and> !y. (!i. 1 \<le> i \<and> i \<le> dimindex(:N)
+ ==> abs((x:real^N)$i - (y:real^N)$i) \<le> e)
+ ==> y \<in> g(x)`
+ STRIP_ASSUME_TAC THENL
+ [FIRST_ASSUM(MP_TAC o SPEC `x:real^N` o GEN_REWRITE_RULE I [gauge]) THEN
+ STRIP_TAC THEN
+ FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [OPEN_CONTAINS_BALL]) THEN
+ DISCH_THEN(MP_TAC o SPEC `x:real^N`) THEN ASM_REWRITE_TAC[] THEN
+ DISCH_THEN(X_CHOOSE_THEN `e:real` STRIP_ASSUME_TAC) THEN
+ EXISTS_TAC `e / 2 / (dimindex(:N))` THEN
+ ASM_SIMP_TAC[REAL_LT_DIV; REAL_OF_NUM_LT; LE_1; DIMINDEX_GE_1;
+ ARITH] THEN
+ X_GEN_TAC `y:real^N` THEN STRIP_TAC THEN
+ MATCH_MP_TAC(SET_RULE `!s. s \<subseteq> t \<and> x \<in> s ==> x \<in> t`) THEN
+ EXISTS_TAC `ball(x:real^N,e)` THEN ASM_REWRITE_TAC[IN_BALL] THEN
+ MATCH_MP_TAC(REAL_ARITH `0 < e \<and> x \<le> e / 2 ==> x < e`) THEN
+ ASM_REWRITE_TAC[dist] THEN MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `sum(1..dimindex(:N)) (\<lambda>i. abs((x - y:real^N)$i))` THEN
+ REWRITE_TAC[NORM_LE_L1] THEN MATCH_MP_TAC SUM_BOUND_GEN THEN
+ ASM_SIMP_TAC[IN_NUMSEG; FINITE_NUMSEG; NUMSEG_EMPTY; NOT_LT;
+ DIMINDEX_GE_1; VECTOR_SUB_COMPONENT; CARD_NUMSEG_1];
+ ALL_TAC] THEN
+ REWRITE_TAC[EXISTS_IN_IMAGE; EXISTS_PAIR_THM; IN_ELIM_PAIR_THM] THEN
+ MP_TAC(SPECL [`1 / 2`; `e / norm(b - a:real^N)`]
+ REAL_ARCH_POW_INV) THEN
+ SUBGOAL_THEN `0 < norm(b - a:real^N)` ASSUME_TAC THENL
+ [ASM_MESON_TAC[VECTOR_SUB_EQ; NORM_POS_LT; INTERVAL_SING]; ALL_TAC] THEN
+ CONV_TAC REAL_RAT_REDUCE_CONV THEN ASM_SIMP_TAC[REAL_LT_DIV] THEN
+ MATCH_MP_TAC MONO_EXISTS THEN X_GEN_TAC `n:num` THEN
+ REWRITE_TAC[real_div; REAL_MUL_LID; REAL_POW_INV] THEN DISCH_TAC THEN
+ SIMP_TAC[IN_ELIM_THM; IN_INTERVAL; SUBSET; LAMBDA_BETA] THEN
+ MATCH_MP_TAC(MESON[]
+ `(!x. Q x ==> R x) \<and> (\<exists>x. P x \<and> Q x) ==> ?x. P x \<and> Q x \<and> R x`) THEN
+ CONJ_TAC THENL
+ [REWRITE_TAC[RIGHT_IMP_FORALL_THM] THEN
+ MAP_EVERY X_GEN_TAC [`w:num^N`; `y:real^N`] THEN
+ REWRITE_TAC[IMP_IMP; AND_FORALL_THM] THEN
+ DISCH_THEN(fun th -> FIRST_X_ASSUM MATCH_MP_TAC THEN MP_TAC th) THEN
+ MATCH_MP_TAC MONO_FORALL THEN X_GEN_TAC `i:num` THEN
+ DISCH_THEN(fun th -> STRIP_TAC THEN MP_TAC th) THEN
+ ASM_REWRITE_TAC[] THEN DISCH_THEN(MP_TAC o MATCH_MP (REAL_ARITH
+ `(a + n \<le> x \<and> x \<le> a + m) \<and>
+ (a + n \<le> y \<and> y \<le> a + m) ==> abs(x - y) \<le> m - n`)) THEN
+ MATCH_MP_TAC(REAL_ARITH
+ `y * z \<le> e
+ ==> a \<le> ((x + 1) * y) * z - ((x * y) * z) ==> a \<le> e`) THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[INTERVAL_NE_EMPTY]) THEN
+ ASM_SIMP_TAC[GSYM REAL_LE_RDIV_EQ; REAL_SUB_LT] THEN
+ FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP
+ (REAL_ARITH `n < e * x ==> 0 \<le> e * (inv y - x) ==> n \<le> e / y`)) THEN
+ MATCH_MP_TAC REAL_LE_MUL THEN ASM_SIMP_TAC[REAL_LT_IMP_LE] THEN
+ REWRITE_TAC[REAL_SUB_LE] THEN MATCH_MP_TAC REAL_LE_INV2 THEN
+ ASM_SIMP_TAC[REAL_SUB_LT] THEN
+ MP_TAC(SPECL [`b - a:real^N`; `i:num`] COMPONENT_LE_NORM) THEN
+ ASM_SIMP_TAC[VECTOR_SUB_COMPONENT] THEN REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ REWRITE_TAC[IN_UNIV; AND_FORALL_THM] THEN
+ REWRITE_TAC[TAUT `(a ==> c) \<and> (a ==> b) \<longleftrightarrow> a ==> b \<and> c`] THEN
+ REWRITE_TAC[GSYM LAMBDA_SKOLEM] THEN X_GEN_TAC `i:num` THEN
+ STRIP_TAC THEN
+ SUBGOAL_THEN `(x:real^N) \<in> {a..b}` MP_TAC THENL
+ [ASM SET_TAC[]; ALL_TAC] THEN REWRITE_TAC[IN_INTERVAL] THEN
+ DISCH_THEN(MP_TAC o SPEC `i:num`) THEN ASM_REWRITE_TAC[] THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[INTERVAL_NE_EMPTY]) THEN STRIP_TAC THEN
+ DISJ_CASES_TAC(MATCH_MP (REAL_ARITH `x \<le> y ==> x = y \/ x < y`)
+ (ASSUME `(x:real^N)$i \<le> (b:real^N)$i`))
+ THENL
+ [EXISTS_TAC `2 EXP n - 1` THEN
+ SIMP_TAC[GSYM REAL_OF_NUM_SUB; GSYM REAL_OF_NUM_LT;
+ EXP_LT_0; LE_1; ARITH] THEN
+ ASM_REWRITE_TAC[REAL_SUB_ADD; REAL_ARITH `a - 1 < a`] THEN
+ MATCH_MP_TAC(REAL_ARITH
+ `1 * (b - a) = x \<and> y \<le> x ==> a + y \<le> b \<and> b \<le> a + x`) THEN
+ ASM_SIMP_TAC[REAL_EQ_MUL_RCANCEL; REAL_LT_IMP_NZ; REAL_LE_RMUL_EQ;
+ REAL_SUB_LT; REAL_LT_INV_EQ; REAL_LT_POW2] THEN
+ SIMP_TAC[GSYM REAL_OF_NUM_POW; REAL_MUL_RINV; REAL_POW_EQ_0;
+ REAL_OF_NUM_EQ; ARITH_EQ] THEN REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ MP_TAC(SPEC `2 pow n * ((x:real^N)$i - (a:real^N)$i) /
+ ((b:real^N)$i - (a:real^N)$i)` FLOOR_POS) THEN
+ ANTS_TAC THENL
+ [ASM_MESON_TAC[REAL_LE_MUL; REAL_LE_MUL; REAL_POW_LE; REAL_POS;
+ REAL_SUB_LE; REAL_LT_IMP_LE; REAL_LE_DIV];
+ ALL_TAC] THEN
+ MATCH_MP_TAC MONO_EXISTS THEN X_GEN_TAC `m:num` THEN
+ REWRITE_TAC[GSYM REAL_OF_NUM_LT; GSYM REAL_OF_NUM_POW] THEN
+ DISCH_THEN(SUBST_ALL_TAC o SYM) THEN
+ REWRITE_TAC[REAL_ARITH `a + b * c \<le> x \<and> x \<le> a + b' * c \<longleftrightarrow>
+ b * c \<le> x - a \<and> x - a \<le> b' * c`] THEN
+ ASM_SIMP_TAC[GSYM REAL_LE_LDIV_EQ; GSYM REAL_LE_RDIV_EQ;
+ REAL_SUB_LT; GSYM real_div] THEN
+ ONCE_REWRITE_TAC[REAL_MUL_SYM] THEN
+ SIMP_TAC[REAL_LE_LDIV_EQ; REAL_LE_RDIV_EQ; REAL_LT_POW2] THEN
+ SIMP_TAC[FLOOR; REAL_LT_IMP_LE] THEN MATCH_MP_TAC REAL_LET_TRANS THEN
+ EXISTS_TAC `((x:real^N)$i - (a:real^N)$i) /
+ ((b:real^N)$i - (a:real^N)$i) *
+ 2 pow n` THEN
+ REWRITE_TAC[FLOOR] THEN GEN_REWRITE_TAC RAND_CONV [GSYM REAL_MUL_LID] THEN
+ ASM_SIMP_TAC[REAL_LT_RMUL_EQ; REAL_LT_POW2] THEN
+ ASM_SIMP_TAC[REAL_LT_LDIV_EQ; REAL_MUL_LID; REAL_SUB_LT] THEN
+ ASM_REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ REWRITE_TAC[FORALL_IN_IMAGE; FORALL_PAIR_THM; IN_ELIM_PAIR_THM] THEN
+ MAP_EVERY X_GEN_TAC [`n:num`; `v:num^N`] THEN
+ REWRITE_TAC[IN_ELIM_THM; IN_UNIV] THEN DISCH_TAC THEN
+ MATCH_MP_TAC FINITE_SUBSET THEN EXISTS_TAC
+ `IMAGE (\<lambda>(n,v).
+ interval[(lambda i. a$i + (v$i) / 2 pow n *
+ ((b:real^N)$i - (a:real^N)$i)):real^N,
+ (lambda i. a$i + ((v$i) + 1) / 2 pow n * (b$i - a$i))])
+ {m,v | m \<in> 0..n \<and>
+ v \<in> {v:num^N | !i. 1 \<le> i \<and> i \<le> dimindex(:N)
+ ==> v$i < 2 EXP m}}` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC FINITE_IMAGE THEN
+ MATCH_MP_TAC FINITE_PRODUCT_DEPENDENT THEN
+ REWRITE_TAC[FINITE_NUMSEG] THEN REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC FINITE_CART THEN REWRITE_TAC[FINITE_NUMSEG_LT];
+ ALL_TAC] THEN
+ GEN_REWRITE_TAC I [SUBSET] THEN
+ REWRITE_TAC[IN_ELIM_THM] THEN ONCE_REWRITE_TAC[IMP_CONJ] THEN
+ REWRITE_TAC[FORALL_IN_IMAGE; FORALL_PAIR_THM; IN_ELIM_PAIR_THM] THEN
+ MAP_EVERY X_GEN_TAC [`m:num`; `w:num^N`] THEN DISCH_TAC THEN
+ DISCH_TAC THEN SIMP_TAC[IN_IMAGE; EXISTS_PAIR_THM; IN_ELIM_PAIR_THM] THEN
+ MAP_EVERY EXISTS_TAC [`m:num`; `w:num^N`] THEN ASM_REWRITE_TAC[] THEN
+ REWRITE_TAC[IN_NUMSEG; GSYM NOT_LT; LT] THEN DISCH_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o GEN_REWRITE_RULE I [SUBSET_INTERVAL]) THEN
+ SIMP_TAC[NOT_IMP; LAMBDA_BETA] THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[INTERVAL_NE_EMPTY]) THEN
+ ASM_SIMP_TAC[REAL_LE_LADD; REAL_LE_RMUL_EQ; REAL_SUB_LT] THEN
+ ASM_SIMP_TAC[REAL_LE_DIV2_EQ; REAL_LT_POW2] THEN
+ REWRITE_TAC[REAL_ARITH `x \<le> x + 1`] THEN
+ DISCH_THEN(MP_TAC o SPEC `1`) THEN
+ REWRITE_TAC[LE_REFL; DIMINDEX_GE_1] THEN
+ DISCH_THEN(MP_TAC o MATCH_MP (REAL_ARITH
+ `w / m \<le> v / n \<and> (v + 1) / n \<le> (w + 1) / m
+ ==> inv n \<le> inv m`)) THEN
+ REWRITE_TAC[REAL_NOT_LE] THEN MATCH_MP_TAC REAL_LT_INV2 THEN
+ ASM_REWRITE_TAC[REAL_LT_POW2] THEN MATCH_MP_TAC REAL_POW_MONO_LT THEN
+ ASM_REWRITE_TAC[] THEN CONV_TAC REAL_RAT_REDUCE_CONV;
+ ALL_TAC] THEN
+ SUBGOAL_THEN
+ `?d. COUNTABLE d \<and>
+ (!k. k \<in> d ==> k \<subseteq> {a..b} \<and> ~(k = {}) \<and>
+ (\<exists>c d:real^N. k = {c..d})) \<and>
+ (!k1 k2. k1 \<in> d \<and> k2 \<in> d
+ ==> k1 \<subseteq> k2 \/ k2 \<subseteq> k1 \/
+ interior k1 \<inter> interior k2 = {}) \<and>
+ (!k. k \<in> d ==> (\<exists>x. x \<in> s \<inter> k \<and> k \<subseteq> g x)) \<and>
+ (!k. k \<in> d ==> FINITE {l | l \<in> d \<and> k \<subseteq> l}) \<and>
+ s \<subseteq> \<Union>d`
+ MP_TAC THENL
+ [FIRST_X_ASSUM(X_CHOOSE_THEN `d:(real^N->bool)->bool` STRIP_ASSUME_TAC) THEN
+ EXISTS_TAC
+ `{k:real^N->bool | k \<in> d \<and> ?x. x \<in> (s \<inter> k) \<and> k \<subseteq> g x}` THEN
+ ASM_SIMP_TAC[IN_ELIM_THM] THEN REPEAT CONJ_TAC THENL
+ [MATCH_MP_TAC COUNTABLE_SUBSET THEN
+ EXISTS_TAC `d:(real^N->bool)->bool` THEN
+ ASM_REWRITE_TAC[] THEN SET_TAC[];
+ X_GEN_TAC `k:real^N->bool` THEN REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC FINITE_SUBSET THEN
+ EXISTS_TAC `{l:real^N->bool | l \<in> d \<and> k \<subseteq> l}` THEN
+ ASM_REWRITE_TAC[] THEN SET_TAC[];
+ ASM SET_TAC[]];
+ ALL_TAC] THEN
+ DISCH_THEN(X_CHOOSE_THEN `d:(real^N->bool)->bool` STRIP_ASSUME_TAC) THEN
+ EXISTS_TAC
+ `{k:real^N->bool | k \<in> d \<and> !k'. k' \<in> d \<and> ~(k = k')
+ ==> ~(k \<subseteq> k')}` THEN
+ ASM_SIMP_TAC[IN_ELIM_THM] THEN REPEAT CONJ_TAC THENL
+ [MATCH_MP_TAC COUNTABLE_SUBSET THEN EXISTS_TAC `d:(real^N->bool)->bool` THEN
+ ASM_REWRITE_TAC[] THEN SET_TAC[];
+ ASM SET_TAC[];
+ ALL_TAC] THEN
+ FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP
+ (REWRITE_RULE[IMP_CONJ] SUBSET_TRANS)) THEN
+ GEN_REWRITE_TAC I [SUBSET] THEN REWRITE_TAC[FORALL_IN_UNIONS] THEN
+ MAP_EVERY X_GEN_TAC [`k:real^N->bool`; `x:real^N`] THEN DISCH_TAC THEN
+ REWRITE_TAC[IN_UNIONS; IN_ELIM_THM] THEN
+ MP_TAC(ISPEC `\k l:real^N->bool. k \<in> d \<and> l \<in> d \<and> l \<subseteq> k \<and> ~(k = l)`
+ WF_FINITE) THEN
+ REWRITE_TAC[WF] THEN ANTS_TAC THENL
+ [CONJ_TAC THENL [SET_TAC[]; ALL_TAC] THEN X_GEN_TAC `l:real^N->bool` THEN
+ ASM_CASES_TAC `(l:real^N->bool) \<in> d` THEN
+ ASM_REWRITE_TAC[EMPTY_GSPEC; FINITE_RULES] THEN
+ MATCH_MP_TAC FINITE_SUBSET THEN
+ EXISTS_TAC `{m:real^N->bool | m \<in> d \<and> l \<subseteq> m}` THEN
+ ASM_SIMP_TAC[] THEN SET_TAC[];
+ ALL_TAC] THEN
+ DISCH_THEN(MP_TAC o SPEC `\l:real^N->bool. l \<in> d \<and> x \<in> l`) THEN
+ REWRITE_TAC[] THEN ANTS_TAC THENL [SET_TAC[]; ALL_TAC] THEN
+ MATCH_MP_TAC MONO_EXISTS THEN ASM SET_TAC[]);; *)
+
+lemma GMEASURABLE_OUTER_INTERVALS_BOUNDED: True .. (*
+ "!s a b:real^N e.
+ gmeasurable s \<and> s \<subseteq> {a..b} \<and> 0 < e
+ ==> ?d. COUNTABLE d \<and>
+ (!k. k \<in> d ==> k \<subseteq> {a..b} \<and> ~(k = {}) \<and>
+ (\<exists>c d. k = {c..d})) \<and>
+ (!k1 k2. k1 \<in> d \<and> k2 \<in> d \<and> ~(k1 = k2)
+ ==> interior k1 \<inter> interior k2 = {}) \<and>
+ s \<subseteq> \<Union>d \<and>
+ gmeasurable (\<Union>d) \<and>
+ gmeasure (\<Union>d) \<le> gmeasure s + e"
+qed lemma lemma = prove
+ (`(!x y. (x,y) \<in> IMAGE (\<lambda>z. f z,g z) s ==> P x y) \<longleftrightarrow>
+ (!z. z \<in> s ==> P (f z) (g z))"
+qed REWRITE_TAC[IN_IMAGE; PAIR_EQ] THEN MESON_TAC[]) in
+ REPEAT GEN_TAC THEN
+ ASM_CASES_TAC `interval[a:real^N,b] = {}` THENL
+ [ASM_REWRITE_TAC[SUBSET_EMPTY] THEN STRIP_TAC THEN
+ EXISTS_TAC `{}:(real^N->bool)->bool` THEN
+ ASM_REWRITE_TAC[NOT_IN_EMPTY; UNIONS_0; MEASURE_EMPTY; REAL_ADD_LID;
+ SUBSET_REFL; COUNTABLE_EMPTY; GMEASURABLE_EMPTY] THEN
+ ASM_SIMP_TAC[REAL_LT_IMP_LE];
+ ALL_TAC] THEN
+ STRIP_TAC THEN ASM_CASES_TAC `interval(a:real^N,b) = {}` THENL
+ [EXISTS_TAC `{interval[a:real^N,b]}` THEN
+ REWRITE_TAC[UNIONS_1; COUNTABLE_SING] THEN
+ ASM_REWRITE_TAC[IMP_CONJ; RIGHT_FORALL_IMP_THM; FORALL_IN_INSERT;
+ NOT_IN_EMPTY; SUBSET_REFL; GMEASURABLE_INTERVAL] THEN
+ CONJ_TAC THENL [MESON_TAC[]; ALL_TAC] THEN
+ SUBGOAL_THEN
+ `measure(interval[a:real^N,b]) = 0 \<and> measure(s:real^N->bool) = 0`
+ (fun th -> ASM_SIMP_TAC[th; REAL_LT_IMP_LE; REAL_ADD_LID]) THEN
+ SUBGOAL_THEN
+ `interval[a:real^N,b] has_gmeasure 0 \<and> (s:real^N->bool) has_gmeasure 0`
+ (fun th -> MESON_TAC[th; MEASURE_UNIQUE]) THEN
+ REWRITE_TAC[HAS_GMEASURE_0] THEN
+ MATCH_MP_TAC(TAUT `a \<and> (a ==> b) ==> a \<and> b`) THEN CONJ_TAC THENL
+ [ASM_REWRITE_TAC[NEGLIGIBLE_INTERVAL];
+ ASM_MESON_TAC[NEGLIGIBLE_SUBSET]];
+ ALL_TAC] THEN
+ FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [measurable]) THEN
+ DISCH_THEN(X_CHOOSE_TAC `m:real`) THEN
+ FIRST_ASSUM(ASSUME_TAC o MATCH_MP MEASURE_UNIQUE) THEN
+ SUBGOAL_THEN
+ `((\<lambda>x:real^N. if x \<in> s then 1 else 0) has_integral (lift m))
+ {a..b}`
+ ASSUME_TAC THENL
+ [ONCE_REWRITE_TAC[GSYM HAS_INTEGRAL_RESTRICT_UNIV] THEN
+ FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [HAS_GMEASURE]) THEN
+ MATCH_MP_TAC(REWRITE_RULE[IMP_CONJ] HAS_INTEGRAL_EQ) THEN
+ ASM SET_TAC[];
+ ALL_TAC] THEN
+ FIRST_ASSUM(ASSUME_TAC o MATCH_MP HAS_INTEGRAL_INTEGRABLE) THEN
+ FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [has_integral]) THEN
+ DISCH_THEN(MP_TAC o SPEC `e:real`) THEN ASM_REWRITE_TAC[] THEN
+ DISCH_THEN(X_CHOOSE_THEN `g:real^N->real^N->bool` STRIP_ASSUME_TAC) THEN
+ MP_TAC(SPECL [`a:real^N`; `b:real^N`; `s:real^N->bool`;
+ `g:real^N->real^N->bool`] COVERING_LEMMA) THEN
+ ASM_REWRITE_TAC[] THEN MATCH_MP_TAC MONO_EXISTS THEN
+ X_GEN_TAC `d:(real^N->bool)->bool` THEN STRIP_TAC THEN ASM_REWRITE_TAC[] THEN
+ MP_TAC(ISPECL [`(\<lambda>x. if x \<in> s then 1 else 0):real^N->real^1`;
+ `a:real^N`; `b:real^N`; `g:real^N->real^N->bool`;
+ `e:real`]
+ HENSTOCK_LEMMA_PART1) THEN
+ ASM_REWRITE_TAC[] THEN
+ FIRST_ASSUM(SUBST1_TAC o MATCH_MP INTEGRAL_UNIQUE) THEN
+ ASM_REWRITE_TAC[] THEN DISCH_THEN(LABEL_TAC "*") THEN
+ SUBGOAL_THEN
+ `!k l:real^N->bool. k \<in> d \<and> l \<in> d \<and> ~(k = l)
+ ==> negligible(k \<inter> l)`
+ ASSUME_TAC THENL
+ [REPEAT STRIP_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPECL [`k:real^N->bool`; `l:real^N->bool`]) THEN
+ ASM_SIMP_TAC[] THEN
+ SUBGOAL_THEN
+ `?x y:real^N u v:real^N. k = {x..y} \<and> l = {u..v}`
+ MP_TAC THENL [ASM_MESON_TAC[SUBSET]; ALL_TAC] THEN
+ DISCH_THEN(REPEAT_TCL CHOOSE_THEN (CONJUNCTS_THEN SUBST_ALL_TAC)) THEN
+ REWRITE_TAC[INTERIOR_CLOSED_INTERVAL] THEN DISCH_TAC THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `(interval[x:real^N,y] DIFF {x<..<y}) UNION
+ (interval[u:real^N,v] DIFF {u<..<v}) UNION
+ (interval (x,y) \<inter> interval (u,v))` THEN
+ CONJ_TAC THENL [ALL_TAC; SET_TAC[]] THEN
+ ASM_REWRITE_TAC[UNION_EMPTY] THEN
+ SIMP_TAC[NEGLIGIBLE_UNION; NEGLIGIBLE_FRONTIER_INTERVAL];
+ ALL_TAC] THEN
+ SUBGOAL_THEN
+ `!D. FINITE D \<and> D \<subseteq> d
+ ==> gmeasurable(\<Union>D :real^N->bool) \<and> measure(\<Union>D) \<le> m + e`
+ ASSUME_TAC THENL
+ [GEN_TAC THEN STRIP_TAC THEN
+ SUBGOAL_THEN
+ `?t:(real^N->bool)->real^N. !k. k \<in> D ==> t(k) \<in> (s \<inter> k) \<and>
+ k \<subseteq> (g(t k))`
+ (CHOOSE_THEN (LABEL_TAC "+")) THENL
+ [REWRITE_TAC[GSYM SKOLEM_THM] THEN ASM SET_TAC[]; ALL_TAC] THEN
+ REMOVE_THEN "*" (MP_TAC o SPEC
+ `IMAGE (\<lambda>k. (t:(real^N->bool)->real^N) k,k) D`) THEN
+ ASM_SIMP_TAC[VSUM_IMAGE; PAIR_EQ] THEN REWRITE_TAC[o_DEF] THEN
+ ANTS_TAC THENL
+ [REWRITE_TAC[tagged_partial_division_of; fine] THEN
+ ASM_SIMP_TAC[FINITE_IMAGE; FORALL_IN_IMAGE] THEN
+ REWRITE_TAC[lemma; RIGHT_FORALL_IMP_THM; IMP_CONJ; PAIR_EQ] THEN
+ ASM_SIMP_TAC[] THEN
+ CONJ_TAC THENL [ASM SET_TAC[]; ASM_MESON_TAC[SUBSET]];
+ ALL_TAC] THEN
+ USE_THEN "+" (MP_TAC o REWRITE_RULE[IN_INTER]) THEN
+ SIMP_TAC[] THEN DISCH_THEN(K ALL_TAC) THEN
+ ASM_SIMP_TAC[VSUM_SUB] THEN
+ SUBGOAL_THEN `D division_of (\<Union>D:real^N->bool)` ASSUME_TAC THENL
+ [REWRITE_TAC[division_of] THEN ASM SET_TAC[]; ALL_TAC] THEN
+ FIRST_ASSUM(ASSUME_TAC o MATCH_MP GMEASURABLE_ELEMENTARY) THEN
+ SUBGOAL_THEN `vsum D (\<lambda>k:real^N->bool. content k % 1) =
+ lift(measure(\<Union>D))`
+ SUBST1_TAC THENL
+ [ONCE_REWRITE_TAC[GSYM _EQ] THEN
+ ASM_SIMP_TAC[LIFT_; _VSUM; o_DEF; _CMUL; _VEC] THEN
+ SIMP_TAC[REAL_MUL_RID; ETA_AX] THEN ASM_MESON_TAC[MEASURE_ELEMENTARY];
+ ALL_TAC] THEN
+ SUBGOAL_THEN
+ `vsum D (\<lambda>k. integral k (\<lambda>x:real^N. if x \<in> s then 1 else 0)) =
+ lift(sum D (\<lambda>k. measure(k \<inter> s)))`
+ SUBST1_TAC THENL
+ [ASM_SIMP_TAC[LIFT_SUM; o_DEF] THEN MATCH_MP_TAC VSUM_EQ THEN
+ X_GEN_TAC `k:real^N->bool` THEN DISCH_TAC THEN REWRITE_TAC[] THEN
+ SUBGOAL_THEN `measurable(k:real^N->bool)` ASSUME_TAC THENL
+ [ASM_MESON_TAC[SUBSET; GMEASURABLE_INTERVAL]; ALL_TAC] THEN
+ ASM_SIMP_TAC[GSYM INTEGRAL_MEASURE_UNIV; GMEASURABLE_INTER] THEN
+ REWRITE_TAC[MESON[IN_INTER]
+ `(if x \<in> k \<inter> s then a else b) =
+ (if x \<in> k then if x \<in> s then a else b else b)`] THEN
+ CONV_TAC SYM_CONV THEN MATCH_MP_TAC INTEGRAL_RESTRICT_UNIV THEN
+ ONCE_REWRITE_TAC[GSYM INTEGRABLE_RESTRICT_UNIV] THEN
+ REWRITE_TAC[MESON[IN_INTER]
+ `(if x \<in> k then if x \<in> s then a else b else b) =
+ (if x \<in> k \<inter> s then a else b)`] THEN
+ ASM_SIMP_TAC[GSYM GMEASURABLE_INTEGRABLE; GMEASURABLE_INTER];
+ ALL_TAC] THEN
+ ASM_REWRITE_TAC[GSYM LIFT_SUB; NORM_LIFT] THEN
+ MATCH_MP_TAC(REAL_ARITH `y \<le> m ==> abs(x - y) \<le> e ==> x \<le> m + e`) THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `measure(\<Union>D \<inter> s:real^N->bool)` THEN
+ CONJ_TAC THENL
+ [ALL_TAC;
+ EXPAND_TAC "m" THEN MATCH_MP_TAC MEASURE_SUBSET THEN
+ ASM_REWRITE_TAC[] THEN CONJ_TAC THENL [ALL_TAC; SET_TAC[]] THEN
+ MATCH_MP_TAC GMEASURABLE_INTER THEN ASM_REWRITE_TAC[]] THEN
+ REWRITE_TAC[INTER_UNIONS] THEN MATCH_MP_TAC REAL_EQ_IMP_LE THEN
+ ONCE_REWRITE_TAC[SIMPLE_IMAGE] THEN CONV_TAC SYM_CONV THEN
+ MATCH_MP_TAC MEASURE_NEGLIGIBLE_UNIONS_IMAGE_STRONG THEN
+ ASM_SIMP_TAC[FINITE_RESTRICT] THEN CONJ_TAC THENL
+ [ASM_MESON_TAC[SUBSET; GMEASURABLE_INTERVAL; GMEASURABLE_INTER];
+ ALL_TAC] THEN
+ MAP_EVERY X_GEN_TAC [`k:real^N->bool`; `l:real^N->bool`] THEN
+ STRIP_TAC THEN MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `k \<inter> l:real^N->bool` THEN ASM_SIMP_TAC[] THEN SET_TAC[];
+ ALL_TAC] THEN
+ ASM_CASES_TAC `FINITE(d:(real^N->bool)->bool)` THENL
+ [ASM_MESON_TAC[SUBSET_REFL]; ALL_TAC] THEN
+ MP_TAC(ISPEC `d:(real^N->bool)->bool` COUNTABLE_AS_INJECTIVE_IMAGE) THEN
+ ASM_REWRITE_TAC[INFINITE] THEN
+ DISCH_THEN(X_CHOOSE_THEN `s:num->real^N->bool`
+ (CONJUNCTS_THEN2 SUBST_ALL_TAC ASSUME_TAC)) THEN
+ MP_TAC(ISPECL [`s:num->real^N->bool`; `m + e:real`]
+ HAS_GMEASURE_COUNTABLE_NEGLIGIBLE_UNIONS) THEN
+ MATCH_MP_TAC(TAUT `a \<and> (a \<and> b ==> c) ==> (a ==> b) ==> c`) THEN
+ REWRITE_TAC[GSYM CONJ_ASSOC] THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[IMP_CONJ; RIGHT_FORALL_IMP_THM;
+ FORALL_IN_IMAGE; IN_UNIV]) THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[IMP_IMP; RIGHT_IMP_FORALL_THM]) THEN
+ REPEAT CONJ_TAC THENL
+ [ASM_MESON_TAC[MEASURABLE_INTERVAL; GMEASURABLE_INTER];
+ ASM_MESON_TAC[];
+ X_GEN_TAC `n:num` THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `IMAGE (s:num->real^N->bool) (0..n)`) THEN
+ SIMP_TAC[FINITE_IMAGE; FINITE_NUMSEG; IMAGE_SUBSET; SUBSET_UNIV] THEN
+ DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC MP_TAC) THEN
+ MATCH_MP_TAC(REAL_ARITH `x = y ==> x \<le> e ==> y \<le> e`) THEN
+ MATCH_MP_TAC MEASURE_NEGLIGIBLE_UNIONS_IMAGE THEN
+ ASM_MESON_TAC[FINITE_NUMSEG; GMEASURABLE_INTERVAL];
+ ALL_TAC] THEN
+ ONCE_REWRITE_TAC[SIMPLE_IMAGE] THEN STRIP_TAC THEN ASM_REWRITE_TAC[] THEN
+ GEN_REWRITE_TAC LAND_CONV [GSYM(CONJUNCT2 LIFT_)] THEN
+ REWRITE_TAC[] THEN
+ MATCH_MP_TAC(ISPEC `sequentially` LIM_COMPONENT_UBOUND) THEN
+ EXISTS_TAC
+ `\n. vsum(from 0 \<inter> (0..n)) (\<lambda>n. lift(measure(s n:real^N->bool)))` THEN
+ ASM_REWRITE_TAC[GSYM sums; TRIVIAL_LIMIT_SEQUENTIALLY] THEN
+ REWRITE_TAC[DIMINDEX_1; ARITH; EVENTUALLY_SEQUENTIALLY] THEN
+ SIMP_TAC[VSUM_COMPONENT; ARITH; DIMINDEX_1] THEN
+ ASM_REWRITE_TAC[GSYM ; LIFT_; FROM_INTER_NUMSEG]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Hence for linear transformation, suffices to check compact intervals. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma GMEASURABLE_LINEAR_IMAGE_INTERVAL: True .. (*
+ "!f a b. linear f ==> gmeasurable(IMAGE f {a..b})"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC GMEASURABLE_CONVEX THEN CONJ_TAC THENL
+ [MATCH_MP_TAC CONVEX_LINEAR_IMAGE THEN
+ ASM_MESON_TAC[CONVEX_INTERVAL];
+ MATCH_MP_TAC BOUNDED_LINEAR_IMAGE THEN
+ ASM_MESON_TAC[BOUNDED_INTERVAL]]);; *)
+
+lemma HAS_GMEASURE_LINEAR_SUFFICIENT: True .. (*
+ "!f:real^N->real^N m.
+ linear f \<and>
+ (!a b. IMAGE f {a..b} has_gmeasure
+ (m * measure{a..b}))
+ ==> !s. gmeasurable s ==> (IMAGE f s) has_gmeasure (m * gmeasure s)"
+qed REPEAT GEN_TAC THEN STRIP_TAC THEN
+ DISJ_CASES_TAC(REAL_ARITH `m < 0 \/ 0 \<le> m`) THENL
+ [FIRST_X_ASSUM(MP_TAC o SPECL [`0:real^N`; `1:real^N`]) THEN
+ DISCH_THEN(MP_TAC o MATCH_MP HAS_GMEASURE_POS_LE) THEN
+ MATCH_MP_TAC(TAUT `~a ==> a ==> b`) THEN
+ MATCH_MP_TAC(REAL_ARITH `0 < --m * x ==> ~(0 \<le> m * x)`) THEN
+ MATCH_MP_TAC REAL_LT_MUL THEN ASM_REWRITE_TAC[REAL_NEG_GT0] THEN
+ REWRITE_TAC[MEASURE_INTERVAL] THEN MATCH_MP_TAC CONTENT_POS_LT THEN
+ SIMP_TAC[VEC_COMPONENT; REAL_LT_01];
+ ALL_TAC] THEN
+ ASM_CASES_TAC `!x y. (f:real^N->real^N) x = f y ==> x = y` THENL
+ [ALL_TAC;
+ SUBGOAL_THEN `!s. negligible(IMAGE (f:real^N->real^N) s)` ASSUME_TAC THENL
+ [ASM_MESON_TAC[NEGLIGIBLE_LINEAR_SINGULAR_IMAGE]; ALL_TAC] THEN
+ SUBGOAL_THEN `m * measure(interval[0:real^N,1]) = 0` MP_TAC THENL
+ [MATCH_MP_TAC(ISPEC `IMAGE (f:real^N->real^N) {0..1}`
+ HAS_GMEASURE_UNIQUE) THEN
+ ASM_REWRITE_TAC[HAS_GMEASURE_0];
+ REWRITE_TAC[REAL_ENTIRE; MEASURE_INTERVAL] THEN
+ MATCH_MP_TAC(TAUT `~b \<and> (a ==> c) ==> a \/ b ==> c`) THEN CONJ_TAC THENL
+ [SIMP_TAC[CONTENT_EQ_0_INTERIOR; INTERIOR_CLOSED_INTERVAL;
+ INTERVAL_NE_EMPTY; VEC_COMPONENT; REAL_LT_01];
+ ASM_SIMP_TAC[REAL_MUL_LZERO; HAS_GMEASURE_0]]]] THEN
+ MP_TAC(ISPEC `f:real^N->real^N` LINEAR_INJECTIVE_ISOMORPHISM) THEN
+ ASM_REWRITE_TAC[] THEN
+ DISCH_THEN(X_CHOOSE_THEN `h:real^N->real^N` STRIP_ASSUME_TAC) THEN
+ UNDISCH_THEN `!x y. (f:real^N->real^N) x = f y ==> x = y` (K ALL_TAC) THEN
+ SUBGOAL_THEN
+ `!s. bounded s \<and> gmeasurable s
+ ==> (IMAGE (f:real^N->real^N) s) has_gmeasure (m * gmeasure s)`
+ ASSUME_TAC THENL
+ [REPEAT STRIP_TAC THEN
+ FIRST_ASSUM(MP_TAC o MATCH_MP BOUNDED_SUBSET_CLOSED_INTERVAL) THEN
+ REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
+ MAP_EVERY X_GEN_TAC [`a:real^N`; `b:real^N`] THEN DISCH_TAC THEN
+ SUBGOAL_THEN
+ `!d. COUNTABLE d \<and>
+ (!k. k \<in> d ==> k \<subseteq> {a..b} \<and> ~(k = {}) \<and>
+ (\<exists>c d. k = {c..d})) \<and>
+ (!k1 k2. k1 \<in> d \<and> k2 \<in> d \<and> ~(k1 = k2)
+ ==> interior k1 \<inter> interior k2 = {})
+ ==> IMAGE (f:real^N->real^N) (\<Union>d) has_gmeasure
+ (m * measure(\<Union>d))`
+ ASSUME_TAC THENL
+ [REWRITE_TAC[IMAGE_UNIONS] THEN REPEAT STRIP_TAC THEN
+ SUBGOAL_THEN
+ `!g:real^N->real^N.
+ linear g
+ ==> !k l. k \<in> d \<and> l \<in> d \<and> ~(k = l)
+ ==> negligible((IMAGE g k) \<inter> (IMAGE g l))`
+ MP_TAC THENL
+ [REPEAT STRIP_TAC THEN
+ ASM_CASES_TAC `!x y. (g:real^N->real^N) x = g y ==> x = y` THENL
+ [ALL_TAC;
+ ASM_MESON_TAC[NEGLIGIBLE_LINEAR_SINGULAR_IMAGE;
+ NEGLIGIBLE_INTER]] THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `frontier(IMAGE (g:real^N->real^N) k \<inter> IMAGE g l) UNION
+ interior(IMAGE g k \<inter> IMAGE g l)` THEN
+ CONJ_TAC THENL
+ [ALL_TAC;
+ REWRITE_TAC[frontier] THEN MATCH_MP_TAC(SET_RULE
+ `s \<subseteq> t ==> s \<subseteq> (t DIFF u) \<union> u`) THEN
+ REWRITE_TAC[CLOSURE_SUBSET]] THEN
+ MATCH_MP_TAC NEGLIGIBLE_UNION THEN CONJ_TAC THENL
+ [MATCH_MP_TAC NEGLIGIBLE_CONVEX_FRONTIER THEN
+ MATCH_MP_TAC CONVEX_INTER THEN CONJ_TAC THEN
+ MATCH_MP_TAC CONVEX_LINEAR_IMAGE THEN ASM_MESON_TAC[CONVEX_INTERVAL];
+ ALL_TAC] THEN
+ REWRITE_TAC[INTERIOR_INTER] THEN MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `IMAGE (g:real^N->real^N) (interior k) INTER
+ IMAGE g (interior l)` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC
+ `IMAGE (g:real^N->real^N) (interior k \<inter> interior l)` THEN
+ CONJ_TAC THENL
+ [ASM_SIMP_TAC[IMAGE_CLAUSES; NEGLIGIBLE_EMPTY]; SET_TAC[]];
+ MATCH_MP_TAC(SET_RULE
+ `s \<subseteq> u \<and> t \<subseteq> v ==> (s \<inter> t) \<subseteq> (u \<inter> v)`) THEN
+ CONJ_TAC THEN MATCH_MP_TAC INTERIOR_IMAGE_SUBSET THEN
+ ASM_MESON_TAC[LINEAR_CONTINUOUS_AT]];
+ ALL_TAC] THEN
+ DISCH_THEN(fun th -> MP_TAC(SPEC `f:real^N->real^N` th) THEN
+ MP_TAC(SPEC `\x:real^N. x` th)) THEN
+ ASM_REWRITE_TAC[LINEAR_ID; SET_RULE `IMAGE (\<lambda>x. x) s = s`] THEN
+ REPEAT STRIP_TAC THEN ASM_CASES_TAC `FINITE(d:(real^N->bool)->bool)` THENL
+ [MP_TAC(ISPECL [`IMAGE (f:real^N->real^N)`; `d:(real^N->bool)->bool`]
+ HAS_GMEASURE_NEGLIGIBLE_UNIONS_IMAGE) THEN
+ ANTS_TAC THENL [ASM_MESON_TAC[measurable]; ALL_TAC] THEN
+ MATCH_MP_TAC EQ_IMP THEN AP_TERM_TAC THEN
+ MATCH_MP_TAC EQ_TRANS THEN
+ EXISTS_TAC `sum d (\<lambda>k:real^N->bool. m * gmeasure k)` THEN CONJ_TAC THENL
+ [MATCH_MP_TAC SUM_EQ THEN ASM_MESON_TAC[MEASURE_UNIQUE]; ALL_TAC] THEN
+ REWRITE_TAC[SUM_LMUL] THEN AP_TERM_TAC THEN
+ CONV_TAC SYM_CONV THEN MATCH_MP_TAC MEASURE_NEGLIGIBLE_UNIONS THEN
+ ASM_REWRITE_TAC[GSYM HAS_GMEASURE_MEASURE] THEN
+ ASM_MESON_TAC[MEASURABLE_INTERVAL];
+ ALL_TAC] THEN
+ MP_TAC(ISPEC `d:(real^N->bool)->bool` COUNTABLE_AS_INJECTIVE_IMAGE) THEN
+ ASM_REWRITE_TAC[INFINITE] THEN
+ DISCH_THEN(X_CHOOSE_THEN `s:num->real^N->bool`
+ (CONJUNCTS_THEN2 SUBST_ALL_TAC ASSUME_TAC)) THEN
+ MP_TAC(ISPEC `s:num->real^N->bool`
+ HAS_GMEASURE_COUNTABLE_NEGLIGIBLE_UNIONS_BOUNDED) THEN
+ MP_TAC(ISPEC `\n:num. IMAGE (f:real^N->real^N) (s n)`
+ HAS_GMEASURE_COUNTABLE_NEGLIGIBLE_UNIONS_BOUNDED) THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[IMP_CONJ; RIGHT_FORALL_IMP_THM;
+ FORALL_IN_IMAGE; IN_UNIV]) THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[IMP_IMP; RIGHT_IMP_FORALL_THM]) THEN
+ ASM_SIMP_TAC[] THEN ONCE_REWRITE_TAC[SIMPLE_IMAGE] THEN ANTS_TAC THENL
+ [REPEAT CONJ_TAC THENL
+ [ASM_MESON_TAC[MEASURABLE_LINEAR_IMAGE_INTERVAL];
+ ASM_MESON_TAC[];
+ ONCE_REWRITE_TAC[GSYM o_DEF] THEN
+ REWRITE_TAC[GSYM IMAGE_UNIONS; IMAGE_o] THEN
+ MATCH_MP_TAC BOUNDED_LINEAR_IMAGE THEN ASM_REWRITE_TAC[] THEN
+ MATCH_MP_TAC BOUNDED_SUBSET THEN REWRITE_TAC[UNIONS_SUBSET] THEN
+ EXISTS_TAC `interval[a:real^N,b]` THEN
+ REWRITE_TAC[BOUNDED_INTERVAL] THEN ASM SET_TAC[]];
+ ALL_TAC] THEN
+ STRIP_TAC THEN ANTS_TAC THENL
+ [REPEAT CONJ_TAC THENL
+ [ASM_MESON_TAC[MEASURABLE_INTERVAL];
+ ASM_MESON_TAC[];
+ MATCH_MP_TAC BOUNDED_SUBSET THEN REWRITE_TAC[UNIONS_SUBSET] THEN
+ EXISTS_TAC `interval[a:real^N,b]` THEN
+ REWRITE_TAC[BOUNDED_INTERVAL] THEN ASM SET_TAC[]];
+ ALL_TAC] THEN
+ STRIP_TAC THEN REWRITE_TAC[GSYM IMAGE_o; o_DEF] THEN
+ SUBGOAL_THEN `m * gmeasure (\<Union>(IMAGE s (:num)):real^N->bool) =
+ measure(\<Union>(IMAGE (\<lambda>x. IMAGE f (s x)) (:num)):real^N->bool)`
+ (fun th -> ASM_REWRITE_TAC[GSYM HAS_GMEASURE_MEASURE; th]) THEN
+ ONCE_REWRITE_TAC[GSYM LIFT_EQ] THEN
+ MATCH_MP_TAC SERIES_UNIQUE THEN
+ EXISTS_TAC `\n:num. lift(measure(IMAGE (f:real^N->real^N) (s n)))` THEN
+ EXISTS_TAC `from 0` THEN ASM_REWRITE_TAC[] THEN MATCH_MP_TAC SUMS_EQ THEN
+ EXISTS_TAC `\n:num. m % lift(measure(s n:real^N->bool))` THEN
+ CONJ_TAC THENL
+ [REWRITE_TAC[GSYM LIFT_CMUL; LIFT_EQ] THEN
+ ASM_MESON_TAC[MEASURE_UNIQUE];
+ REWRITE_TAC[LIFT_CMUL] THEN MATCH_MP_TAC SERIES_CMUL THEN
+ ASM_REWRITE_TAC[]];
+ ALL_TAC] THEN
+ REWRITE_TAC[HAS_GMEASURE_INNER_OUTER_LE] THEN CONJ_TAC THEN
+ X_GEN_TAC `e:real` THEN DISCH_TAC THENL
+ [MP_TAC(ISPECL [`{a..b} DIFF s:real^N->bool`; `a:real^N`;
+ `b:real^N`; `e / (1 + abs m)`] GMEASURABLE_OUTER_INTERVALS_BOUNDED) THEN
+ ANTS_TAC THENL
+ [ASM_SIMP_TAC[MEASURABLE_DIFF; GMEASURABLE_INTERVAL] THEN
+ ASM_SIMP_TAC[REAL_ARITH `0 < 1 + abs x`; REAL_LT_DIV] THEN SET_TAC[];
+ ALL_TAC] THEN
+ DISCH_THEN(X_CHOOSE_THEN `d:(real^N->bool)->bool` STRIP_ASSUME_TAC) THEN
+ EXISTS_TAC `IMAGE f {a..b} DIFF
+ IMAGE (f:real^N->real^N) (\<Union>d)` THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `d:(real^N->bool)->bool`) THEN
+ ASM_SIMP_TAC[IMAGE_SUBSET] THEN DISCH_TAC THEN
+ CONJ_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN CONJ_TAC THENL
+ [ASM_MESON_TAC[MEASURABLE_DIFF; gmeasurable]; ALL_TAC] THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `measure(IMAGE f {a..b}) -
+ measure(IMAGE (f:real^N->real^N) (\<Union>d))` THEN
+ CONJ_TAC THENL
+ [ALL_TAC;
+ MATCH_MP_TAC REAL_EQ_IMP_LE THEN CONV_TAC SYM_CONV THEN
+ MATCH_MP_TAC MEASURE_DIFF_SUBSET THEN
+ REPEAT(CONJ_TAC THENL [ASM_MESON_TAC[measurable]; ALL_TAC]) THEN
+ MATCH_MP_TAC IMAGE_SUBSET THEN ASM_SIMP_TAC[UNIONS_SUBSET]] THEN
+ FIRST_ASSUM(ASSUME_TAC o SPECL [`a:real^N`; `b:real^N`]) THEN
+ REPEAT(FIRST_X_ASSUM(SUBST1_TAC o MATCH_MP MEASURE_UNIQUE)) THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `m * measure(s:real^N->bool) - m * e / (1 + abs m)` THEN
+ CONJ_TAC THENL
+ [REWRITE_TAC[REAL_ARITH `a - x \<le> a - y \<longleftrightarrow> y \<le> x`] THEN
+ REWRITE_TAC[real_div; REAL_MUL_ASSOC] THEN
+ REWRITE_TAC[GSYM real_div] THEN
+ ASM_SIMP_TAC[REAL_LE_LDIV_EQ; REAL_ARITH `0 < 1 + abs x`] THEN
+ GEN_REWRITE_TAC RAND_CONV [REAL_MUL_SYM] THEN
+ ASM_SIMP_TAC[REAL_LE_RMUL_EQ] THEN REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ REWRITE_TAC[GSYM REAL_SUB_LDISTRIB] THEN MATCH_MP_TAC REAL_LE_LMUL THEN
+ ASM_REWRITE_TAC[] THEN
+ FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REAL_ARITH
+ `d \<le> a + e ==> a = i - s ==> s - e \<le> i - d`)) THEN
+ MATCH_MP_TAC MEASURE_DIFF_SUBSET THEN
+ ASM_REWRITE_TAC[MEASURABLE_INTERVAL];
+ MP_TAC(ISPECL [`s:real^N->bool`; `a:real^N`; `b:real^N`;
+ `e / (1 + abs m)`] GMEASURABLE_OUTER_INTERVALS_BOUNDED) THEN
+ ASM_SIMP_TAC[REAL_LT_DIV; REAL_ARITH `0 < 1 + abs x`] THEN
+ DISCH_THEN(X_CHOOSE_THEN `d:(real^N->bool)->bool` STRIP_ASSUME_TAC) THEN
+ EXISTS_TAC `IMAGE (f:real^N->real^N) (\<Union>d)` THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `d:(real^N->bool)->bool`) THEN
+ ASM_SIMP_TAC[IMAGE_SUBSET] THEN
+ SIMP_TAC[HAS_GMEASURE_MEASURABLE_MEASURE] THEN STRIP_TAC THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN
+ EXISTS_TAC `m * measure(s:real^N->bool) + m * e / (1 + abs m)` THEN
+ CONJ_TAC THENL
+ [REWRITE_TAC[GSYM REAL_ADD_LDISTRIB] THEN ASM_SIMP_TAC[REAL_LE_LMUL];
+ REWRITE_TAC[REAL_LE_LADD] THEN
+ REWRITE_TAC[real_div; REAL_MUL_ASSOC] THEN
+ REWRITE_TAC[GSYM real_div] THEN
+ ASM_SIMP_TAC[REAL_LE_LDIV_EQ; REAL_ARITH `0 < 1 + abs x`] THEN
+ GEN_REWRITE_TAC RAND_CONV [REAL_MUL_SYM] THEN
+ ASM_SIMP_TAC[REAL_LE_RMUL_EQ] THEN REAL_ARITH_TAC]];
+ ALL_TAC] THEN
+ REPEAT STRIP_TAC THEN ONCE_REWRITE_TAC[HAS_GMEASURE_LIMIT] THEN
+ X_GEN_TAC `e:real` THEN DISCH_TAC THEN
+ FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [HAS_GMEASURE_MEASURE]) THEN
+ GEN_REWRITE_TAC LAND_CONV [HAS_GMEASURE_LIMIT] THEN
+ DISCH_THEN(MP_TAC o SPEC `e / (1 + abs m)`) THEN
+ ASM_SIMP_TAC[REAL_LT_DIV; REAL_ARITH `0 < 1 + abs x`] THEN
+ DISCH_THEN(X_CHOOSE_THEN `B:real`
+ (CONJUNCTS_THEN2 ASSUME_TAC (LABEL_TAC "*"))) THEN
+ MP_TAC(ISPEC `ball(0:real^N,B)` BOUNDED_SUBSET_CLOSED_INTERVAL) THEN
+ REWRITE_TAC[BOUNDED_BALL; LEFT_IMP_EXISTS_THM] THEN
+ REMOVE_THEN "*" MP_TAC THEN
+ MATCH_MP_TAC MONO_FORALL THEN X_GEN_TAC `c:real^N` THEN
+ MATCH_MP_TAC MONO_FORALL THEN X_GEN_TAC `d:real^N` THEN
+ DISCH_THEN(fun th -> DISCH_TAC THEN MP_TAC th) THEN ASM_REWRITE_TAC[] THEN
+ DISCH_THEN(X_CHOOSE_THEN `z:real` STRIP_ASSUME_TAC) THEN
+ MP_TAC(ISPECL [`interval[c:real^N,d]`; `0:real^N`]
+ BOUNDED_SUBSET_BALL) THEN
+ REWRITE_TAC[BOUNDED_INTERVAL] THEN
+ DISCH_THEN(X_CHOOSE_THEN `D:real` STRIP_ASSUME_TAC) THEN
+ MP_TAC(ISPEC `f:real^N->real^N` LINEAR_BOUNDED_POS) THEN
+ ASM_REWRITE_TAC[] THEN
+ DISCH_THEN(X_CHOOSE_THEN `C:real` STRIP_ASSUME_TAC) THEN
+
+ EXISTS_TAC `D * C` THEN ASM_SIMP_TAC[REAL_LT_MUL] THEN
+ MAP_EVERY X_GEN_TAC [`a:real^N`; `b:real^N`] THEN DISCH_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC
+ `s \<inter> (IMAGE (h:real^N->real^N) {a..b})`) THEN
+ SUBGOAL_THEN
+ `IMAGE (f:real^N->real^N) (s \<inter> IMAGE h (interval [a,b])) =
+ (IMAGE f s) \<inter> {a..b}`
+ SUBST1_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN ANTS_TAC THENL
+ [ASM_SIMP_TAC[BOUNDED_INTER; BOUNDED_LINEAR_IMAGE; BOUNDED_INTERVAL] THEN
+ ASM_SIMP_TAC[MEASURABLE_INTER; GMEASURABLE_LINEAR_IMAGE_INTERVAL];
+ ALL_TAC] THEN
+ DISCH_TAC THEN EXISTS_TAC
+ `m * measure(s \<inter> (IMAGE (h:real^N->real^N) {a..b}))` THEN
+ ASM_REWRITE_TAC[] THEN
+ MATCH_MP_TAC REAL_LET_TRANS THEN EXISTS_TAC `m * e / (1 + abs m)` THEN
+ CONJ_TAC THENL
+ [ALL_TAC;
+ REWRITE_TAC[real_div; REAL_MUL_ASSOC] THEN REWRITE_TAC[GSYM real_div] THEN
+ ASM_SIMP_TAC[REAL_LT_LDIV_EQ; REAL_ARITH `0 < 1 + abs x`] THEN
+ GEN_REWRITE_TAC RAND_CONV [REAL_MUL_SYM] THEN
+ ASM_SIMP_TAC[REAL_LT_RMUL_EQ] THEN REAL_ARITH_TAC] THEN
+ REWRITE_TAC[GSYM REAL_SUB_LDISTRIB; REAL_ABS_MUL] THEN
+ GEN_REWRITE_TAC (LAND_CONV o LAND_CONV) [real_abs] THEN
+ ASM_REWRITE_TAC[] THEN MATCH_MP_TAC REAL_LE_LMUL THEN ASM_REWRITE_TAC[] THEN
+ FIRST_ASSUM(MATCH_MP_TAC o MATCH_MP (REAL_ARITH
+ `abs(z - m) < e ==> z \<le> w \<and> w \<le> m ==> abs(w - m) \<le> e`)) THEN
+ SUBST1_TAC(SYM(MATCH_MP MEASURE_UNIQUE
+ (ASSUME `s \<inter> interval [c:real^N,d] has_gmeasure z`))) THEN
+ CONJ_TAC THEN MATCH_MP_TAC MEASURE_SUBSET THEN
+ ASM_SIMP_TAC[MEASURABLE_INTER; GMEASURABLE_LINEAR_IMAGE_INTERVAL;
+ GMEASURABLE_INTERVAL; INTER_SUBSET] THEN
+ MATCH_MP_TAC(SET_RULE
+ `!v. t \<subseteq> v \<and> v \<subseteq> u ==> s \<inter> t \<subseteq> s \<inter> u`) THEN
+ EXISTS_TAC `ball(0:real^N,D)` THEN ASM_REWRITE_TAC[] THEN
+ MATCH_MP_TAC(SET_RULE
+ `!f. (!x. h(f x) = x) \<and> IMAGE f s \<subseteq> t ==> s \<subseteq> IMAGE h t`) THEN
+ EXISTS_TAC `f:real^N->real^N` THEN ASM_REWRITE_TAC[] THEN
+ MATCH_MP_TAC SUBSET_TRANS THEN EXISTS_TAC `ball(0:real^N,D * C)` THEN
+ ASM_REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_BALL_0] THEN
+ X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+ MATCH_MP_TAC REAL_LET_TRANS THEN EXISTS_TAC `C * norm(x:real^N)` THEN
+ ASM_REWRITE_TAC[] THEN GEN_REWRITE_TAC RAND_CONV [REAL_MUL_SYM] THEN
+ ASM_SIMP_TAC[REAL_LT_LMUL_EQ]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Some inductions by expressing mapping in terms of elementary matrices. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma INDUCT_MATRIX_ROW_OPERATIONS: True .. (*
+ "!P:real^N^N->bool.
+ (!A i. 1 \<le> i \<and> i \<le> dimindex(:N) \<and> row i A = 0 ==> P A) \<and>
+ (!A. (!i j. 1 \<le> i \<and> i \<le> dimindex(:N) \<and>
+ 1 \<le> j \<and> j \<le> dimindex(:N) \<and> ~(i = j)
+ ==> A$i$j = 0) ==> P A) \<and>
+ (!A m n. P A \<and> 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and> ~(m = n)
+ ==> P(lambda i j. A$i$(swap(m,n) j))) \<and>
+ (!A m n c. P A \<and> 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and> ~(m = n)
+ ==> P(lambda i. if i = m then row m A + c % row n A
+ else row i A))
+ ==> !A. P A"
+qed GEN_TAC THEN
+ DISCH_THEN(CONJUNCTS_THEN2 (LABEL_TAC "zero_row") MP_TAC) THEN
+ DISCH_THEN(CONJUNCTS_THEN2 (LABEL_TAC "diagonal") MP_TAC) THEN
+ DISCH_THEN(CONJUNCTS_THEN2 (LABEL_TAC "swap_cols") (LABEL_TAC "row_op")) THEN
+ SUBGOAL_THEN
+ `!k A:real^N^N.
+ (!i j. 1 \<le> i \<and> i \<le> dimindex(:N) \<and>
+ k \<le> j \<and> j \<le> dimindex(:N) \<and> ~(i = j)
+ ==> A$i$j = 0)
+ ==> P A`
+ (fun th -> GEN_TAC THEN MATCH_MP_TAC th THEN
+ EXISTS_TAC `dimindex(:N) + 1` THEN ARITH_TAC) THEN
+ MATCH_MP_TAC num_INDUCTION THEN CONJ_TAC THENL
+ [REPEAT STRIP_TAC THEN USE_THEN "diagonal" MATCH_MP_TAC THEN
+ REPEAT STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN
+ ASM_REWRITE_TAC[LE_0];
+ ALL_TAC] THEN
+ X_GEN_TAC `k:num` THEN DISCH_THEN(LABEL_TAC "ind_hyp") THEN
+ DISJ_CASES_THEN2 SUBST1_TAC ASSUME_TAC (ARITH_RULE `k = 0 \/ 1 \<le> k`) THEN
+ ASM_REWRITE_TAC[ARITH] THEN
+ ASM_CASES_TAC `k \<le> dimindex(:N)` THENL
+ [ALL_TAC;
+ REPEAT STRIP_TAC THEN REMOVE_THEN "ind_hyp" MATCH_MP_TAC THEN
+ ASM_ARITH_TAC] THEN
+ SUBGOAL_THEN
+ `!A:real^N^N.
+ ~(A$k$k = 0) \<and>
+ (!i j. 1 \<le> i \<and> i \<le> dimindex (:N) \<and>
+ SUC k \<le> j \<and> j \<le> dimindex (:N) \<and> ~(i = j)
+ ==> A$i$j = 0)
+ ==> P A`
+ (LABEL_TAC "nonzero_hyp") THENL
+ [ALL_TAC;
+ X_GEN_TAC `A:real^N^N` THEN DISCH_TAC THEN
+ ASM_CASES_TAC `row k (A:real^N^N) = 0` THENL
+ [REMOVE_THEN "zero_row" MATCH_MP_TAC THEN ASM_MESON_TAC[];
+ ALL_TAC] THEN
+ FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE RAND_CONV [CART_EQ]) THEN
+ SIMP_TAC[VEC_COMPONENT; row; LAMBDA_BETA] THEN
+ REWRITE_TAC[NOT_FORALL_THM; NOT_IMP; LEFT_IMP_EXISTS_THM] THEN
+ X_GEN_TAC `l:num` THEN STRIP_TAC THEN
+ ASM_CASES_TAC `l:num = k` THENL
+ [REMOVE_THEN "nonzero_hyp" MATCH_MP_TAC THEN ASM_MESON_TAC[];
+ ALL_TAC] THEN
+ REMOVE_THEN "swap_cols" (MP_TAC o SPECL
+ [`(lambda i j. (A:real^N^N)$i$swap(k,l) j):real^N^N`;
+ `k:num`; `l:num`]) THEN
+ ASM_SIMP_TAC[LAMBDA_BETA] THEN ANTS_TAC THENL
+ [ALL_TAC;
+ MATCH_MP_TAC EQ_IMP THEN AP_TERM_TAC THEN
+ SIMP_TAC[CART_EQ; LAMBDA_BETA] THEN
+ REPEAT STRIP_TAC THEN REWRITE_TAC[swap] THEN
+ REPEAT(COND_CASES_TAC THEN ASM_SIMP_TAC[LAMBDA_BETA])] THEN
+ REMOVE_THEN "nonzero_hyp" MATCH_MP_TAC THEN
+ ONCE_REWRITE_TAC[ARITH_RULE `SUC k \<le> i \<longleftrightarrow> 1 \<le> i \<and> SUC k \<le> i`] THEN
+ ASM_SIMP_TAC[LAMBDA_BETA] THEN
+ ASM_REWRITE_TAC[swap] THEN MAP_EVERY X_GEN_TAC [`i:num`; `j:num`] THEN
+ STRIP_TAC THEN SUBGOAL_THEN `l:num \<le> k` ASSUME_TAC THENL
+ [FIRST_X_ASSUM(MP_TAC o SPECL [`k:num`; `l:num`]) THEN
+ ASM_REWRITE_TAC[] THEN ARITH_TAC;
+ ALL_TAC] THEN
+ REPEAT(COND_CASES_TAC THEN ASM_REWRITE_TAC[]) THEN
+ FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_REWRITE_TAC[] THEN
+ ASM_ARITH_TAC] THEN
+ SUBGOAL_THEN
+ `!l A:real^N^N.
+ ~(A$k$k = 0) \<and>
+ (!i j. 1 \<le> i \<and> i \<le> dimindex (:N) \<and>
+ SUC k \<le> j \<and> j \<le> dimindex (:N) \<and> ~(i = j)
+ ==> A$i$j = 0) \<and>
+ (!i. l \<le> i \<and> i \<le> dimindex(:N) \<and> ~(i = k) ==> A$i$k = 0)
+ ==> P A`
+ MP_TAC THENL
+ [ALL_TAC;
+ DISCH_THEN(MP_TAC o SPEC `dimindex(:N) + 1`) THEN
+ REWRITE_TAC[CONJ_ASSOC; ARITH_RULE `~(n + 1 \<le> i \<and> i \<le> n)`]] THEN
+ MATCH_MP_TAC num_INDUCTION THEN CONJ_TAC THENL
+ [GEN_TAC THEN DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC MP_TAC) THEN
+ DISCH_THEN(CONJUNCTS_THEN2 (LABEL_TAC "main") (LABEL_TAC "aux")) THEN
+ USE_THEN "ind_hyp" MATCH_MP_TAC THEN
+ MAP_EVERY X_GEN_TAC [`i:num`; `j:num`] THEN STRIP_TAC THEN
+ ASM_CASES_TAC `j:num = k` THENL
+ [ASM_REWRITE_TAC[] THEN USE_THEN "aux" MATCH_MP_TAC THEN ASM_ARITH_TAC;
+ REMOVE_THEN "main" MATCH_MP_TAC THEN ASM_ARITH_TAC];
+ ALL_TAC] THEN
+ X_GEN_TAC `l:num` THEN DISCH_THEN(LABEL_TAC "inner_hyp") THEN
+ GEN_TAC THEN DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC MP_TAC) THEN
+ DISCH_THEN(CONJUNCTS_THEN2 (LABEL_TAC "main") (LABEL_TAC "aux")) THEN
+ ASM_CASES_TAC `l:num = k` THENL
+ [REMOVE_THEN "inner_hyp" MATCH_MP_TAC THEN ASM_REWRITE_TAC[] THEN
+ REPEAT STRIP_TAC THEN REMOVE_THEN "aux" MATCH_MP_TAC THEN ASM_ARITH_TAC;
+ ALL_TAC] THEN
+ DISJ_CASES_TAC(ARITH_RULE `l = 0 \/ 1 \<le> l`) THENL
+ [REMOVE_THEN "ind_hyp" MATCH_MP_TAC THEN
+ MAP_EVERY X_GEN_TAC [`i:num`; `j:num`] THEN STRIP_TAC THEN
+ ASM_CASES_TAC `j:num = k` THENL
+ [ASM_REWRITE_TAC[] THEN REMOVE_THEN "aux" MATCH_MP_TAC THEN ASM_ARITH_TAC;
+ REMOVE_THEN "main" MATCH_MP_TAC THEN ASM_ARITH_TAC];
+ ALL_TAC] THEN
+ ASM_CASES_TAC `l \<le> dimindex(:N)` THENL
+ [ALL_TAC;
+ REMOVE_THEN "inner_hyp" MATCH_MP_TAC THEN ASM_REWRITE_TAC[] THEN
+ ASM_ARITH_TAC] THEN
+ REMOVE_THEN "inner_hyp" (MP_TAC o SPECL
+ [`(lambda i. if i = l then row l (A:real^N^N) + --(A$l$k/A$k$k) % row k A
+ else row i A):real^N^N`]) THEN
+ ANTS_TAC THENL
+ [SUBGOAL_THEN `!i. l \<le> i ==> 1 \<le> i` ASSUME_TAC THENL
+ [ASM_ARITH_TAC; ALL_TAC] THEN
+ ONCE_REWRITE_TAC[ARITH_RULE `SUC k \<le> j \<longleftrightarrow> 1 \<le> j \<and> SUC k \<le> j`] THEN
+ ASM_SIMP_TAC[LAMBDA_BETA; row; COND_COMPONENT;
+ VECTOR_ADD_COMPONENT; VECTOR_MUL_COMPONENT] THEN
+ ASM_SIMP_TAC[REAL_FIELD `~(y = 0) ==> x + --(x / y) * y = 0`] THEN
+ REWRITE_TAC[AND_FORALL_THM] THEN X_GEN_TAC `i:num` THEN
+ ASM_CASES_TAC `i:num = l` THEN ASM_REWRITE_TAC[] THENL
+ [REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC(REAL_RING `x = 0 \<and> y = 0 ==> x + z * y = 0`) THEN
+ CONJ_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_ARITH_TAC;
+ REPEAT STRIP_TAC THEN REMOVE_THEN "aux" MATCH_MP_TAC THEN ASM_ARITH_TAC];
+ ALL_TAC] THEN
+ DISCH_TAC THEN REMOVE_THEN "row_op" (MP_TAC o SPECL
+ [`(lambda i. if i = l then row l A + --(A$l$k / A$k$k) % row k A
+ else row i (A:real^N^N)):real^N^N`;
+ `l:num`; `k:num`; `(A:real^N^N)$l$k / A$k$k`]) THEN
+ ASM_REWRITE_TAC[] THEN MATCH_MP_TAC EQ_IMP THEN AP_TERM_TAC THEN
+ ASM_SIMP_TAC[CART_EQ; LAMBDA_BETA; VECTOR_ADD_COMPONENT;
+ VECTOR_MUL_COMPONENT; row; COND_COMPONENT] THEN
+ REPEAT STRIP_TAC THEN COND_CASES_TAC THEN ASM_REWRITE_TAC[] THEN
+ REAL_ARITH_TAC);; *)
+
+lemma INDUCT_MATRIX_ELEMENTARY: True .. (*
+ "!P:real^N^N->bool.
+ (!A B. P A \<and> P B ==> P(A ** B)) \<and>
+ (!A i. 1 \<le> i \<and> i \<le> dimindex(:N) \<and> row i A = 0 ==> P A) \<and>
+ (!A. (!i j. 1 \<le> i \<and> i \<le> dimindex(:N) \<and>
+ 1 \<le> j \<and> j \<le> dimindex(:N) \<and> ~(i = j)
+ ==> A$i$j = 0) ==> P A) \<and>
+ (!m n. 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and> ~(m = n)
+ ==> P(lambda i j. (mat 1:real^N^N)$i$(swap(m,n) j))) \<and>
+ (!m n c. 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and> ~(m = n)
+ ==> P(lambda i j. if i = m \<and> j = n then c
+ else if i = j then 1 else 0))
+ ==> !A. P A"
+qed GEN_TAC THEN
+ DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC MP_TAC) THEN
+ DISCH_THEN(fun th ->
+ MATCH_MP_TAC INDUCT_MATRIX_ROW_OPERATIONS THEN MP_TAC th) THEN
+ REPEAT(MATCH_MP_TAC MONO_AND THEN CONJ_TAC) THEN REWRITE_TAC[] THEN
+ DISCH_THEN(fun th -> X_GEN_TAC `A:real^N^N` THEN MP_TAC th) THEN
+ REPEAT(MATCH_MP_TAC MONO_FORALL THEN GEN_TAC) THEN
+ DISCH_THEN(fun th -> STRIP_TAC THEN MP_TAC th) THEN ASM_REWRITE_TAC[] THEN
+ UNDISCH_TAC `(P:real^N^N->bool) A` THENL
+ [REWRITE_TAC[GSYM IMP_CONJ]; REWRITE_TAC[GSYM IMP_CONJ_ALT]] THEN
+ DISCH_THEN(ANTE_RES_THEN MP_TAC) THEN MATCH_MP_TAC EQ_IMP THEN
+ AP_TERM_TAC THEN REWRITE_TAC[CART_EQ] THEN
+ X_GEN_TAC `i:num` THEN STRIP_TAC THEN
+ X_GEN_TAC `j:num` THEN STRIP_TAC THEN
+ ASM_SIMP_TAC[CART_EQ; LAMBDA_BETA; matrix_mul; row] THENL
+ [ASM_SIMP_TAC[mat; IN_DIMINDEX_SWAP; LAMBDA_BETA] THEN
+ ONCE_REWRITE_TAC[COND_RAND] THEN
+ SIMP_TAC[SUM_DELTA; REAL_MUL_RZERO; REAL_MUL_RID] THEN
+ COND_CASES_TAC THEN REWRITE_TAC[] THEN
+ RULE_ASSUM_TAC(REWRITE_RULE[swap; IN_NUMSEG]) THEN ASM_ARITH_TAC;
+ ALL_TAC] THEN
+ ASM_CASES_TAC `i:num = m` THEN ASM_REWRITE_TAC[] THENL
+ [ALL_TAC;
+ ONCE_REWRITE_TAC[COND_RAND] THEN ONCE_REWRITE_TAC[COND_RATOR] THEN
+ REWRITE_TAC[REAL_MUL_LZERO] THEN
+ GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [EQ_SYM_EQ] THEN
+ ASM_SIMP_TAC[SUM_DELTA; LAMBDA_BETA; IN_NUMSEG; REAL_MUL_LID]] THEN
+ ASM_SIMP_TAC[VECTOR_ADD_COMPONENT; VECTOR_MUL_COMPONENT; LAMBDA_BETA] THEN
+ MATCH_MP_TAC EQ_TRANS THEN
+ EXISTS_TAC
+ `sum {m,n} (\<lambda>k. (if k = n then c else if m = k then 1 else 0) *
+ (A:real^N^N)$k$j)` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC SUM_SUPERSET THEN
+ ASM_SIMP_TAC[SUBSET; IN_INSERT; NOT_IN_EMPTY; DE_MORGAN_THM;
+ IN_NUMSEG; REAL_MUL_LZERO] THEN
+ ASM_ARITH_TAC;
+ ASM_SIMP_TAC[SUM_CLAUSES; FINITE_RULES; IN_INSERT; NOT_IN_EMPTY] THEN
+ REAL_ARITH_TAC]);; *)
+
+lemma INDUCT_MATRIX_ELEMENTARY_ALT: True .. (*
+ "!P:real^N^N->bool.
+ (!A B. P A \<and> P B ==> P(A ** B)) \<and>
+ (!A i. 1 \<le> i \<and> i \<le> dimindex(:N) \<and> row i A = 0 ==> P A) \<and>
+ (!A. (!i j. 1 \<le> i \<and> i \<le> dimindex(:N) \<and>
+ 1 \<le> j \<and> j \<le> dimindex(:N) \<and> ~(i = j)
+ ==> A$i$j = 0) ==> P A) \<and>
+ (!m n. 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and> ~(m = n)
+ ==> P(lambda i j. (mat 1:real^N^N)$i$(swap(m,n) j))) \<and>
+ (!m n. 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and> ~(m = n)
+ ==> P(lambda i j. if i = m \<and> j = n \/ i = j then 1 else 0))
+ ==> !A. P A"
+qed GEN_TAC THEN STRIP_TAC THEN MATCH_MP_TAC INDUCT_MATRIX_ELEMENTARY THEN
+ ASM_REWRITE_TAC[] THEN REPEAT STRIP_TAC THEN
+ ASM_CASES_TAC `c = 0` THENL
+ [FIRST_X_ASSUM(fun th -> MATCH_MP_TAC th THEN
+ MAP_EVERY X_GEN_TAC [`i:num`; `j:num`]) THEN
+ ASM_SIMP_TAC[LAMBDA_BETA; COND_ID];
+ ALL_TAC] THEN
+ SUBGOAL_THEN
+ `(lambda i j. if i = m \<and> j = n then c else if i = j then 1 else 0) =
+ ((lambda i j. if i = j then if j = n then inv c else 1 else 0):real^N^N) **
+ ((lambda i j. if i = m \<and> j = n \/ i = j then 1 else 0):real^N^N) **
+ ((lambda i j. if i = j then if j = n then c else 1 else 0):real^N^N)`
+ SUBST1_TAC THENL
+ [ALL_TAC;
+ REPEAT(MATCH_MP_TAC(ASSUME `!A B:real^N^N. P A \<and> P B ==> P(A ** B)`) THEN
+ CONJ_TAC) THEN
+ ASM_SIMP_TAC[] THEN FIRST_X_ASSUM(fun th -> MATCH_MP_TAC th THEN
+ MAP_EVERY X_GEN_TAC [`i:num`; `j:num`]) THEN
+ ASM_SIMP_TAC[LAMBDA_BETA]] THEN
+ SIMP_TAC[CART_EQ; matrix_mul; LAMBDA_BETA] THEN
+ X_GEN_TAC `i:num` THEN STRIP_TAC THEN
+ X_GEN_TAC `j:num` THEN STRIP_TAC THEN
+ ASM_SIMP_TAC[SUM_DELTA; IN_NUMSEG; REAL_ARITH
+ `(if p then 1 else 0) * (if q then c else 0) =
+ if q then if p then c else 0 else 0`] THEN
+ REWRITE_TAC[REAL_ARITH
+ `(if p then x else 0) * y = (if p then x * y else 0)`] THEN
+ GEN_REWRITE_TAC (RAND_CONV o ONCE_DEPTH_CONV) [EQ_SYM_EQ] THEN
+ ASM_SIMP_TAC[SUM_DELTA; IN_NUMSEG] THEN
+ ASM_CASES_TAC `i:num = m` THEN ASM_REWRITE_TAC[] THEN
+ ASM_CASES_TAC `j:num = n` THEN ASM_REWRITE_TAC[REAL_MUL_LID; EQ_SYM_EQ] THEN
+ ASM_CASES_TAC `i:num = n` THEN
+ ASM_SIMP_TAC[REAL_MUL_LINV; REAL_MUL_LID; REAL_MUL_RZERO]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* The same thing in mapping form (might have been easier all along). *)
+(* ------------------------------------------------------------------------- *)
+
+lemma INDUCT_LINEAR_ELEMENTARY: True .. (*
+ "!P. (!f g. linear f \<and> linear g \<and> P f \<and> P g ==> P(f o g)) \<and>
+ (!f i. linear f \<and> 1 \<le> i \<and> i \<le> dimindex(:N) \<and> (!x. (f x)$i = 0)
+ ==> P f) \<and>
+ (!c. P(\<lambda>x. lambda i. c i * x$i)) \<and>
+ (!m n. 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and> ~(m = n)
+ ==> P(\<lambda>x. lambda i. x$swap(m,n) i)) \<and>
+ (!m n. 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and> ~(m = n)
+ ==> P(\<lambda>x. lambda i. if i = m then x$m + x$n else x$i))
+ ==> !f:real^N->real^N. linear f ==> P f"
+qed GEN_TAC THEN
+ MP_TAC(ISPEC `\A:real^N^N. P(\<lambda>x:real^N. A ** x):bool`
+ INDUCT_MATRIX_ELEMENTARY_ALT) THEN
+ REWRITE_TAC[] THEN MATCH_MP_TAC MONO_IMP THEN CONJ_TAC THENL
+ [ALL_TAC;
+ DISCH_TAC THEN X_GEN_TAC `f:real^N->real^N` THEN DISCH_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `matrix(f:real^N->real^N)`) THEN
+ ASM_SIMP_TAC[MATRIX_WORKS] THEN REWRITE_TAC[ETA_AX]] THEN
+ MATCH_MP_TAC MONO_AND THEN CONJ_TAC THENL
+ [DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`A:real^N^N`; `B:real^N^N`] THEN
+ STRIP_TAC THEN FIRST_X_ASSUM(MP_TAC o SPECL
+ [`\x:real^N. (A:real^N^N) ** x`; `\x:real^N. (B:real^N^N) ** x`]) THEN
+ ASM_REWRITE_TAC[MATRIX_VECTOR_MUL_LINEAR; o_DEF] THEN
+ REWRITE_TAC[MATRIX_VECTOR_MUL_ASSOC];
+ ALL_TAC] THEN
+ MATCH_MP_TAC MONO_AND THEN CONJ_TAC THENL
+ [DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`A:real^N^N`; `m:num`] THEN
+ STRIP_TAC THEN FIRST_X_ASSUM(MP_TAC o SPECL
+ [`\x:real^N. (A:real^N^N) ** x`; `m:num`]) THEN
+ ASM_REWRITE_TAC[MATRIX_VECTOR_MUL_LINEAR] THEN
+ DISCH_THEN MATCH_MP_TAC THEN
+ UNDISCH_TAC `row m (A:real^N^N) = 0` THEN
+ ASM_SIMP_TAC[CART_EQ; row; LAMBDA_BETA; VEC_COMPONENT; matrix_vector_mul;
+ REAL_MUL_LZERO; SUM_0];
+ ALL_TAC] THEN
+ MATCH_MP_TAC MONO_AND THEN CONJ_TAC THENL
+ [DISCH_TAC THEN X_GEN_TAC `A:real^N^N` THEN STRIP_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `\i. (A:real^N^N)$i$i`) THEN
+ MATCH_MP_TAC EQ_IMP THEN AP_TERM_TAC THEN
+ ASM_SIMP_TAC[CART_EQ; matrix_vector_mul; FUN_EQ_THM; LAMBDA_BETA] THEN
+ MAP_EVERY X_GEN_TAC [`x:real^N`; `i:num`] THEN STRIP_TAC THEN
+ MATCH_MP_TAC EQ_TRANS THEN EXISTS_TAC
+ `sum(1..dimindex(:N)) (\<lambda>j. if j = i then (A:real^N^N)$i$j * (x:real^N)$j
+ else 0)` THEN
+ CONJ_TAC THENL [ASM_SIMP_TAC[SUM_DELTA; IN_NUMSEG]; ALL_TAC] THEN
+ MATCH_MP_TAC SUM_EQ_NUMSEG THEN X_GEN_TAC `j:num` THEN STRIP_TAC THEN
+ ASM_SIMP_TAC[] THEN COND_CASES_TAC THEN ASM_SIMP_TAC[REAL_MUL_LZERO];
+ ALL_TAC] THEN
+ MATCH_MP_TAC MONO_AND THEN CONJ_TAC THEN
+ MATCH_MP_TAC MONO_FORALL THEN X_GEN_TAC `m:num` THEN
+ MATCH_MP_TAC MONO_FORALL THEN X_GEN_TAC `n:num` THEN
+ DISCH_THEN(fun th -> STRIP_TAC THEN MP_TAC th) THEN
+ ASM_REWRITE_TAC[] THEN MATCH_MP_TAC EQ_IMP THEN AP_TERM_TAC THEN
+ ASM_SIMP_TAC[CART_EQ; matrix_vector_mul; FUN_EQ_THM; LAMBDA_BETA;
+ mat; IN_DIMINDEX_SWAP]
+ THENL
+ [ONCE_REWRITE_TAC[SWAP_GALOIS] THEN ONCE_REWRITE_TAC[COND_RAND] THEN
+ ONCE_REWRITE_TAC[COND_RATOR] THEN
+ SIMP_TAC[SUM_DELTA; REAL_MUL_LID; REAL_MUL_LZERO; IN_NUMSEG] THEN
+ REPEAT STRIP_TAC THEN REWRITE_TAC[swap] THEN
+ COND_CASES_TAC THEN ASM_REWRITE_TAC[] THEN ASM_ARITH_TAC;
+ MAP_EVERY X_GEN_TAC [`x:real^N`; `i:num`] THEN STRIP_TAC THEN
+ ASM_CASES_TAC `i:num = m` THEN ASM_REWRITE_TAC[] THEN
+ ONCE_REWRITE_TAC[COND_RAND] THEN ONCE_REWRITE_TAC[COND_RATOR] THEN
+ GEN_REWRITE_TAC (RAND_CONV o ONCE_DEPTH_CONV) [EQ_SYM_EQ] THEN
+ ASM_SIMP_TAC[SUM_DELTA; REAL_MUL_LZERO; REAL_MUL_LID; IN_NUMSEG] THEN
+ MATCH_MP_TAC EQ_TRANS THEN EXISTS_TAC
+ `sum {m,n} (\<lambda>j. if n = j \/ j = m then (x:real^N)$j else 0)` THEN
+ CONJ_TAC THENL
+ [SIMP_TAC[SUM_CLAUSES; FINITE_RULES; IN_INSERT; NOT_IN_EMPTY] THEN
+ ASM_REWRITE_TAC[REAL_ADD_RID];
+ CONV_TAC SYM_CONV THEN MATCH_MP_TAC SUM_SUPERSET THEN
+ ASM_SIMP_TAC[SUBSET; IN_INSERT; NOT_IN_EMPTY; DE_MORGAN_THM;
+ IN_NUMSEG; REAL_MUL_LZERO] THEN
+ ASM_ARITH_TAC]]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Hence the effect of an arbitrary linear map on a gmeasurable set. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma LAMBDA_SWAP_GALOIS: True .. (*
+ "!x:real^N y:real^N.
+ 1 \<le> m \<and> m \<le> dimindex(:N) \<and> 1 \<le> n \<and> n \<le> dimindex(:N)
+ ==> (x = (lambda i. y$swap(m,n) i) \<longleftrightarrow>
+ (lambda i. x$swap(m,n) i) = y)"
+qed SIMP_TAC[CART_EQ; LAMBDA_BETA; IN_DIMINDEX_SWAP] THEN
+ REPEAT GEN_TAC THEN STRIP_TAC THEN EQ_TAC THEN
+ DISCH_TAC THEN X_GEN_TAC `i:num` THEN STRIP_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `swap(m,n) (i:num)`) THEN
+ ASM_SIMP_TAC[IN_DIMINDEX_SWAP] THEN
+ ASM_MESON_TAC[REWRITE_RULE[FUN_EQ_THM; o_THM; I_THM] SWAP_IDEMPOTENT]);; *)
+
+lemma LAMBDA_ADD_GALOIS: True .. (*
+ "!x:real^N y:real^N.
+ 1 \<le> m \<and> m \<le> dimindex(:N) \<and> 1 \<le> n \<and> n \<le> dimindex(:N) \<and>
+ ~(m = n)
+ ==> (x = (lambda i. if i = m then y$m + y$n else y$i) \<longleftrightarrow>
+ (lambda i. if i = m then x$m - x$n else x$i) = y)"
+qed SIMP_TAC[CART_EQ; LAMBDA_BETA] THEN
+ REPEAT GEN_TAC THEN STRIP_TAC THEN EQ_TAC THEN
+ DISCH_TAC THEN X_GEN_TAC `i:num` THEN STRIP_TAC THEN
+ FIRST_ASSUM(MP_TAC o SPEC `n:num`) THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `i:num`) THEN
+ ASM_REWRITE_TAC[] THEN COND_CASES_TAC THEN ASM_REWRITE_TAC[] THEN
+ REAL_ARITH_TAC);; *)
+
+lemma HAS_GMEASURE_SHEAR_INTERVAL: True .. (*
+ "!a b:real^N m n.
+ 1 \<le> m \<and> m \<le> dimindex(:N) \<and>
+ 1 \<le> n \<and> n \<le> dimindex(:N) \<and>
+ ~(m = n) \<and> ~({a..b} = {}) \<and>
+ 0 \<le> a$n
+ ==> (IMAGE (\<lambda>x. (lambda i. if i = m then x$m + x$n else x$i))
+ {a..b}:real^N->bool)
+ has_gmeasure gmeasure (interval [a,b])"
+qed lemma lemma = prove
+ (`!s t u v:real^N->bool.
+ gmeasurable s \<and> gmeasurable t \<and> gmeasurable u \<and>
+ negligible(s \<inter> t) \<and> negligible(s \<inter> u) \<and>
+ negligible(t \<inter> u) \<and>
+ s \<union> t \<union> u = v
+ ==> v has_gmeasure (measure s) + (measure t) + (measure u)"
+qed REPEAT STRIP_TAC THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_MEASURABLE_MEASURE; GMEASURABLE_UNION] THEN
+ FIRST_X_ASSUM(SUBST1_TAC o SYM) THEN
+ ASM_SIMP_TAC[MEASURE_UNION; GMEASURABLE_UNION] THEN
+ ASM_SIMP_TAC[MEASURE_EQ_0; UNION_OVER_INTER; MEASURE_UNION;
+ GMEASURABLE_UNION; NEGLIGIBLE_INTER; GMEASURABLE_INTER] THEN
+ REAL_ARITH_TAC)
+ and lemma' = prove
+ (`!s t u a.
+ gmeasurable s \<and> gmeasurable t \<and>
+ s \<union> (IMAGE (\<lambda>x. a + x) t) = u \<and>
+ negligible(s \<inter> (IMAGE (\<lambda>x. a + x) t))
+ ==> gmeasure s + gmeasure t = gmeasure u"
+qed REPEAT STRIP_TAC THEN FIRST_X_ASSUM(SUBST_ALL_TAC o SYM) THEN
+ ASM_SIMP_TAC[MEASURE_NEGLIGIBLE_UNION; GMEASURABLE_TRANSLATION;
+ MEASURE_TRANSLATION]) in
+ REWRITE_TAC[INTERVAL_NE_EMPTY] THEN REPEAT STRIP_TAC THEN
+ SUBGOAL_THEN
+ `linear((\<lambda>x. lambda i. if i = m then x$m + x$n else x$i):real^N->real^N)`
+ ASSUME_TAC THENL
+ [ASM_SIMP_TAC[linear; LAMBDA_BETA; VECTOR_ADD_COMPONENT;
+ VECTOR_MUL_COMPONENT; CART_EQ] THEN REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ MP_TAC(ISPECL
+ [`IMAGE (\<lambda>x. lambda i. if i = m then x$m + x$n else x$i)
+ (interval[a:real^N,b]):real^N->bool`;
+ `interval[a,(lambda i. if i = m then (b:real^N)$m + b$n else b$i)] INTER
+ {x:real^N | (basis m - basis n) dot x \<le> a$m}`;
+ `interval[a,(lambda i. if i = m then b$m + b$n else b$i)] INTER
+ {x:real^N | (basis m - basis n) dot x >= (b:real^N)$m}`;
+ `interval[a:real^N,
+ (lambda i. if i = m then (b:real^N)$m + b$n else b$i)]`]
+ lemma) THEN
+ ANTS_TAC THENL
+ [ASM_SIMP_TAC[CONVEX_LINEAR_IMAGE; CONVEX_INTERVAL;
+ CONVEX_HALFSPACE_LE; CONVEX_HALFSPACE_GE;
+ CONVEX_INTER; GMEASURABLE_CONVEX; BOUNDED_INTER;
+ BOUNDED_LINEAR_IMAGE; BOUNDED_INTERVAL] THEN
+ REWRITE_TAC[INTER] THEN
+ REWRITE_TAC[EXTENSION; IN_UNION; IN_INTER; IN_IMAGE] THEN
+ ASM_SIMP_TAC[LAMBDA_ADD_GALOIS; UNWIND_THM1] THEN
+ ASM_SIMP_TAC[IN_INTERVAL; IN_ELIM_THM; LAMBDA_BETA;
+ DOT_BASIS; DOT_LSUB] THEN
+ ONCE_REWRITE_TAC[MESON[]
+ `(!i:num. P i) \<longleftrightarrow> P m \<and> (!i. ~(i = m) ==> P i)`] THEN
+ ASM_SIMP_TAC[] THEN
+ REWRITE_TAC[TAUT `(p \<and> x) \<and> (q \<and> x) \<and> r \<longleftrightarrow> x \<and> p \<and> q \<and> r`;
+ TAUT `(p \<and> x) \<and> q \<and> (r \<and> x) \<longleftrightarrow> x \<and> p \<and> q \<and> r`;
+ TAUT `((p \<and> x) \<and> q) \<and> (r \<and> x) \<and> s \<longleftrightarrow>
+ x \<and> p \<and> q \<and> r \<and> s`;
+ TAUT `(a \<and> x \/ (b \<and> x) \<and> c \/ (d \<and> x) \<and> e \<longleftrightarrow> f \<and> x) \<longleftrightarrow>
+ x ==> (a \/ b \<and> c \/ d \<and> e \<longleftrightarrow> f)`] THEN
+ ONCE_REWRITE_TAC[SET_RULE
+ `{x | P x \<and> Q x} = {x | P x} \<inter> {x | Q x}`] THEN
+ REWRITE_TAC[CONJ_ASSOC] THEN CONJ_TAC THENL
+ [ALL_TAC;
+ GEN_TAC THEN DISCH_THEN(MP_TAC o SPEC `n:num`) THEN
+ ASM_REWRITE_TAC[] THEN ASM_REAL_ARITH_TAC] THEN
+ REWRITE_TAC[GSYM CONJ_ASSOC] THEN REPEAT CONJ_TAC THEN
+ MATCH_MP_TAC NEGLIGIBLE_INTER THEN DISJ2_TAC THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THENL
+ [EXISTS_TAC `{x:real^N | (basis m - basis n) dot x = (a:real^N)$m}`;
+ EXISTS_TAC `{x:real^N | (basis m - basis n) dot x = (b:real^N)$m}`;
+ EXISTS_TAC `{x:real^N | (basis m - basis n) dot x = (b:real^N)$m}`]
+ THEN (CONJ_TAC THENL
+ [MATCH_MP_TAC NEGLIGIBLE_HYPERPLANE THEN
+ REWRITE_TAC[VECTOR_SUB_EQ] THEN
+ ASM_MESON_TAC[BASIS_INJ];
+ ASM_SIMP_TAC[DOT_LSUB; DOT_BASIS; SUBSET; IN_ELIM_THM;
+ NOT_IN_EMPTY] THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `m:num`) THEN ASM_REWRITE_TAC[] THEN
+ ASM_REAL_ARITH_TAC]);
+ ALL_TAC] THEN
+ ASM_SIMP_TAC[HAS_GMEASURE_MEASURABLE_MEASURE;
+ GMEASURABLE_LINEAR_IMAGE_INTERVAL;
+ GMEASURABLE_INTERVAL] THEN
+ MP_TAC(ISPECL
+ [`interval[a,(lambda i. if i = m then (b:real^N)$m + b$n else b$i)] INTER
+ {x:real^N | (basis m - basis n) dot x \<le> a$m}`;
+ `interval[a,(lambda i. if i = m then b$m + b$n else b$i)] INTER
+ {x:real^N | (basis m - basis n) dot x >= (b:real^N)$m}`;
+ `interval[a:real^N,
+ (lambda i. if i = m then (a:real^N)$m + b$n
+ else (b:real^N)$i)]`;
+ `(lambda i. if i = m then (a:real^N)$m - (b:real^N)$m
+ else 0):real^N`]
+ lemma') THEN
+ ANTS_TAC THENL
+ [ASM_SIMP_TAC[CONVEX_LINEAR_IMAGE; CONVEX_INTERVAL;
+ CONVEX_HALFSPACE_LE; CONVEX_HALFSPACE_GE;
+ CONVEX_INTER; GMEASURABLE_CONVEX; BOUNDED_INTER;
+ BOUNDED_LINEAR_IMAGE; BOUNDED_INTERVAL] THEN
+ REWRITE_TAC[INTER] THEN
+ REWRITE_TAC[EXTENSION; IN_UNION; IN_INTER; IN_IMAGE] THEN
+ ONCE_REWRITE_TAC[VECTOR_ARITH `x:real^N = (lambda i. p i) + y \<longleftrightarrow>
+ x - (lambda i. p i) = y`] THEN
+ ASM_SIMP_TAC[IN_INTERVAL; IN_ELIM_THM; LAMBDA_BETA;
+ DOT_BASIS; DOT_LSUB; UNWIND_THM1;
+ VECTOR_SUB_COMPONENT] THEN
+ ONCE_REWRITE_TAC[MESON[]
+ `(!i:num. P i) \<longleftrightarrow> P m \<and> (!i. ~(i = m) ==> P i)`] THEN
+ ASM_SIMP_TAC[REAL_SUB_RZERO] THEN CONJ_TAC THENL
+ [X_GEN_TAC `x:real^N` THEN
+ FIRST_ASSUM(MP_TAC o SPEC `n:num`) THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `m:num`) THEN
+ ASM_REWRITE_TAC[] THEN
+ ASM_CASES_TAC
+ `!i. ~(i = m)
+ ==> 1 \<le> i \<and> i \<le> dimindex (:N)
+ ==> (a:real^N)$i \<le> (x:real^N)$i \<and>
+ x$i \<le> (b:real^N)$i` THEN
+ ASM_REWRITE_TAC[] THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `n:num`) THEN
+ ASM_REWRITE_TAC[] THEN ASM_REAL_ARITH_TAC;
+ ONCE_REWRITE_TAC[TAUT `((a \<and> b) \<and> c) \<and> (d \<and> e) \<and> f \<longleftrightarrow>
+ (b \<and> e) \<and> a \<and> c \<and> d \<and> f`] THEN
+ ONCE_REWRITE_TAC[SET_RULE
+ `{x | P x \<and> Q x} = {x | P x} \<inter> {x | Q x}`] THEN
+ MATCH_MP_TAC NEGLIGIBLE_INTER THEN DISJ2_TAC THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `{x:real^N | (basis m - basis n) dot x = (a:real^N)$m}`
+ THEN CONJ_TAC THENL
+ [MATCH_MP_TAC NEGLIGIBLE_HYPERPLANE THEN
+ REWRITE_TAC[VECTOR_SUB_EQ] THEN
+ ASM_MESON_TAC[BASIS_INJ];
+ ASM_SIMP_TAC[DOT_LSUB; DOT_BASIS; SUBSET; IN_ELIM_THM;
+ NOT_IN_EMPTY] THEN
+ FIRST_ASSUM(MP_TAC o SPEC `n:num`) THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `m:num`) THEN
+ ASM_REWRITE_TAC[] THEN ASM_REAL_ARITH_TAC]];
+ ALL_TAC] THEN
+ DISCH_THEN SUBST1_TAC THEN MATCH_MP_TAC(REAL_ARITH
+ `a = b + c ==> a = x + b ==> x = c`) THEN
+ ASM_SIMP_TAC[MEASURE_INTERVAL; CONTENT_CLOSED_INTERVAL_CASES;
+ LAMBDA_BETA] THEN
+ REPEAT(COND_CASES_TAC THENL
+ [ALL_TAC;
+ FIRST_X_ASSUM(MP_TAC o check (is_neg o concl)) THEN
+ MATCH_MP_TAC(TAUT `p ==> ~p ==> q`) THEN
+ X_GEN_TAC `i:num` THEN STRIP_TAC THEN
+ COND_CASES_TAC THEN ASM_SIMP_TAC[] THEN
+ FIRST_ASSUM(MP_TAC o SPEC `n:num`) THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `m:num`) THEN
+ ASM_REWRITE_TAC[] THEN ASM_REAL_ARITH_TAC]) THEN
+ SUBGOAL_THEN `1..dimindex(:N) = m INSERT ((1..dimindex(:N)) DELETE m)`
+ SUBST1_TAC THENL
+ [REWRITE_TAC[EXTENSION; IN_INSERT; IN_DELETE; IN_NUMSEG] THEN
+ ASM_ARITH_TAC;
+ ALL_TAC] THEN
+ SIMP_TAC[PRODUCT_CLAUSES; FINITE_DELETE; FINITE_NUMSEG] THEN
+ ASM_SIMP_TAC[IN_DELETE] THEN
+ MATCH_MP_TAC(REAL_RING
+ `s1 = s3 \<and> s2 = s3
+ ==> ((bm + bn) - am) * s1 =
+ ((am + bn) - am) * s2 + (bm - am) * s3`) THEN
+ CONJ_TAC THEN MATCH_MP_TAC PRODUCT_EQ THEN
+ SIMP_TAC[IN_DELETE] THEN REAL_ARITH_TAC);; *)
+
+lemma HAS_GMEASURE_LINEAR_IMAGE: True .. (*
+ "!f:real^N->real^N s.
+ linear f \<and> gmeasurable s
+ ==> (IMAGE f s) has_gmeasure (abs(det(matrix f)) * gmeasure s)"
+qed REWRITE_TAC[IMP_CONJ; RIGHT_FORALL_IMP_THM] THEN
+ MATCH_MP_TAC INDUCT_LINEAR_ELEMENTARY THEN REPEAT CONJ_TAC THENL
+ [MAP_EVERY X_GEN_TAC [`f:real^N->real^N`; `g:real^N->real^N`] THEN
+ REPLICATE_TAC 2 (DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC MP_TAC)) THEN
+ DISCH_THEN(fun th -> REPEAT STRIP_TAC THEN MP_TAC th) THEN
+ DISCH_THEN(CONJUNCTS_THEN2
+ (MP_TAC o SPEC `IMAGE (g:real^N->real^N) s`)
+ (MP_TAC o SPEC `s:real^N->bool`)) THEN
+ ASM_REWRITE_TAC[] THEN
+ GEN_REWRITE_TAC LAND_CONV [HAS_GMEASURE_MEASURABLE_MEASURE] THEN
+ STRIP_TAC THEN ASM_SIMP_TAC[MATRIX_COMPOSE; DET_MUL; REAL_ABS_MUL] THEN
+ REWRITE_TAC[IMAGE_o; GSYM REAL_MUL_ASSOC];
+
+ MAP_EVERY X_GEN_TAC [`f:real^N->real^N`; `m:num`] THEN STRIP_TAC THEN
+ SUBGOAL_THEN `~(!x y. (f:real^N->real^N) x = f y ==> x = y)`
+ ASSUME_TAC THENL
+ [ASM_SIMP_TAC[LINEAR_SINGULAR_INTO_HYPERPLANE] THEN
+ EXISTS_TAC `basis m:real^N` THEN
+ ASM_SIMP_TAC[BASIS_NONZERO; DOT_BASIS];
+ ALL_TAC] THEN
+ MP_TAC(ISPEC `matrix f:real^N^N` INVERTIBLE_DET_NZ) THEN
+ ASM_SIMP_TAC[INVERTIBLE_LEFT_INVERSE; MATRIX_LEFT_INVERTIBLE_INJECTIVE;
+ MATRIX_WORKS; REAL_ABS_NUM; REAL_MUL_LZERO] THEN
+ DISCH_THEN(K ALL_TAC) THEN REWRITE_TAC[HAS_GMEASURE_0] THEN
+ ASM_SIMP_TAC[NEGLIGIBLE_LINEAR_SINGULAR_IMAGE];
+
+ MAP_EVERY X_GEN_TAC [`c:num->real`; `s:real^N->bool`] THEN
+ DISCH_TAC THEN
+ FIRST_ASSUM(ASSUME_TAC o REWRITE_RULE[HAS_GMEASURE_MEASURE]) THEN
+ FIRST_ASSUM(MP_TAC o SPEC `c:num->real` o
+ MATCH_MP HAS_GMEASURE_STRETCH) THEN
+ MATCH_MP_TAC EQ_IMP THEN AP_TERM_TAC THEN
+ AP_THM_TAC THEN AP_TERM_TAC THEN AP_TERM_TAC THEN
+ SIMP_TAC[matrix; LAMBDA_BETA] THEN
+ W(MP_TAC o PART_MATCH (lhs o rand) DET_DIAGONAL o rand o snd) THEN
+ SIMP_TAC[LAMBDA_BETA; BASIS_COMPONENT; REAL_MUL_RZERO] THEN
+ DISCH_THEN(K ALL_TAC) THEN MATCH_MP_TAC PRODUCT_EQ_NUMSEG THEN
+ REWRITE_TAC[REAL_MUL_RID];
+
+ MAP_EVERY X_GEN_TAC [`m:num`; `n:num`] THEN STRIP_TAC THEN
+ MATCH_MP_TAC HAS_GMEASURE_LINEAR_SUFFICIENT THEN
+ ASM_SIMP_TAC[linear; LAMBDA_BETA; IN_DIMINDEX_SWAP; VECTOR_ADD_COMPONENT;
+ VECTOR_MUL_COMPONENT; CART_EQ] THEN
+ MAP_EVERY X_GEN_TAC [`a:real^N`; `b:real^N`] THEN
+ SUBGOAL_THEN `matrix (\<lambda>x:real^N. lambda i. x$swap (m,n) i):real^N^N =
+ transp(lambda i j. (mat 1:real^N^N)$i$swap (m,n) j)`
+ SUBST1_TAC THENL
+ [ASM_SIMP_TAC[MATRIX_EQ; LAMBDA_BETA; IN_DIMINDEX_SWAP;
+ matrix_vector_mul; CART_EQ; matrix; mat; basis;
+ COND_COMPONENT; transp] THEN
+ REWRITE_TAC[EQ_SYM_EQ];
+ ALL_TAC] THEN
+ REWRITE_TAC[DET_TRANSP] THEN
+ W(MP_TAC o PART_MATCH (lhs o rand) DET_PERMUTE_COLUMNS o
+ rand o lhand o rand o snd) THEN
+ ASM_SIMP_TAC[PERMUTES_SWAP; IN_NUMSEG; ETA_AX] THEN
+ DISCH_THEN(K ALL_TAC) THEN
+ REWRITE_TAC[DET_I; REAL_ABS_SIGN; REAL_MUL_RID; REAL_MUL_LID] THEN
+ ASM_CASES_TAC `interval[a:real^N,b] = {}` THENL
+ [ASM_SIMP_TAC[IMAGE_CLAUSES; HAS_GMEASURE_EMPTY; MEASURE_EMPTY];
+ ALL_TAC] THEN
+ SUBGOAL_THEN
+ `~(IMAGE (\<lambda>x:real^N. lambda i. x$swap (m,n) i)
+ {a..b}:real^N->bool = {})`
+ MP_TAC THENL [ASM_REWRITE_TAC[IMAGE_EQ_EMPTY]; ALL_TAC] THEN
+ SUBGOAL_THEN
+ `IMAGE (\<lambda>x:real^N. lambda i. x$swap (m,n) i)
+ {a..b}:real^N->bool =
+ interval[(lambda i. a$swap (m,n) i),
+ (lambda i. b$swap (m,n) i)]`
+ SUBST1_TAC THENL
+ [REWRITE_TAC[EXTENSION; IN_INTERVAL; IN_IMAGE] THEN
+ ASM_SIMP_TAC[LAMBDA_SWAP_GALOIS; UNWIND_THM1] THEN
+ SIMP_TAC[LAMBDA_BETA] THEN GEN_TAC THEN EQ_TAC THEN
+ DISCH_TAC THEN X_GEN_TAC `i:num` THEN STRIP_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `swap(m,n) (i:num)`) THEN
+ ASM_SIMP_TAC[IN_DIMINDEX_SWAP] THEN
+ ASM_MESON_TAC[REWRITE_RULE[FUN_EQ_THM; o_THM; I_THM] SWAP_IDEMPOTENT];
+ ALL_TAC] THEN
+ REWRITE_TAC[HAS_GMEASURE_MEASURABLE_MEASURE; GMEASURABLE_INTERVAL] THEN
+ REWRITE_TAC[MEASURE_INTERVAL] THEN
+ ASM_SIMP_TAC[CONTENT_CLOSED_INTERVAL; GSYM INTERVAL_NE_EMPTY] THEN
+ DISCH_THEN(K ALL_TAC) THEN SIMP_TAC[LAMBDA_BETA] THEN
+ ASM_SIMP_TAC[GSYM VECTOR_SUB_COMPONENT; IN_DIMINDEX_SWAP] THEN
+ MP_TAC(ISPECL [`\i. (b - a:real^N)$i`; `swap(m:num,n)`; `1..dimindex(:N)`]
+ (GSYM PRODUCT_PERMUTE)) THEN
+ REWRITE_TAC[o_DEF] THEN DISCH_THEN MATCH_MP_TAC THEN
+ ASM_SIMP_TAC[PERMUTES_SWAP; IN_NUMSEG];
+
+ MAP_EVERY X_GEN_TAC [`m:num`; `n:num`] THEN STRIP_TAC THEN
+ MATCH_MP_TAC HAS_GMEASURE_LINEAR_SUFFICIENT THEN
+ MATCH_MP_TAC(TAUT `a \<and> (a ==> b) ==> a \<and> b`) THEN CONJ_TAC THENL
+ [ASM_SIMP_TAC[linear; LAMBDA_BETA; VECTOR_ADD_COMPONENT;
+ VECTOR_MUL_COMPONENT; CART_EQ] THEN REAL_ARITH_TAC;
+ DISCH_TAC] THEN
+ MAP_EVERY X_GEN_TAC [`a:real^N`; `b:real^N`] THEN
+ SUBGOAL_THEN
+ `det(matrix(\<lambda>x. lambda i. if i = m then (x:real^N)$m + x$n
+ else x$i):real^N^N) = 1`
+ SUBST1_TAC THENL
+ [ASM_SIMP_TAC[matrix; basis; COND_COMPONENT; LAMBDA_BETA] THEN
+ FIRST_ASSUM(DISJ_CASES_TAC o MATCH_MP (ARITH_RULE
+ `~(m:num = n) ==> m < n \/ n < m`))
+ THENL
+ [W(MP_TAC o PART_MATCH (lhs o rand) DET_UPPERTRIANGULAR o lhs o snd);
+ W(MP_TAC o PART_MATCH (lhs o rand) DET_LOWERTRIANGULAR o lhs o snd)]
+ THEN ASM_SIMP_TAC[LAMBDA_BETA; BASIS_COMPONENT; COND_COMPONENT;
+ matrix; REAL_ADD_RID; COND_ID;
+ PRODUCT_CONST_NUMSEG; REAL_POW_ONE] THEN
+ DISCH_THEN MATCH_MP_TAC THEN
+ REPEAT GEN_TAC THEN REPEAT(COND_CASES_TAC THEN ASM_REWRITE_TAC[]) THEN
+ ASM_ARITH_TAC;
+ ALL_TAC] THEN
+ REWRITE_TAC[REAL_ABS_NUM; REAL_MUL_LID] THEN
+ ASM_CASES_TAC `interval[a:real^N,b] = {}` THENL
+ [ASM_SIMP_TAC[IMAGE_CLAUSES; HAS_GMEASURE_EMPTY; MEASURE_EMPTY];
+ ALL_TAC] THEN
+ SUBGOAL_THEN
+ `IMAGE (\<lambda>x. lambda i. if i = m then x$m + x$n else x$i) (interval [a,b]) =
+ IMAGE (\<lambda>x:real^N. (lambda i. if i = m \/ i = n then a$n else 0) +
+ x)
+ (IMAGE (\<lambda>x:real^N. lambda i. if i = m then x$m + x$n else x$i)
+ (IMAGE (\<lambda>x. (lambda i. if i = n then --(a$n) else 0) + x)
+ {a..b}))`
+ SUBST1_TAC THENL
+ [REWRITE_TAC[GSYM IMAGE_o] THEN AP_THM_TAC THEN AP_TERM_TAC THEN
+ ASM_SIMP_TAC[FUN_EQ_THM; o_THM; VECTOR_ADD_COMPONENT; LAMBDA_BETA;
+ CART_EQ] THEN
+ MAP_EVERY X_GEN_TAC [`x:real^N`; `i:num`] THEN
+ STRIP_TAC THEN ASM_CASES_TAC `i:num = m` THEN ASM_REWRITE_TAC[] THEN
+ ASM_CASES_TAC `i:num = n` THEN ASM_REWRITE_TAC[] THEN REAL_ARITH_TAC;
+ ALL_TAC] THEN
+ MATCH_MP_TAC HAS_GMEASURE_TRANSLATION THEN
+ SUBGOAL_THEN
+ `measure{a..b} =
+ measure(IMAGE (\<lambda>x:real^N. (lambda i. if i = n then --(a$n) else 0) + x)
+ {a..b}:real^N->bool)`
+ SUBST1_TAC THENL
+ [CONV_TAC SYM_CONV THEN MATCH_MP_TAC MEASURE_TRANSLATION THEN
+ REWRITE_TAC[MEASURABLE_INTERVAL];
+ ALL_TAC] THEN
+ SUBGOAL_THEN
+ `~(IMAGE (\<lambda>x:real^N. (lambda i. if i = n then --(a$n) else 0) + x)
+ {a..b}:real^N->bool = {})`
+ MP_TAC THENL [ASM_SIMP_TAC[IMAGE_EQ_EMPTY]; ALL_TAC] THEN
+ ONCE_REWRITE_TAC[VECTOR_ARITH `c + x = 1 % x + c`] THEN
+ ASM_REWRITE_TAC[IMAGE_AFFINITY_INTERVAL; REAL_POS] THEN
+ DISCH_TAC THEN MATCH_MP_TAC HAS_GMEASURE_SHEAR_INTERVAL THEN
+ ASM_SIMP_TAC[LAMBDA_BETA; VECTOR_ADD_COMPONENT; VECTOR_MUL_COMPONENT] THEN
+ REAL_ARITH_TAC]);; *)
+
+lemma GMEASURABLE_LINEAR_IMAGE: True .. (*
+ "!f:real^N->real^N s.
+ linear f \<and> gmeasurable s ==> gmeasurable(IMAGE f s)"
+qed REPEAT GEN_TAC THEN
+ DISCH_THEN(MP_TAC o MATCH_MP HAS_GMEASURE_LINEAR_IMAGE) THEN
+ SIMP_TAC[HAS_GMEASURE_MEASURABLE_MEASURE]);; *)
+
+lemma MEASURE_LINEAR_IMAGE: True .. (*
+ "!f:real^N->real^N s.
+ linear f \<and> gmeasurable s
+ ==> measure(IMAGE f s) = abs(det(matrix f)) * gmeasure s"
+qed REPEAT GEN_TAC THEN
+ DISCH_THEN(MP_TAC o MATCH_MP HAS_GMEASURE_LINEAR_IMAGE) THEN
+ SIMP_TAC[HAS_GMEASURE_MEASURABLE_MEASURE]);; *)
+
+lemma HAS_GMEASURE_LINEAR_IMAGE_SAME: True .. (*
+ "!f s. linear f \<and> gmeasurable s \<and> abs(det(matrix f)) = 1
+ ==> (IMAGE f s) has_gmeasure (measure s)"
+qed MESON_TAC[HAS_GMEASURE_LINEAR_IMAGE; REAL_MUL_LID]);; *)
+
+lemma MEASURE_LINEAR_IMAGE_SAME: True .. (*
+ "!f:real^N->real^N s.
+ linear f \<and> gmeasurable s \<and> abs(det(matrix f)) = 1
+ ==> measure(IMAGE f s) = gmeasure s"
+qed REPEAT GEN_TAC THEN
+ DISCH_THEN(MP_TAC o MATCH_MP HAS_GMEASURE_LINEAR_IMAGE_SAME) THEN
+ SIMP_TAC[HAS_GMEASURE_MEASURABLE_MEASURE]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* gmeasure of a standard simplex. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma CONGRUENT_IMAGE_STD_SIMPLEX: True .. (*
+ "!p. p permutes 1..dimindex(:N)
+ ==> {x:real^N | 0 \<le> x$(p 1) \<and> x$(p(dimindex(:N))) \<le> 1 \<and>
+ (!i. 1 \<le> i \<and> i < dimindex(:N)
+ ==> x$(p i) \<le> x$(p(i + 1)))} =
+ IMAGE (\<lambda>x:real^N. lambda i. sum(1..inverse p(i)) (\<lambda>j. x$j))
+ {x | (!i. 1 \<le> i \<and> i \<le> dimindex (:N) ==> 0 \<le> x$i) \<and>
+ sum (1..dimindex (:N)) (\<lambda>i. x$i) \<le> 1}"
+qed REPEAT STRIP_TAC THEN MATCH_MP_TAC SUBSET_ANTISYM THEN CONJ_TAC THENL
+ [ALL_TAC;
+ REWRITE_TAC[SUBSET; FORALL_IN_IMAGE] THEN X_GEN_TAC `x:real^N` THEN
+ ASM_SIMP_TAC[IN_ELIM_THM; LAMBDA_BETA; LAMBDA_BETA_PERM; LE_REFL;
+ ARITH_RULE `i < n ==> i \<le> n \<and> i + 1 \<le> n`;
+ ARITH_RULE `1 \<le> n + 1`; DIMINDEX_GE_1] THEN
+ STRIP_TAC THEN
+ FIRST_ASSUM(fun th -> REWRITE_TAC[MATCH_MP PERMUTES_INVERSES th]) THEN
+ ASM_SIMP_TAC[SUM_SING_NUMSEG; DIMINDEX_GE_1; LE_REFL] THEN
+ REWRITE_TAC[GSYM ADD1; SUM_CLAUSES_NUMSEG; ARITH_RULE `1 \<le> SUC n`] THEN
+ ASM_SIMP_TAC[REAL_LE_ADDR] THEN REPEAT STRIP_TAC THEN
+ FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_ARITH_TAC] THEN
+ REWRITE_TAC[SUBSET; IN_IMAGE; IN_ELIM_THM] THEN X_GEN_TAC `x:real^N` THEN
+ STRIP_TAC THEN
+ EXISTS_TAC `(lambda i. if i = 1 then x$(p 1)
+ else (x:real^N)$p(i) - x$p(i - 1)):real^N` THEN
+ ASM_SIMP_TAC[IN_ELIM_THM; LAMBDA_BETA; LAMBDA_BETA_PERM; LE_REFL;
+ ARITH_RULE `i < n ==> i \<le> n \<and> i + 1 \<le> n`;
+ ARITH_RULE `1 \<le> n + 1`; DIMINDEX_GE_1; CART_EQ] THEN
+ REPEAT CONJ_TAC THENL
+ [X_GEN_TAC `i:num` THEN STRIP_TAC THEN
+ SUBGOAL_THEN `1 \<le> inverse (p:num->num) i \<and>
+ !x. x \<le> inverse p i ==> x \<le> dimindex(:N)`
+ ASSUME_TAC THENL
+ [ASM_MESON_TAC[PERMUTES_INVERSE; IN_NUMSEG; LE_TRANS; PERMUTES_IN_IMAGE];
+ ASM_SIMP_TAC[LAMBDA_BETA] THEN ASM_SIMP_TAC[SUM_CLAUSES_LEFT; ARITH]] THEN
+ SIMP_TAC[ARITH_RULE `2 \<le> n ==> ~(n = 1)`] THEN
+ GEN_REWRITE_TAC (RAND_CONV o RAND_CONV o BINDER_CONV)
+ [GSYM REAL_MUL_LID] THEN
+ ONCE_REWRITE_TAC[SUM_PARTIAL_PRE] THEN
+ REWRITE_TAC[REAL_SUB_REFL; REAL_MUL_RZERO; SUM_0; COND_ID] THEN
+ REWRITE_TAC[REAL_MUL_LID; ARITH; REAL_SUB_RZERO] THEN
+ FIRST_ASSUM(DISJ_CASES_TAC o MATCH_MP (ARITH_RULE
+ `1 \<le> p ==> p = 1 \/ 2 \<le> p`) o CONJUNCT1) THEN
+ ASM_SIMP_TAC[ARITH] THEN
+ FIRST_ASSUM(fun th -> REWRITE_TAC[MATCH_MP PERMUTES_INVERSES th]) THEN
+ REWRITE_TAC[REAL_ADD_RID] THEN TRY REAL_ARITH_TAC THEN
+ ASM_MESON_TAC[PERMUTES_INVERSE_EQ; PERMUTES_INVERSE];
+
+ X_GEN_TAC `i:num` THEN STRIP_TAC THEN COND_CASES_TAC THEN
+ ASM_REWRITE_TAC[REAL_SUB_LE] THEN
+ FIRST_X_ASSUM(MP_TAC o SPEC `i - 1`) THEN
+ ASM_SIMP_TAC[SUB_ADD] THEN DISCH_THEN MATCH_MP_TAC THEN ASM_ARITH_TAC;
+
+ SIMP_TAC[SUM_CLAUSES_LEFT; DIMINDEX_GE_1; ARITH;
+ ARITH_RULE `2 \<le> n ==> ~(n = 1)`] THEN
+ GEN_REWRITE_TAC (LAND_CONV o RAND_CONV o RAND_CONV o BINDER_CONV)
+ [GSYM REAL_MUL_LID] THEN
+ ONCE_REWRITE_TAC[SUM_PARTIAL_PRE] THEN
+ REWRITE_TAC[REAL_SUB_REFL; REAL_MUL_RZERO; SUM_0; COND_ID] THEN
+ REWRITE_TAC[REAL_MUL_LID; ARITH; REAL_SUB_RZERO] THEN
+ COND_CASES_TAC THEN ASM_SIMP_TAC[REAL_ADD_RID] THEN
+ ASM_REWRITE_TAC[REAL_ARITH `x + y - x:real = y`] THEN
+ ASM_MESON_TAC[DIMINDEX_GE_1;
+ ARITH_RULE `1 \<le> n \<and> ~(2 \<le> n) ==> n = 1`]]);; *)
+
+lemma HAS_GMEASURE_IMAGE_STD_SIMPLEX: True .. (*
+ "!p. p permutes 1..dimindex(:N)
+ ==> {x:real^N | 0 \<le> x$(p 1) \<and> x$(p(dimindex(:N))) \<le> 1 \<and>
+ (!i. 1 \<le> i \<and> i < dimindex(:N)
+ ==> x$(p i) \<le> x$(p(i + 1)))}
+ has_gmeasure
+ (measure (convex hull
+ (0 INSERT {basis i:real^N | 1 \<le> i \<and> i \<le> dimindex(:N)})))"
+qed REPEAT STRIP_TAC THEN ASM_SIMP_TAC[CONGRUENT_IMAGE_STD_SIMPLEX] THEN
+ ASM_SIMP_TAC[GSYM STD_SIMPLEX] THEN
+ MATCH_MP_TAC HAS_GMEASURE_LINEAR_IMAGE_SAME THEN
+ REPEAT CONJ_TAC THENL
+ [REWRITE_TAC[linear; CART_EQ] THEN
+ ASM_SIMP_TAC[LAMBDA_BETA; VECTOR_ADD_COMPONENT; VECTOR_MUL_COMPONENT;
+ GSYM SUM_ADD_NUMSEG; GSYM SUM_LMUL] THEN
+ REPEAT STRIP_TAC THEN MATCH_MP_TAC SUM_EQ_NUMSEG THEN
+ REPEAT STRIP_TAC THEN REWRITE_TAC[] THENL
+ [MATCH_MP_TAC VECTOR_ADD_COMPONENT;
+ MATCH_MP_TAC VECTOR_MUL_COMPONENT] THEN
+ ASM_MESON_TAC[PERMUTES_INVERSE; IN_NUMSEG; LE_TRANS; PERMUTES_IN_IMAGE];
+ MATCH_MP_TAC GMEASURABLE_CONVEX THEN REWRITE_TAC[CONVEX_CONVEX_HULL] THEN
+ MATCH_MP_TAC BOUNDED_CONVEX_HULL THEN REWRITE_TAC[BOUNDED_INSERT] THEN
+ ONCE_REWRITE_TAC[SIMPLE_IMAGE_GEN] THEN
+ MATCH_MP_TAC FINITE_IMP_BOUNDED THEN MATCH_MP_TAC FINITE_IMAGE THEN
+ REWRITE_TAC[GSYM numseg; FINITE_NUMSEG];
+ MATCH_MP_TAC EQ_TRANS THEN EXISTS_TAC
+ `abs(det
+ ((lambda i. ((lambda i j. if j \<le> i then 1 else 0):real^N^N)
+ $inverse p i)
+ :real^N^N))` THEN
+ CONJ_TAC THENL
+ [AP_TERM_TAC THEN AP_TERM_TAC THEN REWRITE_TAC[CART_EQ] THEN
+ ASM_SIMP_TAC[matrix; LAMBDA_BETA; BASIS_COMPONENT; COND_COMPONENT;
+ LAMBDA_BETA_PERM; PERMUTES_INVERSE] THEN
+ X_GEN_TAC `i:num` THEN STRIP_TAC THEN
+ X_GEN_TAC `j:num` THEN STRIP_TAC THEN MATCH_MP_TAC EQ_TRANS THEN
+ EXISTS_TAC `sum (1..inverse (p:num->num) i)
+ (\<lambda>k. if k = j then 1 else 0)` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC SUM_EQ THEN
+ ASM_SIMP_TAC[IN_NUMSEG; PERMUTES_IN_IMAGE; basis] THEN
+ REPEAT STRIP_TAC THEN MATCH_MP_TAC LAMBDA_BETA THEN
+ ASM_MESON_TAC[PERMUTES_IN_IMAGE; IN_NUMSEG; LE_TRANS;
+ PERMUTES_INVERSE];
+ ASM_SIMP_TAC[SUM_DELTA; IN_NUMSEG]];
+ ALL_TAC] THEN
+ ASM_SIMP_TAC[PERMUTES_INVERSE; DET_PERMUTE_ROWS; ETA_AX] THEN
+ REWRITE_TAC[REAL_ABS_MUL; REAL_ABS_SIGN; REAL_MUL_LID] THEN
+ MATCH_MP_TAC(REAL_ARITH `x = 1 ==> abs x = 1`) THEN
+ ASM_SIMP_TAC[DET_LOWERTRIANGULAR; GSYM NOT_LT; LAMBDA_BETA] THEN
+ REWRITE_TAC[LT_REFL; PRODUCT_CONST_NUMSEG; REAL_POW_ONE]]);; *)
+
+lemma HAS_GMEASURE_STD_SIMPLEX: True .. (*
+ "(convex hull (0:real^N INSERT {basis i | 1 \<le> i \<and> i \<le> dimindex(:N)}))
+ has_gmeasure inv((FACT(dimindex(:N))))"
+qed lemma lemma = prove
+ (`!f:num->real. (!i. 1 \<le> i \<and> i < n ==> f i \<le> f(i + 1)) \<longleftrightarrow>
+ (!i j. 1 \<le> i \<and> i \<le> j \<and> j \<le> n ==> f i \<le> f j)"
+qed GEN_TAC THEN EQ_TAC THEN DISCH_TAC THENL
+ [GEN_TAC THEN INDUCT_TAC THEN
+ SIMP_TAC[LE; REAL_LE_REFL] THEN
+ STRIP_TAC THEN ASM_SIMP_TAC[REAL_LE_REFL] THEN
+ MATCH_MP_TAC REAL_LE_TRANS THEN EXISTS_TAC `(f:num->real) j` THEN
+ ASM_SIMP_TAC[ARITH_RULE `SUC x \<le> y ==> x \<le> y`] THEN
+ REWRITE_TAC[ADD1] THEN FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_ARITH_TAC;
+ REPEAT STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN ASM_ARITH_TAC]) in
+ MP_TAC(ISPECL
+ [`\p. {x:real^N | 0 \<le> x$(p 1) \<and> x$(p(dimindex(:N))) \<le> 1 \<and>
+ (!i. 1 \<le> i \<and> i < dimindex(:N)
+ ==> x$(p i) \<le> x$(p(i + 1)))}`;
+ `{p | p permutes 1..dimindex(:N)}`]
+ HAS_GMEASURE_NEGLIGIBLE_UNIONS_IMAGE) THEN
+ ASM_SIMP_TAC[REWRITE_RULE[HAS_GMEASURE_MEASURABLE_MEASURE]
+ HAS_GMEASURE_IMAGE_STD_SIMPLEX; IN_ELIM_THM] THEN
+ ASM_SIMP_TAC[SUM_CONST; FINITE_PERMUTATIONS; FINITE_NUMSEG;
+ CARD_PERMUTATIONS; CARD_NUMSEG_1] THEN
+ ANTS_TAC THENL
+ [MAP_EVERY X_GEN_TAC [`p:num->num`; `q:num->num`] THEN STRIP_TAC THEN
+ SUBGOAL_THEN `?i. i \<in> 1..dimindex(:N) \<and> ~(p i:num = q i)` MP_TAC THENL
+ [ASM_MESON_TAC[permutes; FUN_EQ_THM]; ALL_TAC] THEN
+ GEN_REWRITE_TAC LAND_CONV [num_WOP] THEN
+ REWRITE_TAC[TAUT `a ==> ~(b \<and> ~c) \<longleftrightarrow> a \<and> b ==> c`] THEN
+ REWRITE_TAC[IN_NUMSEG] THEN
+ DISCH_THEN(X_CHOOSE_THEN `k:num` STRIP_ASSUME_TAC) THEN
+ MATCH_MP_TAC NEGLIGIBLE_SUBSET THEN
+ EXISTS_TAC `{x:real^N | (basis(p(k:num)) - basis(q k)) dot x = 0}` THEN
+ CONJ_TAC THENL
+ [MATCH_MP_TAC NEGLIGIBLE_HYPERPLANE THEN REWRITE_TAC[VECTOR_SUB_EQ] THEN
+ MATCH_MP_TAC BASIS_NE THEN ASM_MESON_TAC[PERMUTES_IN_IMAGE; IN_NUMSEG];
+ ALL_TAC] THEN
+ REWRITE_TAC[SUBSET; IN_INTER; IN_ELIM_THM; DOT_LSUB; VECTOR_SUB_EQ] THEN
+ ASM_SIMP_TAC[DOT_BASIS; GSYM IN_NUMSEG; PERMUTES_IN_IMAGE] THEN
+ SUBGOAL_THEN `?l. (q:num->num) l = p(k:num)` STRIP_ASSUME_TAC THENL
+ [ASM_MESON_TAC[permutes]; ALL_TAC] THEN
+ SUBGOAL_THEN `1 \<le> l \<and> l \<le> dimindex(:N)` STRIP_ASSUME_TAC THENL
+ [ASM_MESON_TAC[PERMUTES_IN_IMAGE; IN_NUMSEG]; ALL_TAC] THEN
+ SUBGOAL_THEN `k:num < l` ASSUME_TAC THENL
+ [REWRITE_TAC[GSYM NOT_LE] THEN REWRITE_TAC[LE_LT] THEN
+ ASM_MESON_TAC[PERMUTES_INJECTIVE; IN_NUMSEG];
+ ALL_TAC] THEN
+ SUBGOAL_THEN `?m. (p:num->num) m = q(k:num)` STRIP_ASSUME_TAC THENL
+ [ASM_MESON_TAC[permutes]; ALL_TAC] THEN
+ SUBGOAL_THEN `1 \<le> m \<and> m \<le> dimindex(:N)` STRIP_ASSUME_TAC THENL
+ [ASM_MESON_TAC[PERMUTES_IN_IMAGE; IN_NUMSEG]; ALL_TAC] THEN
+ SUBGOAL_THEN `k:num < m` ASSUME_TAC THENL
+ [REWRITE_TAC[GSYM NOT_LE] THEN REWRITE_TAC[LE_LT] THEN
+ ASM_MESON_TAC[PERMUTES_INJECTIVE; IN_NUMSEG];
+ ALL_TAC] THEN
+ X_GEN_TAC `x:real^N` THEN REWRITE_TAC[lemma] THEN STRIP_TAC THEN
+ FIRST_X_ASSUM(MP_TAC o SPECL [`k:num`; `l:num`]) THEN
+ FIRST_X_ASSUM(MP_TAC o SPECL [`k:num`; `m:num`]) THEN
+ ASM_SIMP_TAC[LT_IMP_LE; IMP_IMP; REAL_LE_ANTISYM; REAL_SUB_0] THEN
+ MATCH_MP_TAC EQ_IMP THEN BINOP_TAC THEN
+ ASM_MESON_TAC[PERMUTES_IN_IMAGE; IN_NUMSEG; DOT_BASIS];
+ ALL_TAC] THEN
+ REWRITE_TAC[HAS_GMEASURE_MEASURABLE_MEASURE] THEN
+ DISCH_THEN(ASSUME_TAC o CONJUNCT2) THEN CONJ_TAC THENL
+ [MATCH_MP_TAC GMEASURABLE_CONVEX THEN REWRITE_TAC[CONVEX_CONVEX_HULL] THEN
+ MATCH_MP_TAC BOUNDED_CONVEX_HULL THEN REWRITE_TAC[BOUNDED_INSERT] THEN
+ ONCE_REWRITE_TAC[SIMPLE_IMAGE_GEN] THEN
+ MATCH_MP_TAC FINITE_IMP_BOUNDED THEN MATCH_MP_TAC FINITE_IMAGE THEN
+ REWRITE_TAC[GSYM numseg; FINITE_NUMSEG];
+ ALL_TAC] THEN
+ ASM_SIMP_TAC[REAL_FIELD `~(y = 0) ==> (x = inv y \<longleftrightarrow> y * x = 1)`;
+ REAL_OF_NUM_EQ; FACT_NZ] THEN
+ FIRST_X_ASSUM(SUBST1_TAC o SYM) THEN MATCH_MP_TAC EQ_TRANS THEN
+ EXISTS_TAC `measure(interval[0:real^N,1])` THEN CONJ_TAC THENL
+ [AP_TERM_TAC; REWRITE_TAC[MEASURE_INTERVAL; CONTENT_UNIT]] THEN
+ REWRITE_TAC[lemma] THEN MATCH_MP_TAC SUBSET_ANTISYM THEN CONJ_TAC THENL
+ [REWRITE_TAC[SUBSET; FORALL_IN_UNIONS; FORALL_IN_IMAGE; IMP_CONJ;
+ RIGHT_FORALL_IMP_THM; IN_ELIM_THM] THEN
+ SIMP_TAC[IMP_IMP; IN_INTERVAL; LAMBDA_BETA; VEC_COMPONENT] THEN
+ X_GEN_TAC `p:num->num` THEN STRIP_TAC THEN X_GEN_TAC `x:real^N` THEN
+ STRIP_TAC THEN X_GEN_TAC `i:num` THEN REPEAT STRIP_TAC THEN
+ MATCH_MP_TAC REAL_LE_TRANS THENL
+ [EXISTS_TAC `(x:real^N)$(p 1)`;
+ EXISTS_TAC `(x:real^N)$(p(dimindex(:N)))`] THEN
+ ASM_REWRITE_TAC[] THEN
+ FIRST_ASSUM(MP_TAC o SPEC `i:num` o MATCH_MP PERMUTES_SURJECTIVE) THEN
+ ASM_MESON_TAC[LE_REFL; PERMUTES_IN_IMAGE; IN_NUMSEG];
+ ALL_TAC] THEN
+ REWRITE_TAC[SET_RULE `s \<subseteq> UNIONS(IMAGE f t) \<longleftrightarrow>
+ !x. x \<in> s ==> ?y. y \<in> t \<and> x \<in> f y`] THEN
+ X_GEN_TAC `x:real^N` THEN REWRITE_TAC[IN_INTERVAL; IN_ELIM_THM] THEN
+ SIMP_TAC[VEC_COMPONENT] THEN DISCH_TAC THEN
+ MP_TAC(ISPEC `\i j. ~((x:real^N)$j \<le> x$i)` TOPOLOGICAL_SORT) THEN
+ REWRITE_TAC[REAL_NOT_LE; REAL_NOT_LT] THEN
+ ANTS_TAC THENL [REAL_ARITH_TAC; ALL_TAC] THEN
+ DISCH_THEN(MP_TAC o SPECL [`dimindex(:N)`; `1..dimindex(:N)`]) THEN
+ REWRITE_TAC[HAS_SIZE_NUMSEG_1; EXTENSION; IN_IMAGE; IN_NUMSEG] THEN
+ DISCH_THEN(X_CHOOSE_THEN `f:num->num` (CONJUNCTS_THEN2
+ (ASSUME_TAC o GSYM) ASSUME_TAC)) THEN
+ EXISTS_TAC `\i. if i \<in> 1..dimindex(:N) then f(i) else i` THEN
+ REWRITE_TAC[] THEN ONCE_REWRITE_TAC[ARITH_RULE
+ `1 \<le> i \<and> i \<le> j \<and> j \<le> n \<longleftrightarrow>
+ 1 \<le> i \<and> 1 \<le> j \<and> i \<le> n \<and> j \<le> n \<and> i \<le> j`] THEN
+ ASM_SIMP_TAC[IN_NUMSEG; LE_REFL; DIMINDEX_GE_1] THEN
+ CONJ_TAC THENL
+ [ALL_TAC;
+ ASM_MESON_TAC[LE_REFL; DIMINDEX_GE_1; LE_LT; REAL_LE_LT]] THEN
+ SIMP_TAC[PERMUTES_FINITE_SURJECTIVE; FINITE_NUMSEG] THEN
+ SIMP_TAC[IN_NUMSEG] THEN ASM_MESON_TAC[]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Hence the gmeasure of a general simplex. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma HAS_GMEASURE_SIMPLEX_0: True .. (*
+ "!l:(real^N)list.
+ LENGTH l = dimindex(:N)
+ ==> (convex hull (0 INSERT set_of_list l)) has_gmeasure
+ abs(det(vector l)) / (FACT(dimindex(:N)))"
+qed REPEAT STRIP_TAC THEN
+ SUBGOAL_THEN
+ `0 INSERT (set_of_list l) =
+ IMAGE (\<lambda>x:real^N. transp(vector l:real^N^N) ** x)
+ (0 INSERT {basis i:real^N | 1 \<le> i \<and> i \<le> dimindex(:N)})`
+ SUBST1_TAC THENL
+ [ONCE_REWRITE_TAC[SIMPLE_IMAGE_GEN] THEN
+ REWRITE_TAC[IMAGE_CLAUSES; GSYM IMAGE_o; o_DEF] THEN
+ REWRITE_TAC[MATRIX_VECTOR_MUL_RZERO] THEN AP_TERM_TAC THEN
+ SIMP_TAC[matrix_vector_mul; vector; transp; LAMBDA_BETA; basis] THEN
+ ONCE_REWRITE_TAC[COND_RAND] THEN
+ SIMP_TAC[REAL_MUL_RZERO; SUM_DELTA] THEN
+ REWRITE_TAC[EXTENSION; IN_IMAGE; IN_ELIM_THM; IN_NUMSEG] THEN
+ ONCE_REWRITE_TAC[TAUT `a \<and> b \<and> c \<longleftrightarrow> ~(b \<and> c ==> ~a)`] THEN
+ X_GEN_TAC `y:real^N` THEN SIMP_TAC[LAMBDA_BETA; REAL_MUL_RID] THEN
+ SIMP_TAC[CART_EQ; LAMBDA_BETA] THEN
+ REWRITE_TAC[NOT_IMP; REAL_MUL_RID; GSYM CART_EQ] THEN
+ ASM_REWRITE_TAC[IN_SET_OF_LIST; MEM_EXISTS_EL] THEN
+ EQ_TAC THEN DISCH_THEN(X_CHOOSE_THEN `i:num` STRIP_ASSUME_TAC) THENL
+ [EXISTS_TAC `SUC i`; EXISTS_TAC `i - 1`] THEN
+ ASM_REWRITE_TAC[SUC_SUB1] THEN ASM_ARITH_TAC;
+ ALL_TAC] THEN
+ ASM_SIMP_TAC[GSYM CONVEX_HULL_LINEAR_IMAGE; MATRIX_VECTOR_MUL_LINEAR] THEN
+ SUBGOAL_THEN
+ `det(vector l:real^N^N) = det(matrix(\<lambda>x:real^N. transp(vector l) ** x))`
+ SUBST1_TAC THENL
+ [REWRITE_TAC[MATRIX_OF_MATRIX_VECTOR_MUL; DET_TRANSP]; ALL_TAC] THEN
+ REWRITE_TAC[real_div] THEN
+ ASM_SIMP_TAC[GSYM(REWRITE_RULE[HAS_GMEASURE_MEASURABLE_MEASURE]
+ HAS_GMEASURE_STD_SIMPLEX)] THEN
+ MATCH_MP_TAC HAS_GMEASURE_LINEAR_IMAGE THEN
+ REWRITE_TAC[MATRIX_VECTOR_MUL_LINEAR] THEN
+ MATCH_MP_TAC GMEASURABLE_CONVEX THEN REWRITE_TAC[CONVEX_CONVEX_HULL] THEN
+ MATCH_MP_TAC BOUNDED_CONVEX_HULL THEN REWRITE_TAC[BOUNDED_INSERT] THEN
+ ONCE_REWRITE_TAC[SIMPLE_IMAGE_GEN] THEN
+ MATCH_MP_TAC FINITE_IMP_BOUNDED THEN MATCH_MP_TAC FINITE_IMAGE THEN
+ REWRITE_TAC[GSYM numseg; FINITE_NUMSEG]);; *)
+
+lemma HAS_GMEASURE_SIMPLEX: True .. (*
+ "!a l:(real^N)list.
+ LENGTH l = dimindex(:N)
+ ==> (convex hull (set_of_list(CONS a l))) has_gmeasure
+ abs(det(vector(MAP (\<lambda>x. x - a) l))) / (FACT(dimindex(:N)))"
+qed REPEAT STRIP_TAC THEN
+ MP_TAC(ISPEC `MAP (\<lambda>x:real^N. x - a) l` HAS_GMEASURE_SIMPLEX_0) THEN
+ ASM_REWRITE_TAC[LENGTH_MAP; set_of_list] THEN
+ DISCH_THEN(MP_TAC o SPEC `a:real^N` o MATCH_MP HAS_GMEASURE_TRANSLATION) THEN
+ REWRITE_TAC[GSYM CONVEX_HULL_TRANSLATION] THEN
+ MATCH_MP_TAC EQ_IMP THEN AP_THM_TAC THEN AP_TERM_TAC THEN
+ REWRITE_TAC[IMAGE_CLAUSES; VECTOR_ADD_RID; SET_OF_LIST_MAP] THEN
+ REWRITE_TAC[GSYM IMAGE_o; o_DEF; VECTOR_ARITH `a + x - a:real^N = x`;
+ SET_RULE `IMAGE (\<lambda>x. x) s = s`]);; *)
+
+lemma GMEASURABLE_SIMPLEX: True .. (*
+ "!l. gmeasurable(convex hull (set_of_list l))"
+qed GEN_TAC THEN
+ MATCH_MP_TAC GMEASURABLE_CONVEX THEN REWRITE_TAC[CONVEX_CONVEX_HULL] THEN
+ MATCH_MP_TAC BOUNDED_CONVEX_HULL THEN
+ MATCH_MP_TAC FINITE_IMP_BOUNDED THEN REWRITE_TAC[FINITE_SET_OF_LIST]);; *)
+
+lemma MEASURE_SIMPLEX: True .. (*
+ "!a l:(real^N)list.
+ LENGTH l = dimindex(:N)
+ ==> measure(convex hull (set_of_list(CONS a l))) =
+ abs(det(vector(MAP (\<lambda>x. x - a) l))) / (FACT(dimindex(:N)))"
+qed MESON_TAC[HAS_GMEASURE_SIMPLEX; HAS_GMEASURE_MEASURABLE_MEASURE]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Area of a triangle. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma HAS_GMEASURE_TRIANGLE: True .. (*
+ "!a b c:real^2.
+ convex hull {a,b,c} has_gmeasure
+ abs((b$1 - a$1) * (c$2 - a$2) - (b$2 - a$2) * (c$1 - a$1)) / 2"
+qed REPEAT STRIP_TAC THEN
+ MP_TAC(ISPECL [`a:real^2`; `[b;c]:(real^2)list`] HAS_GMEASURE_SIMPLEX) THEN
+ REWRITE_TAC[LENGTH; DIMINDEX_2; ARITH; set_of_list; MAP] THEN
+ CONV_TAC NUM_REDUCE_CONV THEN SIMP_TAC[DET_2; VECTOR_2] THEN
+ SIMP_TAC[VECTOR_SUB_COMPONENT; DIMINDEX_2; ARITH]);; *)
+
+lemma GMEASURABLE_TRIANGLE: True .. (*
+ "!a b c:real^N. gmeasurable(convex hull {a,b,c})"
+qed REPEAT GEN_TAC THEN
+ MATCH_MP_TAC GMEASURABLE_CONVEX THEN REWRITE_TAC[CONVEX_CONVEX_HULL] THEN
+ MATCH_MP_TAC BOUNDED_CONVEX_HULL THEN MATCH_MP_TAC FINITE_IMP_BOUNDED THEN
+ REWRITE_TAC[FINITE_INSERT; FINITE_RULES]);; *)
+
+lemma MEASURE_TRIANGLE: True .. (*
+ "!a b c:real^2.
+ measure(convex hull {a,b,c}) =
+ abs((b$1 - a$1) * (c$2 - a$2) - (b$2 - a$2) * (c$1 - a$1)) / 2"
+qed REWRITE_TAC[REWRITE_RULE[HAS_GMEASURE_MEASURABLE_MEASURE]
+ HAS_GMEASURE_TRIANGLE]);; *)
+
+(* ------------------------------------------------------------------------- *)
+(* Volume of a tetrahedron. *)
+(* ------------------------------------------------------------------------- *)
+
+lemma HAS_GMEASURE_TETRAHEDRON: True .. (*
+ "!a b c d:real^3.
+ convex hull {a,b,c,d} has_gmeasure
+ abs((b$1 - a$1) * (c$2 - a$2) * (d$3 - a$3) +
+ (b$2 - a$2) * (c$3 - a$3) * (d$1 - a$1) +
+ (b$3 - a$3) * (c$1 - a$1) * (d$2 - a$2) -
+ (b$1 - a$1) * (c$3 - a$3) * (d$2 - a$2) -
+ (b$2 - a$2) * (c$1 - a$1) * (d$3 - a$3) -
+ (b$3 - a$3) * (c$2 - a$2) * (d$1 - a$1)) /
+ 6"
+qed REPEAT STRIP_TAC THEN
+ MP_TAC(ISPECL [`a:real^3`; `[b;c;d]:(real^3)list`] HAS_GMEASURE_SIMPLEX) THEN
+ REWRITE_TAC[LENGTH; DIMINDEX_3; ARITH; set_of_list; MAP] THEN
+ CONV_TAC NUM_REDUCE_CONV THEN SIMP_TAC[DET_3; VECTOR_3] THEN
+ SIMP_TAC[VECTOR_SUB_COMPONENT; DIMINDEX_3; ARITH]);; *)
+
+lemma GMEASURABLE_TETRAHEDRON: True .. (*
+ "!a b c d:real^N. gmeasurable(convex hull {a,b,c,d})"
+qed REPEAT GEN_TAC THEN
+ MATCH_MP_TAC GMEASURABLE_CONVEX THEN REWRITE_TAC[CONVEX_CONVEX_HULL] THEN
+ MATCH_MP_TAC BOUNDED_CONVEX_HULL THEN MATCH_MP_TAC FINITE_IMP_BOUNDED THEN
+ REWRITE_TAC[FINITE_INSERT; FINITE_RULES]);; *)
+
+lemma MEASURE_TETRAHEDRON: True .. (*
+ "!a b c d:real^3.
+ measure(convex hull {a,b,c,d}) =
+ abs((b$1 - a$1) * (c$2 - a$2) * (d$3 - a$3) +
+ (b$2 - a$2) * (c$3 - a$3) * (d$1 - a$1) +
+ (b$3 - a$3) * (c$1 - a$1) * (d$2 - a$2) -
+ (b$1 - a$1) * (c$3 - a$3) * (d$2 - a$2) -
+ (b$2 - a$2) * (c$1 - a$1) * (d$3 - a$3) -
+ (b$3 - a$3) * (c$2 - a$2) * (d$1 - a$1)) / 6"
+qed REWRITE_TAC[REWRITE_RULE[HAS_GMEASURE_MEASURABLE_MEASURE]
+ HAS_GMEASURE_TETRAHEDRON]);; *)
+
+end
--- a/src/HOL/Multivariate_Analysis/Integration.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Integration.thy Tue Aug 24 08:22:17 2010 +0200
@@ -5287,10 +5287,10 @@
qed finally show ?case . qed qed
lemma nonnegative_absolutely_integrable: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
- assumes "\<forall>x\<in>s. \<forall>i. 0 \<le> f(x)$$i" "f integrable_on s"
+ assumes "\<forall>x\<in>s. \<forall>i<DIM('m). 0 \<le> f(x)$$i" "f integrable_on s"
shows "f absolutely_integrable_on s"
unfolding absolutely_integrable_abs_eq apply rule defer
- apply(rule integrable_eq[of _ f]) using assms by auto
+ apply(rule integrable_eq[of _ f]) using assms apply-apply(subst euclidean_eq) by auto
lemma absolutely_integrable_integrable_bound: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'm::ordered_euclidean_space"
assumes "\<forall>x\<in>s. norm(f x) \<le> g x" "f integrable_on s" "g integrable_on s"
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1,5 +1,5 @@
theory Multivariate_Analysis
-imports Integration Fashoda
+imports Fashoda Gauge_Measure
begin
end
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Tue Aug 24 08:22:17 2010 +0200
@@ -5105,6 +5105,60 @@
shows "open {x::'a::euclidean_space. x$$i > a}"
using open_halfspace_gt[of a "(basis i)::'a"] unfolding euclidean_component_def .
+text{* Instantiation for intervals on @{text ordered_euclidean_space} *}
+
+lemma eucl_lessThan_eq_halfspaces:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "{..<a} = (\<Inter>i<DIM('a). {x. x $$ i < a $$ i})"
+ by (auto simp: eucl_less[where 'a='a])
+
+lemma eucl_greaterThan_eq_halfspaces:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "{a<..} = (\<Inter>i<DIM('a). {x. a $$ i < x $$ i})"
+ by (auto simp: eucl_less[where 'a='a])
+
+lemma eucl_atMost_eq_halfspaces:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "{.. a} = (\<Inter>i<DIM('a). {x. x $$ i \<le> a $$ i})"
+ by (auto simp: eucl_le[where 'a='a])
+
+lemma eucl_atLeast_eq_halfspaces:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "{a ..} = (\<Inter>i<DIM('a). {x. a $$ i \<le> x $$ i})"
+ by (auto simp: eucl_le[where 'a='a])
+
+lemma open_eucl_lessThan[simp, intro]:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "open {..< a}"
+ by (auto simp: eucl_lessThan_eq_halfspaces open_halfspace_component_lt)
+
+lemma open_eucl_greaterThan[simp, intro]:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "open {a <..}"
+ by (auto simp: eucl_greaterThan_eq_halfspaces open_halfspace_component_gt)
+
+lemma closed_eucl_atMost[simp, intro]:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "closed {.. a}"
+ unfolding eucl_atMost_eq_halfspaces
+proof (safe intro!: closed_INT)
+ fix i :: nat
+ have "- {x::'a. x $$ i \<le> a $$ i} = {x. a $$ i < x $$ i}" by auto
+ then show "closed {x::'a. x $$ i \<le> a $$ i}"
+ by (simp add: closed_def open_halfspace_component_gt)
+qed
+
+lemma closed_eucl_atLeast[simp, intro]:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "closed {a ..}"
+ unfolding eucl_atLeast_eq_halfspaces
+proof (safe intro!: closed_INT)
+ fix i :: nat
+ have "- {x::'a. a $$ i \<le> x $$ i} = {x. x $$ i < a $$ i}" by auto
+ then show "closed {x::'a. a $$ i \<le> x $$ i}"
+ by (simp add: closed_def open_halfspace_component_lt)
+qed
+
text{* This gives a simple derivation of limit component bounds. *}
lemma Lim_component_le: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
--- a/src/HOL/Probability/Borel.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Probability/Borel.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1,242 +1,199 @@
-header {*Borel Sets*}
+(* Author: Armin Heller, Johannes Hoelzl, TU Muenchen *)
+
+header {*Borel spaces*}
theory Borel
- imports Measure
+ imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis
begin
-text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
-
-definition borel_space where
- "borel_space = sigma (UNIV::real set) (range (\<lambda>a::real. {x. x \<le> a}))"
+section "Generic Borel spaces"
-definition borel_measurable where
- "borel_measurable a = measurable a borel_space"
+definition "borel_space = sigma (UNIV::'a::topological_space set) open"
+abbreviation "borel_measurable M \<equiv> measurable M borel_space"
-definition indicator_fn where
- "indicator_fn s = (\<lambda>x. if x \<in> s then 1 else (0::real))"
+interpretation borel_space: sigma_algebra borel_space
+ using sigma_algebra_sigma by (auto simp: borel_space_def)
lemma in_borel_measurable:
"f \<in> borel_measurable M \<longleftrightarrow>
- sigma_algebra M \<and>
- (\<forall>s \<in> sets (sigma UNIV (range (\<lambda>a::real. {x. x \<le> a}))).
- f -` s \<inter> space M \<in> sets M)"
-apply (auto simp add: borel_measurable_def measurable_def borel_space_def)
-apply (metis PowD UNIV_I Un_commute sigma_algebra_sigma subset_Pow_Union subset_UNIV subset_Un_eq)
-done
-
-lemma (in sigma_algebra) borel_measurable_const:
- "(\<lambda>x. c) \<in> borel_measurable M"
- by (auto simp add: in_borel_measurable prems)
-
-lemma (in sigma_algebra) borel_measurable_indicator:
- assumes a: "a \<in> sets M"
- shows "indicator_fn a \<in> borel_measurable M"
-apply (auto simp add: in_borel_measurable indicator_fn_def prems)
-apply (metis Diff_eq Int_commute a compl_sets)
-done
+ (\<forall>S \<in> sets (sigma UNIV open).
+ f -` S \<inter> space M \<in> sets M)"
+ by (auto simp add: measurable_def borel_space_def)
-lemma Collect_eq: "{w \<in> X. f w \<le> a} = {w. f w \<le> a} \<inter> X"
- by (metis Collect_conj_eq Collect_mem_eq Int_commute)
+lemma in_borel_measurable_borel_space:
+ "f \<in> borel_measurable M \<longleftrightarrow>
+ (\<forall>S \<in> sets borel_space.
+ f -` S \<inter> space M \<in> sets M)"
+ by (auto simp add: measurable_def borel_space_def)
-lemma (in measure_space) borel_measurable_le_iff:
- "f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
-proof
- assume f: "f \<in> borel_measurable M"
- { fix a
- have "{w \<in> space M. f w \<le> a} \<in> sets M" using f
- apply (auto simp add: in_borel_measurable sigma_def Collect_eq)
- apply (drule_tac x="{x. x \<le> a}" in bspec, auto)
- apply (metis equalityE rangeI subsetD sigma_sets.Basic)
- done
- }
- thus "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M" by blast
-next
- assume "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M"
- thus "f \<in> borel_measurable M"
- apply (simp add: borel_measurable_def borel_space_def Collect_eq)
- apply (rule measurable_sigma, auto)
- done
+lemma space_borel_space[simp]: "space borel_space = UNIV"
+ unfolding borel_space_def by auto
+
+lemma borel_space_open[simp]:
+ assumes "open A" shows "A \<in> sets borel_space"
+proof -
+ have "A \<in> open" unfolding mem_def using assms .
+ thus ?thesis unfolding borel_space_def sigma_def by (auto intro!: sigma_sets.Basic)
qed
-lemma Collect_less_le:
- "{w \<in> X. f w < g w} = (\<Union>n. {w \<in> X. f w \<le> g w - inverse(real(Suc n))})"
- proof auto
- fix w
- assume w: "f w < g w"
- hence nz: "g w - f w \<noteq> 0"
- by arith
- with w have "real(Suc(natceiling(inverse(g w - f w)))) > inverse(g w - f w)"
- 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)))))
- < inverse(inverse(g w - f w))"
- by (metis less_iff_diff_less_0 less_imp_inverse_less linorder_neqE_linordered_idom nz positive_imp_inverse_positive order_antisym less_le w)
- hence "inverse(real(Suc(natceiling(inverse(g w - f w))))) < g w - f w"
- by (metis inverse_inverse_eq order_less_le_trans order_refl)
- thus "\<exists>n. f w \<le> g w - inverse(real(Suc n))" using w
- by (rule_tac x="natceiling(inverse(g w - f w))" in exI, auto)
- next
- fix w n
- assume le: "f w \<le> g w - inverse(real(Suc n))"
- hence "0 < inverse(real(Suc n))"
- by simp
- thus "f w < g w" using le
- by arith
- qed
-
-
-lemma (in sigma_algebra) sigma_le_less:
- assumes M: "!!a::real. {w \<in> space M. f w \<le> a} \<in> sets M"
- shows "{w \<in> space M. f w < a} \<in> sets M"
+lemma borel_space_closed[simp]:
+ assumes "closed A" shows "A \<in> sets borel_space"
proof -
- show ?thesis using Collect_less_le [of "space M" f "\<lambda>x. a"]
- by (auto simp add: countable_UN M)
+ have "space borel_space - (- A) \<in> sets borel_space"
+ using assms unfolding closed_def by (blast intro: borel_space_open)
+ thus ?thesis by simp
qed
-lemma (in sigma_algebra) sigma_less_ge:
- assumes M: "!!a::real. {w \<in> space M. f w < a} \<in> sets M"
- shows "{w \<in> space M. a \<le> f w} \<in> sets M"
-proof -
- have "{w \<in> space M. a \<le> f w} = space M - {w \<in> space M. f w < a}"
- by auto
- thus ?thesis using M
- by auto
-qed
-
-lemma (in sigma_algebra) sigma_ge_gr:
- assumes M: "!!a::real. {w \<in> space M. a \<le> f w} \<in> sets M"
- shows "{w \<in> space M. a < f w} \<in> sets M"
-proof -
- show ?thesis using Collect_less_le [of "space M" "\<lambda>x. a" f]
- by (auto simp add: countable_UN le_diff_eq M)
+lemma (in sigma_algebra) borel_measurable_vimage:
+ fixes f :: "'a \<Rightarrow> 'x::t2_space"
+ assumes borel: "f \<in> borel_measurable M"
+ shows "f -` {x} \<inter> space M \<in> sets M"
+proof (cases "x \<in> f ` space M")
+ case True then obtain y where "x = f y" by auto
+ from closed_sing[of "f y"]
+ have "{f y} \<in> sets borel_space" by (rule borel_space_closed)
+ with assms show ?thesis
+ unfolding in_borel_measurable_borel_space `x = f y` by auto
+next
+ case False hence "f -` {x} \<inter> space M = {}" by auto
+ thus ?thesis by auto
qed
-lemma (in sigma_algebra) sigma_gr_le:
- assumes M: "!!a::real. {w \<in> space M. a < f w} \<in> sets M"
- shows "{w \<in> space M. f w \<le> a} \<in> sets M"
-proof -
- have "{w \<in> space M. f w \<le> a} = space M - {w \<in> space M. a < f w}"
- by auto
- thus ?thesis
- by (simp add: M compl_sets)
-qed
+lemma (in sigma_algebra) borel_measurableI:
+ fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space"
+ assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
+ shows "f \<in> borel_measurable M"
+ unfolding borel_space_def
+proof (rule measurable_sigma)
+ fix S :: "'x set" assume "S \<in> open" thus "f -` S \<inter> space M \<in> sets M"
+ using assms[of S] by (simp add: mem_def)
+qed simp_all
-lemma (in measure_space) borel_measurable_gr_iff:
- "f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
-proof (auto simp add: borel_measurable_le_iff sigma_gr_le)
- fix u
- assume M: "\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M"
- have "{w \<in> space M. u < f w} = space M - {w \<in> space M. f w \<le> u}"
- by auto
- thus "{w \<in> space M. u < f w} \<in> sets M"
- by (force simp add: compl_sets countable_UN M)
-qed
+lemma borel_space_singleton[simp, intro]:
+ fixes x :: "'a::t1_space"
+ shows "A \<in> sets borel_space \<Longrightarrow> insert x A \<in> sets borel_space"
+ proof (rule borel_space.insert_in_sets)
+ show "{x} \<in> sets borel_space"
+ using closed_sing[of x] by (rule borel_space_closed)
+ qed simp
+
+lemma (in sigma_algebra) borel_measurable_const[simp, intro]:
+ "(\<lambda>x. c) \<in> borel_measurable M"
+ by (auto intro!: measurable_const)
-lemma (in measure_space) borel_measurable_less_iff:
- "f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
-proof (auto simp add: borel_measurable_le_iff sigma_le_less)
- fix u
- assume M: "\<forall>a. {w \<in> space M. f w < a} \<in> sets M"
- have "{w \<in> space M. f w \<le> u} = space M - {w \<in> space M. u < f w}"
- by auto
- thus "{w \<in> space M. f w \<le> u} \<in> sets M"
- using Collect_less_le [of "space M" "\<lambda>x. u" f]
- by (force simp add: compl_sets countable_UN le_diff_eq sigma_less_ge M)
-qed
+lemma (in sigma_algebra) borel_measurable_indicator:
+ assumes A: "A \<in> sets M"
+ shows "indicator A \<in> borel_measurable M"
+ unfolding indicator_def_raw using A
+ by (auto intro!: measurable_If_set borel_measurable_const)
-lemma (in measure_space) borel_measurable_ge_iff:
- "f \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
-proof (auto simp add: borel_measurable_less_iff sigma_le_less sigma_ge_gr sigma_gr_le)
- fix u
- assume M: "\<forall>a. {w \<in> space M. f w < a} \<in> sets M"
- have "{w \<in> space M. u \<le> f w} = space M - {w \<in> space M. f w < u}"
- by auto
- thus "{w \<in> space M. u \<le> f w} \<in> sets M"
- by (force simp add: compl_sets countable_UN M)
+lemma borel_measurable_translate:
+ assumes "A \<in> sets borel_space" and trans: "\<And>B. open B \<Longrightarrow> f -` B \<in> sets borel_space"
+ shows "f -` A \<in> sets borel_space"
+proof -
+ have "A \<in> sigma_sets UNIV open" using assms
+ by (simp add: borel_space_def sigma_def)
+ thus ?thesis
+ proof (induct rule: sigma_sets.induct)
+ case (Basic a) thus ?case using trans[of a] by (simp add: mem_def)
+ next
+ case (Compl a)
+ moreover have "UNIV \<in> sets borel_space"
+ by (metis borel_space.top borel_space_def_raw mem_def space_sigma)
+ ultimately show ?case
+ by (auto simp: vimage_Diff borel_space.Diff)
+ qed (auto simp add: vimage_UN)
qed
-lemma (in measure_space) affine_borel_measurable:
- assumes g: "g \<in> borel_measurable M"
- shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"
-proof (cases rule: linorder_cases [of b 0])
- case equal thus ?thesis
- by (simp add: borel_measurable_const)
-next
- case less
- {
- fix w c
- have "a + g w * b \<le> c \<longleftrightarrow> g w * b \<le> c - a"
- by auto
- also have "... \<longleftrightarrow> (c-a)/b \<le> g w" using less
- by (metis divide_le_eq less less_asym)
- finally have "a + g w * b \<le> c \<longleftrightarrow> (c-a)/b \<le> g w" .
- }
- hence "\<And>w c. a + g w * b \<le> c \<longleftrightarrow> (c-a)/b \<le> g w" .
- thus ?thesis using less g
- by (simp add: borel_measurable_ge_iff [of g])
- (simp add: borel_measurable_le_iff)
-next
- case greater
- hence 0: "\<And>x c. (g x * b \<le> c - a) \<longleftrightarrow> (g x \<le> (c - a) / b)"
- by (metis mult_imp_le_div_pos le_divide_eq)
- have 1: "\<And>x c. (a + g x * b \<le> c) \<longleftrightarrow> (g x * b \<le> c - a)"
- by auto
- from greater g
- show ?thesis
- by (simp add: borel_measurable_le_iff 0 1)
-qed
+section "Borel spaces on euclidean spaces"
+
+lemma lessThan_borel[simp, intro]:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "{..< a} \<in> sets borel_space"
+ by (blast intro: borel_space_open)
+
+lemma greaterThan_borel[simp, intro]:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "{a <..} \<in> sets borel_space"
+ by (blast intro: borel_space_open)
+
+lemma greaterThanLessThan_borel[simp, intro]:
+ fixes a b :: "'a\<Colon>ordered_euclidean_space"
+ shows "{a<..<b} \<in> sets borel_space"
+ by (blast intro: borel_space_open)
+
+lemma atMost_borel[simp, intro]:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "{..a} \<in> sets borel_space"
+ by (blast intro: borel_space_closed)
+
+lemma atLeast_borel[simp, intro]:
+ fixes a :: "'a\<Colon>ordered_euclidean_space"
+ shows "{a..} \<in> sets borel_space"
+ by (blast intro: borel_space_closed)
+
+lemma atLeastAtMost_borel[simp, intro]:
+ fixes a b :: "'a\<Colon>ordered_euclidean_space"
+ shows "{a..b} \<in> sets borel_space"
+ by (blast intro: borel_space_closed)
-definition
- nat_to_rat_surj :: "nat \<Rightarrow> rat" where
- "nat_to_rat_surj n = (let (i,j) = prod_decode n
- in Fract (int_decode i) (int_decode j))"
+lemma greaterThanAtMost_borel[simp, intro]:
+ fixes a b :: "'a\<Colon>ordered_euclidean_space"
+ shows "{a<..b} \<in> sets borel_space"
+ unfolding greaterThanAtMost_def by blast
+
+lemma atLeastLessThan_borel[simp, intro]:
+ fixes a b :: "'a\<Colon>ordered_euclidean_space"
+ shows "{a..<b} \<in> sets borel_space"
+ unfolding atLeastLessThan_def by blast
+
+lemma hafspace_less_borel[simp, intro]:
+ fixes a :: real
+ shows "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel_space"
+ by (auto intro!: borel_space_open open_halfspace_component_gt)
-lemma nat_to_rat_surj: "surj nat_to_rat_surj"
-proof (auto simp add: surj_def nat_to_rat_surj_def)
- fix y
- show "\<exists>x. y = (\<lambda>(i, j). Fract (int_decode i) (int_decode j)) (prod_decode x)"
- proof (cases y)
- case (Fract a b)
- obtain i where i: "int_decode i = a" using surj_int_decode
- by (metis surj_def)
- obtain j where j: "int_decode j = b" using surj_int_decode
- by (metis surj_def)
- obtain n where n: "prod_decode n = (i,j)" using surj_prod_decode
- by (metis surj_def)
+lemma hafspace_greater_borel[simp, intro]:
+ fixes a :: real
+ shows "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel_space"
+ by (auto intro!: borel_space_open open_halfspace_component_lt)
- from Fract i j n show ?thesis
- by (metis prod.cases)
- qed
-qed
+lemma hafspace_less_eq_borel[simp, intro]:
+ fixes a :: real
+ shows "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel_space"
+ by (auto intro!: borel_space_closed closed_halfspace_component_ge)
-lemma rats_enumeration: "\<rat> = range (of_rat o nat_to_rat_surj)"
- using nat_to_rat_surj
- by (auto simp add: image_def surj_def) (metis Rats_cases)
+lemma hafspace_greater_eq_borel[simp, intro]:
+ fixes a :: real
+ shows "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel_space"
+ by (auto intro!: borel_space_closed closed_halfspace_component_le)
-lemma (in measure_space) borel_measurable_less_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_less[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "{w \<in> space M. f w < g w} \<in> sets M"
proof -
have "{w \<in> space M. f w < g w} =
- (\<Union>r\<in>\<rat>. {w \<in> space M. f w < r} \<inter> {w \<in> space M. r < g w })"
- by (auto simp add: Rats_dense_in_real)
- thus ?thesis using f g
- by (simp add: borel_measurable_less_iff [of f]
- borel_measurable_gr_iff [of g])
- (blast intro: gen_countable_UN [OF rats_enumeration])
+ (\<Union>r. (f -` {..< of_rat r} \<inter> space M) \<inter> (g -` {of_rat r <..} \<inter> space M))"
+ using Rats_dense_in_real by (auto simp add: Rats_def)
+ then show ?thesis using f g
+ by simp (blast intro: measurable_sets)
qed
-
-lemma (in measure_space) borel_measurable_leq_borel_measurable:
+
+lemma (in sigma_algebra) borel_measurable_le[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "{w \<in> space M. f w \<le> g w} \<in> sets M"
proof -
- have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}"
- by auto
- thus ?thesis using f g
- by (simp add: borel_measurable_less_borel_measurable compl_sets)
+ have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}"
+ by auto
+ thus ?thesis using f g
+ by simp blast
qed
-lemma (in measure_space) borel_measurable_eq_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_eq[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "{w \<in> space M. f w = g w} \<in> sets M"
@@ -244,40 +201,512 @@
have "{w \<in> space M. f w = g w} =
{w \<in> space M. f w \<le> g w} \<inter> {w \<in> space M. g w \<le> f w}"
by auto
- thus ?thesis using f g
- by (simp add: borel_measurable_leq_borel_measurable Int)
+ thus ?thesis using f g by auto
qed
-lemma (in measure_space) borel_measurable_neq_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_neq[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
proof -
have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}"
by auto
- thus ?thesis using f g
- by (simp add: borel_measurable_eq_borel_measurable compl_sets)
+ thus ?thesis using f g by auto
+qed
+
+subsection "Borel space equals sigma algebras over intervals"
+
+lemma rational_boxes:
+ fixes x :: "'a\<Colon>ordered_euclidean_space"
+ assumes "0 < e"
+ shows "\<exists>a b. (\<forall>i. a $$ i \<in> \<rat>) \<and> (\<forall>i. b $$ i \<in> \<rat>) \<and> x \<in> {a <..< b} \<and> {a <..< b} \<subseteq> ball x e"
+proof -
+ def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
+ then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos)
+ have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x $$ i \<and> x $$ i - y < e'" (is "\<forall>i. ?th i")
+ proof
+ fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e
+ show "?th i" by auto
+ qed
+ from choice[OF this] guess a .. note a = this
+ have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x $$ i < y \<and> y - x $$ i < e'" (is "\<forall>i. ?th i")
+ proof
+ fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e
+ show "?th i" by auto
+ qed
+ from choice[OF this] guess b .. note b = this
+ { fix y :: 'a assume *: "Chi a < y" "y < Chi b"
+ have "dist x y = sqrt (\<Sum>i<DIM('a). (dist (x $$ i) (y $$ i))\<twosuperior>)"
+ unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
+ also have "\<dots> < sqrt (\<Sum>i<DIM('a). e^2 / real (DIM('a)))"
+ proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
+ fix i assume i: "i \<in> {..<DIM('a)}"
+ have "a i < y$$i \<and> y$$i < b i" using * i eucl_less[where 'a='a] by auto
+ moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto
+ moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto
+ ultimately have "\<bar>x$$i - y$$i\<bar> < 2 * e'" by auto
+ then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))"
+ unfolding e'_def by (auto simp: dist_real_def)
+ then have "(dist (x $$ i) (y $$ i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
+ by (rule power_strict_mono) auto
+ then show "(dist (x $$ i) (y $$ i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
+ by (simp add: power_divide)
+ qed auto
+ also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive)
+ finally have "dist x y < e" . }
+ with a b show ?thesis
+ apply (rule_tac exI[of _ "Chi a"])
+ apply (rule_tac exI[of _ "Chi b"])
+ using eucl_less[where 'a='a] by auto
+qed
+
+lemma ex_rat_list:
+ fixes x :: "'a\<Colon>ordered_euclidean_space"
+ assumes "\<And> i. x $$ i \<in> \<rat>"
+ shows "\<exists> r. length r = DIM('a) \<and> (\<forall> i < DIM('a). of_rat (r ! i) = x $$ i)"
+proof -
+ have "\<forall>i. \<exists>r. x $$ i = of_rat r" using assms unfolding Rats_def by blast
+ from choice[OF this] guess r ..
+ then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"])
+qed
+
+lemma open_UNION:
+ fixes M :: "'a\<Colon>ordered_euclidean_space set"
+ assumes "open M"
+ shows "M = UNION {(a, b) | a b. {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)} \<subseteq> M}
+ (\<lambda> (a, b). {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)})"
+ (is "M = UNION ?idx ?box")
+proof safe
+ fix x assume "x \<in> M"
+ obtain e where e: "e > 0" "ball x e \<subseteq> M"
+ using openE[OF assms `x \<in> M`] by auto
+ then obtain a b where ab: "x \<in> {a <..< b}" "\<And>i. a $$ i \<in> \<rat>" "\<And>i. b $$ i \<in> \<rat>" "{a <..< b} \<subseteq> ball x e"
+ using rational_boxes[OF e(1)] by blast
+ then obtain p q where pq: "length p = DIM ('a)"
+ "length q = DIM ('a)"
+ "\<forall> i < DIM ('a). of_rat (p ! i) = a $$ i \<and> of_rat (q ! i) = b $$ i"
+ using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast
+ hence p: "Chi (of_rat \<circ> op ! p) = a"
+ using euclidean_eq[of "Chi (of_rat \<circ> op ! p)" a]
+ unfolding o_def by auto
+ from pq have q: "Chi (of_rat \<circ> op ! q) = b"
+ using euclidean_eq[of "Chi (of_rat \<circ> op ! q)" b]
+ unfolding o_def by auto
+ have "x \<in> ?box (p, q)"
+ using p q ab by auto
+ thus "x \<in> UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto
+qed auto
+
+lemma halfspace_span_open:
+ "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))
+ \<subseteq> sets borel_space"
+ by (auto intro!: borel_space.sigma_sets_subset[simplified] borel_space_open
+ open_halfspace_component_lt simp: sets_sigma)
+
+lemma halfspace_lt_in_halfspace:
+ "{x\<Colon>'a. x $$ i < a} \<in> sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))"
+ unfolding sets_sigma by (rule sigma_sets.Basic) auto
+
+lemma halfspace_gt_in_halfspace:
+ "{x\<Colon>'a. a < x $$ i} \<in> sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))"
+ (is "?set \<in> sets ?SIGMA")
+proof -
+ interpret sigma_algebra ?SIGMA by (rule sigma_algebra_sigma) simp
+ have *: "?set = (\<Union>n. space ?SIGMA - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
+ proof (safe, simp_all add: not_less)
+ fix x assume "a < x $$ i"
+ with reals_Archimedean[of "x $$ i - a"]
+ obtain n where "a + 1 / real (Suc n) < x $$ i"
+ by (auto simp: inverse_eq_divide field_simps)
+ then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i"
+ by (blast intro: less_imp_le)
+ next
+ fix x n
+ have "a < a + 1 / real (Suc n)" by auto
+ also assume "\<dots> \<le> x"
+ finally show "a < x" .
+ qed
+ show "?set \<in> sets ?SIGMA" unfolding *
+ by (safe intro!: countable_UN Diff halfspace_lt_in_halfspace)
qed
-lemma (in measure_space) borel_measurable_add_borel_measurable:
+lemma (in sigma_algebra) sets_sigma_subset:
+ assumes "A = space M"
+ assumes "B \<subseteq> sets M"
+ shows "sets (sigma A B) \<subseteq> sets M"
+ by (unfold assms sets_sigma, rule sigma_sets_subset, rule assms)
+
+lemma open_span_halfspace:
+ "sets borel_space \<subseteq> sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i < a})))"
+ (is "_ \<subseteq> sets ?SIGMA")
+proof (unfold borel_space_def, rule sigma_algebra.sets_sigma_subset, safe)
+ show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) simp
+ then interpret sigma_algebra ?SIGMA .
+ fix S :: "'a set" assume "S \<in> open" then have "open S" unfolding mem_def .
+ from open_UNION[OF this]
+ obtain I where *: "S =
+ (\<Union>(a, b)\<in>I.
+ (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
+ (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
+ unfolding greaterThanLessThan_def
+ unfolding eucl_greaterThan_eq_halfspaces[where 'a='a]
+ unfolding eucl_lessThan_eq_halfspaces[where 'a='a]
+ by blast
+ show "S \<in> sets ?SIGMA"
+ unfolding *
+ by (auto intro!: countable_UN Int countable_INT halfspace_lt_in_halfspace halfspace_gt_in_halfspace)
+qed auto
+
+lemma halfspace_span_halfspace_le:
+ "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))) \<subseteq>
+ sets (sigma UNIV (range (\<lambda> (a, i). {x. x $$ i \<le> a})))"
+ (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+ show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+ then interpret sigma_algebra ?SIGMA .
+ fix a i
+ have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
+ proof (safe, simp_all)
+ fix x::'a assume *: "x$$i < a"
+ with reals_Archimedean[of "a - x$$i"]
+ obtain n where "x $$ i < a - 1 / (real (Suc n))"
+ by (auto simp: field_simps inverse_eq_divide)
+ then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))"
+ by (blast intro: less_imp_le)
+ next
+ fix x::'a and n
+ assume "x$$i \<le> a - 1 / real (Suc n)"
+ also have "\<dots> < a" by auto
+ finally show "x$$i < a" .
+ qed
+ show "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
+ by (safe intro!: countable_UN)
+ (auto simp: sets_sigma intro!: sigma_sets.Basic)
+qed auto
+
+lemma halfspace_span_halfspace_ge:
+ "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))) \<subseteq>
+ sets (sigma UNIV (range (\<lambda> (a, i). {x. a \<le> x $$ i})))"
+ (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+ show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+ then interpret sigma_algebra ?SIGMA .
+ fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
+ show "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
+ by (safe intro!: Diff)
+ (auto simp: sets_sigma intro!: sigma_sets.Basic)
+qed auto
+
+lemma halfspace_le_span_halfspace_gt:
+ "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
+ sets (sigma UNIV (range (\<lambda> (a, i). {x. a < x $$ i})))"
+ (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+ show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+ then interpret sigma_algebra ?SIGMA .
+ fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
+ show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
+ by (safe intro!: Diff)
+ (auto simp: sets_sigma intro!: sigma_sets.Basic)
+qed auto
+
+lemma halfspace_le_span_atMost:
+ "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
+ sets (sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})))"
+ (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+ show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+ then interpret sigma_algebra ?SIGMA .
+ fix a i
+ show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+ proof cases
+ assume "i < DIM('a)"
+ then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
+ proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)
+ fix x
+ from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
+ then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
+ by (subst (asm) Max_le_iff) auto
+ then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
+ by (auto intro!: exI[of _ k])
+ qed
+ show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
+ by (safe intro!: countable_UN)
+ (auto simp: sets_sigma intro!: sigma_sets.Basic)
+ next
+ assume "\<not> i < DIM('a)"
+ then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+ using top by auto
+ qed
+qed auto
+
+lemma halfspace_le_span_greaterThan:
+ "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
+ sets (sigma UNIV (range (\<lambda>a. {a<..})))"
+ (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+ show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+ then interpret sigma_algebra ?SIGMA .
+ fix a i
+ show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+ proof cases
+ assume "i < DIM('a)"
+ have "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
+ also have *: "{x::'a. a < x$$i} = (\<Union>k::nat. {(\<chi>\<chi> n. if n = i then a else -real k) <..})" using `i <DIM('a)`
+ proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
+ fix x
+ from real_arch_lt[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
+ guess k::nat .. note k = this
+ { fix i assume "i < DIM('a)"
+ then have "-x$$i < real k"
+ using k by (subst (asm) Max_less_iff) auto
+ then have "- real k < x$$i" by simp }
+ then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
+ by (auto intro!: exI[of _ k])
+ qed
+ finally show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+ apply (simp only:)
+ apply (safe intro!: countable_UN Diff)
+ by (auto simp: sets_sigma intro!: sigma_sets.Basic)
+ next
+ assume "\<not> i < DIM('a)"
+ then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+ using top by auto
+ qed
+qed auto
+
+lemma atMost_span_atLeastAtMost:
+ "sets (sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))) \<subseteq>
+ sets (sigma UNIV (range (\<lambda>(a,b). {a..b})))"
+ (is "_ \<subseteq> sets ?SIGMA")
+proof (rule sigma_algebra.sets_sigma_subset, safe)
+ show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+ then interpret sigma_algebra ?SIGMA .
+ fix a::'a
+ have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
+ proof (safe, simp_all add: eucl_le[where 'a='a])
+ fix x
+ from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
+ guess k::nat .. note k = this
+ { fix i assume "i < DIM('a)"
+ with k have "- x$$i \<le> real k"
+ by (subst (asm) Max_le_iff) (auto simp: field_simps)
+ then have "- real k \<le> x$$i" by simp }
+ then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
+ by (auto intro!: exI[of _ k])
+ qed
+ show "{..a} \<in> sets ?SIGMA" unfolding *
+ by (safe intro!: countable_UN)
+ (auto simp: sets_sigma intro!: sigma_sets.Basic)
+qed auto
+
+lemma borel_space_eq_greaterThanLessThan:
+ "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, b). {a <..< (b :: 'a \<Colon> ordered_euclidean_space)})))"
+ (is "_ = sets ?SIGMA")
+proof (rule antisym)
+ show "sets ?SIGMA \<subseteq> sets borel_space"
+ by (rule borel_space.sets_sigma_subset) auto
+ show "sets borel_space \<subseteq> sets ?SIGMA"
+ unfolding borel_space_def
+ proof (rule sigma_algebra.sets_sigma_subset, safe)
+ show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+ then interpret sigma_algebra ?SIGMA .
+ fix M :: "'a set" assume "M \<in> open"
+ then have "open M" by (simp add: mem_def)
+ show "M \<in> sets ?SIGMA"
+ apply (subst open_UNION[OF `open M`])
+ apply (safe intro!: countable_UN)
+ by (auto simp add: sigma_def intro!: sigma_sets.Basic)
+ qed auto
+qed
+
+lemma borel_space_eq_atMost:
+ "sets borel_space = sets (sigma UNIV (range (\<lambda> a. {.. a::'a\<Colon>ordered_euclidean_space})))"
+ (is "_ = sets ?SIGMA")
+proof (rule antisym)
+ show "sets borel_space \<subseteq> sets ?SIGMA"
+ using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace
+ by auto
+ show "sets ?SIGMA \<subseteq> sets borel_space"
+ by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_atLeastAtMost:
+ "sets borel_space = sets (sigma UNIV (range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space, b). {a .. b})))"
+ (is "_ = sets ?SIGMA")
+proof (rule antisym)
+ show "sets borel_space \<subseteq> sets ?SIGMA"
+ using atMost_span_atLeastAtMost halfspace_le_span_atMost
+ halfspace_span_halfspace_le open_span_halfspace
+ by auto
+ show "sets ?SIGMA \<subseteq> sets borel_space"
+ by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_greaterThan:
+ "sets borel_space = sets (sigma UNIV (range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {a <..})))"
+ (is "_ = sets ?SIGMA")
+proof (rule antisym)
+ show "sets borel_space \<subseteq> sets ?SIGMA"
+ using halfspace_le_span_greaterThan
+ halfspace_span_halfspace_le open_span_halfspace
+ by auto
+ show "sets ?SIGMA \<subseteq> sets borel_space"
+ by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_halfspace_le:
+ "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i \<le> a})))"
+ (is "_ = sets ?SIGMA")
+proof (rule antisym)
+ show "sets borel_space \<subseteq> sets ?SIGMA"
+ using open_span_halfspace halfspace_span_halfspace_le by auto
+ show "sets ?SIGMA \<subseteq> sets borel_space"
+ by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_halfspace_less:
+ "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i < a})))"
+ (is "_ = sets ?SIGMA")
+proof (rule antisym)
+ show "sets borel_space \<subseteq> sets ?SIGMA"
+ using open_span_halfspace .
+ show "sets ?SIGMA \<subseteq> sets borel_space"
+ by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_halfspace_gt:
+ "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a < x$$i})))"
+ (is "_ = sets ?SIGMA")
+proof (rule antisym)
+ show "sets borel_space \<subseteq> sets ?SIGMA"
+ using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto
+ show "sets ?SIGMA \<subseteq> sets borel_space"
+ by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma borel_space_eq_halfspace_ge:
+ "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a \<le> x$$i})))"
+ (is "_ = sets ?SIGMA")
+proof (rule antisym)
+ show "sets borel_space \<subseteq> sets ?SIGMA"
+ using halfspace_span_halfspace_ge open_span_halfspace by auto
+ show "sets ?SIGMA \<subseteq> sets borel_space"
+ by (rule borel_space.sets_sigma_subset) auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_halfspacesI:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+ assumes "sets borel_space = sets (sigma UNIV (range F))"
+ and "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
+ and "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
+ shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
+proof safe
+ fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
+ then show "S a i \<in> sets M" unfolding assms
+ by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1) sigma_def)
+next
+ assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
+ { fix a i have "S a i \<in> sets M"
+ proof cases
+ assume "i < DIM('c)"
+ with a show ?thesis unfolding assms(2) by simp
+ next
+ assume "\<not> i < DIM('c)"
+ from assms(3)[OF this] show ?thesis .
+ qed }
+ then have "f \<in> measurable M (sigma UNIV (range F))"
+ by (auto intro!: measurable_sigma simp: assms(2))
+ then show "f \<in> borel_measurable M" unfolding measurable_def
+ unfolding assms(1) by simp
+qed
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_le:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+ shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i \<le> a} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_le]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_less:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i < a} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_less]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_ge:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+ shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a \<le> f w $$ i} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_ge]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_greater:
+ fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_gt]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_le:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
+ using borel_measurable_iff_halfspace_le[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_less:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
+ using borel_measurable_iff_halfspace_less[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_ge:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
+ using borel_measurable_iff_halfspace_ge[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_greater:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
+ using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
+
+subsection "Borel measurable operators"
+
+lemma (in sigma_algebra) affine_borel_measurable_vector:
+ fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
+ assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
+proof (rule borel_measurableI)
+ fix S :: "'x set" assume "open S"
+ show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"
+ proof cases
+ assume "b \<noteq> 0"
+ with `open S` have "((\<lambda>x. (- a + x) /\<^sub>R b) ` S) \<in> open" (is "?S \<in> open")
+ by (auto intro!: open_affinity simp: scaleR.add_right mem_def)
+ hence "?S \<in> sets borel_space"
+ unfolding borel_space_def by (auto simp: sigma_def intro!: sigma_sets.Basic)
+ moreover
+ from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
+ apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all)
+ ultimately show ?thesis using assms unfolding in_borel_measurable_borel_space
+ by auto
+ qed simp
+qed
+
+lemma (in sigma_algebra) affine_borel_measurable:
+ fixes g :: "'a \<Rightarrow> real"
+ assumes g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"
+ using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute)
+
+lemma (in sigma_algebra) borel_measurable_add[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
proof -
- have 1:"!!a. {w \<in> space M. a \<le> f w + g w} = {w \<in> space M. a + (g w) * -1 \<le> f w}"
+ have 1: "\<And>a. {w\<in>space M. a \<le> f w + g w} = {w \<in> space M. a + g w * -1 \<le> f w}"
by auto
- have "!!a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M"
- by (rule affine_borel_measurable [OF g])
- hence "!!a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f
- by (rule borel_measurable_leq_borel_measurable)
- hence "!!a. {w \<in> space M. a \<le> f w + g w} \<in> sets M"
- by (simp add: 1)
- thus ?thesis
- by (simp add: borel_measurable_ge_iff)
+ have "\<And>a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M"
+ by (rule affine_borel_measurable [OF g])
+ then have "\<And>a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f
+ by auto
+ then have "\<And>a. {w \<in> space M. a \<le> f w + g w} \<in> sets M"
+ by (simp add: 1)
+ then show ?thesis
+ by (simp add: borel_measurable_iff_ge)
qed
-
-lemma (in measure_space) borel_measurable_square:
+lemma (in sigma_algebra) borel_measurable_square:
+ fixes f :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
shows "(\<lambda>x. (f x)^2) \<in> borel_measurable M"
proof -
@@ -286,21 +715,21 @@
have "{w \<in> space M. (f w)\<twosuperior> \<le> a} \<in> sets M"
proof (cases rule: linorder_cases [of a 0])
case less
- hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = {}"
+ hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = {}"
by auto (metis less order_le_less_trans power2_less_0)
also have "... \<in> sets M"
- by (rule empty_sets)
+ by (rule empty_sets)
finally show ?thesis .
next
case equal
- hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
+ hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
{w \<in> space M. f w \<le> 0} \<inter> {w \<in> space M. 0 \<le> f w}"
by auto
also have "... \<in> sets M"
- apply (insert f)
- apply (rule Int)
- apply (simp add: borel_measurable_le_iff)
- apply (simp add: borel_measurable_ge_iff)
+ apply (insert f)
+ apply (rule Int)
+ apply (simp add: borel_measurable_iff_le)
+ apply (simp add: borel_measurable_iff_ge)
done
finally show ?thesis .
next
@@ -309,329 +738,536 @@
by (metis abs_le_interval_iff abs_of_pos greater real_sqrt_abs
real_sqrt_le_iff real_sqrt_power)
hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
- {w \<in> space M. -(sqrt a) \<le> f w} \<inter> {w \<in> space M. f w \<le> sqrt a}"
+ {w \<in> space M. -(sqrt a) \<le> f w} \<inter> {w \<in> space M. f w \<le> sqrt a}"
using greater by auto
also have "... \<in> sets M"
- apply (insert f)
- apply (rule Int)
- apply (simp add: borel_measurable_ge_iff)
- apply (simp add: borel_measurable_le_iff)
+ apply (insert f)
+ apply (rule Int)
+ apply (simp add: borel_measurable_iff_ge)
+ apply (simp add: borel_measurable_iff_le)
done
finally show ?thesis .
qed
}
- thus ?thesis by (auto simp add: borel_measurable_le_iff)
+ thus ?thesis by (auto simp add: borel_measurable_iff_le)
qed
lemma times_eq_sum_squares:
fixes x::real
shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4"
-by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric])
+by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric])
-
-lemma (in measure_space) borel_measurable_uminus_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_uminus[simp, intro]:
+ fixes g :: "'a \<Rightarrow> real"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. - g x) \<in> borel_measurable M"
proof -
have "(\<lambda>x. - g x) = (\<lambda>x. 0 + (g x) * -1)"
by simp
- also have "... \<in> borel_measurable M"
- by (fast intro: affine_borel_measurable g)
+ also have "... \<in> borel_measurable M"
+ by (fast intro: affine_borel_measurable g)
finally show ?thesis .
qed
-lemma (in measure_space) borel_measurable_times_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_times[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
proof -
have 1: "(\<lambda>x. 0 + (f x + g x)\<twosuperior> * inverse 4) \<in> borel_measurable M"
- by (fast intro: affine_borel_measurable borel_measurable_square
- borel_measurable_add_borel_measurable f g)
- have "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) =
+ using assms by (fast intro: affine_borel_measurable borel_measurable_square)
+ have "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) =
(\<lambda>x. 0 + ((f x + -g x) ^ 2 * inverse -4))"
by (simp add: minus_divide_right)
- also have "... \<in> borel_measurable M"
- by (fast intro: affine_borel_measurable borel_measurable_square
- borel_measurable_add_borel_measurable
- borel_measurable_uminus_borel_measurable f g)
+ also have "... \<in> borel_measurable M"
+ using f g by (fast intro: affine_borel_measurable borel_measurable_square f g)
finally have 2: "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) \<in> borel_measurable M" .
show ?thesis
- apply (simp add: times_eq_sum_squares diff_minus)
- using 1 2 apply (simp add: borel_measurable_add_borel_measurable)
- done
+ apply (simp add: times_eq_sum_squares diff_minus)
+ using 1 2 by simp
qed
-lemma (in measure_space) borel_measurable_diff_borel_measurable:
+lemma (in sigma_algebra) borel_measurable_diff[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes f: "f \<in> borel_measurable M"
assumes g: "g \<in> borel_measurable M"
shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-unfolding diff_minus
- by (fast intro: borel_measurable_add_borel_measurable
- borel_measurable_uminus_borel_measurable f g)
+ unfolding diff_minus using assms by fast
-lemma (in measure_space) borel_measurable_setsum_borel_measurable:
- assumes s: "finite s"
- 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
-proof (induct s)
- case empty
- thus ?case
- by (simp add: borel_measurable_const)
-next
- case (insert x s)
- thus ?case
- by (auto simp add: borel_measurable_add_borel_measurable)
-qed
+lemma (in sigma_algebra) borel_measurable_setsum[simp, intro]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
+ assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
+proof cases
+ assume "finite S"
+ thus ?thesis using assms by induct auto
+qed simp
-lemma (in measure_space) borel_measurable_cong:
- assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
- shows "f \<in> borel_measurable M \<longleftrightarrow> g \<in> borel_measurable M"
-using assms unfolding in_borel_measurable by (simp cong: vimage_inter_cong)
-
-lemma (in measure_space) borel_measurable_inverse:
+lemma (in sigma_algebra) borel_measurable_inverse[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes "f \<in> borel_measurable M"
shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
- unfolding borel_measurable_ge_iff
-proof (safe, rule linorder_cases)
- fix a :: real assume "0 < a"
- { fix w
- from `0 < a` have "a \<le> inverse (f w) \<longleftrightarrow> 0 < f w \<and> f w \<le> 1 / a"
- by (metis inverse_eq_divide inverse_inverse_eq le_imp_inverse_le
- less_le_trans zero_less_divide_1_iff) }
- hence "{w \<in> space M. a \<le> inverse (f w)} =
- {w \<in> space M. 0 < f w} \<inter> {w \<in> space M. f w \<le> 1 / a}" by auto
- with Int assms[unfolded borel_measurable_gr_iff]
- assms[unfolded borel_measurable_le_iff]
- show "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M" by simp
-next
- fix a :: real assume "0 = a"
- { fix w have "a \<le> inverse (f w) \<longleftrightarrow> 0 \<le> f w"
- unfolding `0 = a`[symmetric] by auto }
- thus "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M"
- using assms[unfolded borel_measurable_ge_iff] by simp
-next
- fix a :: real assume "a < 0"
- { fix w
- from `a < 0` have "a \<le> inverse (f w) \<longleftrightarrow> f w \<le> 1 / a \<or> 0 \<le> f w"
- apply (cases "0 \<le> f w")
- apply (metis inverse_eq_divide linorder_not_le xt1(8) xt1(9)
- zero_le_divide_1_iff)
- apply (metis inverse_eq_divide inverse_inverse_eq inverse_le_imp_le_neg
- linorder_not_le order_refl order_trans)
- done }
- hence "{w \<in> space M. a \<le> inverse (f w)} =
- {w \<in> space M. f w \<le> 1 / a} \<union> {w \<in> space M. 0 \<le> f w}" by auto
- with Un assms[unfolded borel_measurable_ge_iff]
- assms[unfolded borel_measurable_le_iff]
- show "{w \<in> space M. a \<le> inverse (f w)} \<in> sets M" by simp
+ unfolding borel_measurable_iff_ge unfolding inverse_eq_divide
+proof safe
+ fix a :: real
+ have *: "{w \<in> space M. a \<le> 1 / f w} =
+ ({w \<in> space M. 0 < f w} \<inter> {w \<in> space M. a * f w \<le> 1}) \<union>
+ ({w \<in> space M. f w < 0} \<inter> {w \<in> space M. 1 \<le> a * f w}) \<union>
+ ({w \<in> space M. f w = 0} \<inter> {w \<in> space M. a \<le> 0})" by (auto simp: le_divide_eq)
+ show "{w \<in> space M. a \<le> 1 / f w} \<in> sets M" using assms unfolding *
+ by (auto intro!: Int Un)
qed
-lemma (in measure_space) borel_measurable_divide:
+lemma (in sigma_algebra) borel_measurable_divide[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
assumes "f \<in> borel_measurable M"
and "g \<in> borel_measurable M"
shows "(\<lambda>x. f x / g x) \<in> borel_measurable M"
unfolding field_divide_inverse
- by (rule borel_measurable_inverse borel_measurable_times_borel_measurable assms)+
+ by (rule borel_measurable_inverse borel_measurable_times assms)+
+
+lemma (in sigma_algebra) borel_measurable_max[intro, simp]:
+ fixes f g :: "'a \<Rightarrow> real"
+ assumes "f \<in> borel_measurable M"
+ assumes "g \<in> borel_measurable M"
+ shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
+ unfolding borel_measurable_iff_le
+proof safe
+ fix a
+ have "{x \<in> space M. max (g x) (f x) \<le> a} =
+ {x \<in> space M. g x \<le> a} \<inter> {x \<in> space M. f x \<le> a}" by auto
+ thus "{x \<in> space M. max (g x) (f x) \<le> a} \<in> sets M"
+ using assms unfolding borel_measurable_iff_le
+ by (auto intro!: Int)
+qed
+
+lemma (in sigma_algebra) borel_measurable_min[intro, simp]:
+ fixes f g :: "'a \<Rightarrow> real"
+ assumes "f \<in> borel_measurable M"
+ assumes "g \<in> borel_measurable M"
+ shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
+ unfolding borel_measurable_iff_ge
+proof safe
+ fix a
+ have "{x \<in> space M. a \<le> min (g x) (f x)} =
+ {x \<in> space M. a \<le> g x} \<inter> {x \<in> space M. a \<le> f x}" by auto
+ thus "{x \<in> space M. a \<le> min (g x) (f x)} \<in> sets M"
+ using assms unfolding borel_measurable_iff_ge
+ by (auto intro!: Int)
+qed
+
+lemma (in sigma_algebra) borel_measurable_abs[simp, intro]:
+ assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
+proof -
+ have *: "\<And>x. \<bar>f x\<bar> = max 0 (f x) + max 0 (- f x)" by (simp add: max_def)
+ show ?thesis unfolding * using assms by auto
+qed
+
+section "Borel space over the real line with infinity"
-lemma (in measure_space) borel_measurable_vimage:
- assumes borel: "f \<in> borel_measurable M"
- shows "f -` {X} \<inter> space M \<in> sets M"
-proof -
- have "{w \<in> space M. f w = X} = {w. f w = X} \<inter> space M" by auto
- with borel_measurable_eq_borel_measurable[OF borel borel_measurable_const, of X]
- show ?thesis unfolding vimage_def by simp
+lemma borel_space_Real_measurable:
+ "A \<in> sets borel_space \<Longrightarrow> Real -` A \<in> sets borel_space"
+proof (rule borel_measurable_translate)
+ fix B :: "pinfreal set" assume "open B"
+ then obtain T x where T: "open T" "Real ` (T \<inter> {0..}) = B - {\<omega>}" and
+ x: "\<omega> \<in> B \<Longrightarrow> 0 \<le> x" "\<omega> \<in> B \<Longrightarrow> {Real x <..} \<subseteq> B"
+ unfolding open_pinfreal_def by blast
+
+ have "Real -` B = Real -` (B - {\<omega>})" by auto
+ also have "\<dots> = Real -` (Real ` (T \<inter> {0..}))" using T by simp
+ also have "\<dots> = (if 0 \<in> T then T \<union> {.. 0} else T \<inter> {0..})"
+ apply (auto simp add: Real_eq_Real image_iff)
+ apply (rule_tac x="max 0 x" in bexI)
+ by (auto simp: max_def)
+ finally show "Real -` B \<in> sets borel_space"
+ using `open T` by auto
+qed simp
+
+lemma borel_space_real_measurable:
+ "A \<in> sets borel_space \<Longrightarrow> (real -` A :: pinfreal set) \<in> sets borel_space"
+proof (rule borel_measurable_translate)
+ fix B :: "real set" assume "open B"
+ { fix x have "0 < real x \<longleftrightarrow> (\<exists>r>0. x = Real r)" by (cases x) auto }
+ note Ex_less_real = this
+ have *: "real -` B = (if 0 \<in> B then real -` (B \<inter> {0 <..}) \<union> {0, \<omega>} else real -` (B \<inter> {0 <..}))"
+ by (force simp: Ex_less_real)
+
+ have "open (real -` (B \<inter> {0 <..}) :: pinfreal set)"
+ unfolding open_pinfreal_def using `open B`
+ by (auto intro!: open_Int exI[of _ "B \<inter> {0 <..}"] simp: image_iff Ex_less_real)
+ then show "(real -` B :: pinfreal set) \<in> sets borel_space" unfolding * by auto
+qed simp
+
+lemma (in sigma_algebra) borel_measurable_Real[intro, simp]:
+ assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
+ unfolding in_borel_measurable_borel_space
+proof safe
+ fix S :: "pinfreal set" assume "S \<in> sets borel_space"
+ from borel_space_Real_measurable[OF this]
+ have "(Real \<circ> f) -` S \<inter> space M \<in> sets M"
+ using assms
+ unfolding vimage_compose in_borel_measurable_borel_space
+ by auto
+ thus "(\<lambda>x. Real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
qed
-section "Monotone convergence"
-
-definition mono_convergent where
- "mono_convergent u f s \<equiv>
- (\<forall>x\<in>s. incseq (\<lambda>n. u n x)) \<and>
- (\<forall>x \<in> s. (\<lambda>i. u i x) ----> f x)"
-
-definition "upclose f g x = max (f x) (g x)"
+lemma (in sigma_algebra) borel_measurable_real[intro, simp]:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. real (f x)) \<in> borel_measurable M"
+ unfolding in_borel_measurable_borel_space
+proof safe
+ fix S :: "real set" assume "S \<in> sets borel_space"
+ from borel_space_real_measurable[OF this]
+ have "(real \<circ> f) -` S \<inter> space M \<in> sets M"
+ using assms
+ unfolding vimage_compose in_borel_measurable_borel_space
+ by auto
+ thus "(\<lambda>x. real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
+qed
-primrec mon_upclose where
-"mon_upclose 0 u = u 0" |
-"mon_upclose (Suc n) u = upclose (u (Suc n)) (mon_upclose n u)"
-
-lemma mono_convergentD:
- assumes "mono_convergent u f s" and "x \<in> s"
- shows "incseq (\<lambda>n. u n x)" and "(\<lambda>i. u i x) ----> f x"
- using assms unfolding mono_convergent_def by auto
+lemma (in sigma_algebra) borel_measurable_Real_eq:
+ assumes "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+ shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
+proof
+ have [simp]: "(\<lambda>x. Real (f x)) -` {\<omega>} \<inter> space M = {}"
+ by auto
+ assume "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
+ hence "(\<lambda>x. real (Real (f x))) \<in> borel_measurable M"
+ by (rule borel_measurable_real)
+ moreover have "\<And>x. x \<in> space M \<Longrightarrow> real (Real (f x)) = f x"
+ using assms by auto
+ ultimately show "f \<in> borel_measurable M"
+ by (simp cong: measurable_cong)
+qed auto
-lemma mono_convergentI:
- assumes "\<And>x. x \<in> s \<Longrightarrow> incseq (\<lambda>n. u n x)"
- assumes "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>i. u i x) ----> f x"
- shows "mono_convergent u f s"
- using assms unfolding mono_convergent_def by auto
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq_real:
+ "f \<in> borel_measurable M \<longleftrightarrow>
+ ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<omega>} \<inter> space M \<in> sets M)"
+proof safe
+ assume "f \<in> borel_measurable M"
+ then show "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
+ by (auto intro: borel_measurable_vimage borel_measurable_real)
+next
+ assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
+ have "f -` {\<omega>} \<inter> space M = {x\<in>space M. f x = \<omega>}" by auto
+ with * have **: "{x\<in>space M. f x = \<omega>} \<in> sets M" by simp
+ have f: "f = (\<lambda>x. if f x = \<omega> then \<omega> else Real (real (f x)))"
+ by (simp add: expand_fun_eq Real_real)
+ show "f \<in> borel_measurable M"
+ apply (subst f)
+ apply (rule measurable_If)
+ using * ** by auto
+qed
+
+lemma (in sigma_algebra) less_eq_ge_measurable:
+ fixes f :: "'a \<Rightarrow> 'c::linorder"
+ shows "{x\<in>space M. a < f x} \<in> sets M \<longleftrightarrow> {x\<in>space M. f x \<le> a} \<in> sets M"
+proof
+ assume "{x\<in>space M. f x \<le> a} \<in> sets M"
+ moreover have "{x\<in>space M. a < f x} = space M - {x\<in>space M. f x \<le> a}" by auto
+ ultimately show "{x\<in>space M. a < f x} \<in> sets M" by auto
+next
+ assume "{x\<in>space M. a < f x} \<in> sets M"
+ moreover have "{x\<in>space M. f x \<le> a} = space M - {x\<in>space M. a < f x}" by auto
+ ultimately show "{x\<in>space M. f x \<le> a} \<in> sets M" by auto
+qed
-lemma (in measure_space) mono_convergent_borel_measurable:
- assumes u: "!!n. u n \<in> borel_measurable M"
- assumes mc: "mono_convergent u f (space M)"
- shows "f \<in> borel_measurable M"
-proof -
- {
- fix a
- have 1: "!!w. w \<in> space M & f w <= a \<longleftrightarrow> w \<in> space M & (\<forall>i. u i w <= a)"
+lemma (in sigma_algebra) greater_eq_le_measurable:
+ fixes f :: "'a \<Rightarrow> 'c::linorder"
+ shows "{x\<in>space M. f x < a} \<in> sets M \<longleftrightarrow> {x\<in>space M. a \<le> f x} \<in> sets M"
+proof
+ assume "{x\<in>space M. a \<le> f x} \<in> sets M"
+ moreover have "{x\<in>space M. f x < a} = space M - {x\<in>space M. a \<le> f x}" by auto
+ ultimately show "{x\<in>space M. f x < a} \<in> sets M" by auto
+next
+ assume "{x\<in>space M. f x < a} \<in> sets M"
+ moreover have "{x\<in>space M. a \<le> f x} = space M - {x\<in>space M. f x < a}" by auto
+ ultimately show "{x\<in>space M. a \<le> f x} \<in> sets M" by auto
+qed
+
+lemma (in sigma_algebra) less_eq_le_pinfreal_measurable:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ shows "(\<forall>a. {x\<in>space M. a < f x} \<in> sets M) \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
+proof
+ assume a: "\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M"
+ show "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
+ proof
+ fix a show "{x \<in> space M. a < f x} \<in> sets M"
+ proof (cases a)
+ case (preal r)
+ have "{x\<in>space M. a < f x} = (\<Union>i. {x\<in>space M. a + inverse (of_nat (Suc i)) \<le> f x})"
proof safe
- fix w i
- assume w: "w \<in> space M" and f: "f w \<le> a"
- hence "u i w \<le> f w"
- by (auto intro: SEQ.incseq_le
- simp add: mc [unfolded mono_convergent_def])
- thus "u i w \<le> a" using f
+ fix x assume "a < f x" and [simp]: "x \<in> space M"
+ with ex_pinfreal_inverse_of_nat_Suc_less[of "f x - a"]
+ obtain n where "a + inverse (of_nat (Suc n)) < f x"
+ by (cases "f x", auto simp: pinfreal_minus_order)
+ then have "a + inverse (of_nat (Suc n)) \<le> f x" by simp
+ then show "x \<in> (\<Union>i. {x \<in> space M. a + inverse (of_nat (Suc i)) \<le> f x})"
by auto
next
- fix w
- assume w: "w \<in> space M" and u: "\<forall>i. u i w \<le> a"
- thus "f w \<le> a"
- by (metis LIMSEQ_le_const2 mc [unfolded mono_convergent_def])
+ fix i x assume [simp]: "x \<in> space M"
+ have "a < a + inverse (of_nat (Suc i))" using preal by auto
+ also assume "a + inverse (of_nat (Suc i)) \<le> f x"
+ finally show "a < f x" .
qed
- have "{w \<in> space M. f w \<le> a} = {w \<in> space M. (\<forall>i. u i w <= a)}"
- by (simp add: 1)
- also have "... = (\<Inter>i. {w \<in> space M. u i w \<le> a})"
- by auto
- also have "... \<in> sets M" using u
- by (auto simp add: borel_measurable_le_iff intro: countable_INT)
- finally have "{w \<in> space M. f w \<le> a} \<in> sets M" .
- }
- thus ?thesis
- by (auto simp add: borel_measurable_le_iff)
-qed
-
-lemma mono_convergent_le:
- assumes "mono_convergent u f s" and "t \<in> s"
- shows "u n t \<le> f t"
-using mono_convergentD[OF assms] by (auto intro!: incseq_le)
-
-lemma mon_upclose_ex:
- fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ('b\<Colon>linorder)"
- shows "\<exists>n \<le> m. mon_upclose m u x = u n x"
-proof (induct m)
- case (Suc m)
- then obtain n where "n \<le> m" and *: "mon_upclose m u x = u n x" by blast
- thus ?case
- proof (cases "u n x \<le> u (Suc m) x")
- case True with min_max.sup_absorb1 show ?thesis
- by (auto simp: * upclose_def intro!: exI[of _ "Suc m"])
- next
- case False
- with min_max.sup_absorb2 `n \<le> m` show ?thesis
- by (auto simp: * upclose_def intro!: exI[of _ n] min_max.sup_absorb2)
+ with a show ?thesis by auto
+ qed simp
qed
-qed simp
-
-lemma mon_upclose_all:
- fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ('b\<Colon>linorder)"
- assumes "m \<le> n"
- shows "u m x \<le> mon_upclose n u x"
-using assms proof (induct n)
- case 0 thus ?case by auto
next
- case (Suc n)
- show ?case
- proof (cases "m = Suc n")
- case True thus ?thesis by (simp add: upclose_def)
- next
- case False
- hence "m \<le> n" using `m \<le> Suc n` by simp
- from Suc.hyps[OF this]
- show ?thesis by (auto simp: upclose_def intro!: min_max.le_supI2)
+ assume a': "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
+ then have a: "\<forall>a. {x \<in> space M. f x \<le> a} \<in> sets M" unfolding less_eq_ge_measurable .
+ show "\<forall>a. {x \<in> space M. a \<le> f x} \<in> sets M" unfolding greater_eq_le_measurable[symmetric]
+ proof
+ fix a show "{x \<in> space M. f x < a} \<in> sets M"
+ proof (cases a)
+ case (preal r)
+ show ?thesis
+ proof cases
+ assume "a = 0" then show ?thesis by simp
+ next
+ assume "a \<noteq> 0"
+ have "{x\<in>space M. f x < a} = (\<Union>i. {x\<in>space M. f x \<le> a - inverse (of_nat (Suc i))})"
+ proof safe
+ fix x assume "f x < a" and [simp]: "x \<in> space M"
+ with ex_pinfreal_inverse_of_nat_Suc_less[of "a - f x"]
+ obtain n where "inverse (of_nat (Suc n)) < a - f x"
+ using preal by (cases "f x") auto
+ then have "f x \<le> a - inverse (of_nat (Suc n)) "
+ using preal by (cases "f x") (auto split: split_if_asm)
+ then show "x \<in> (\<Union>i. {x \<in> space M. f x \<le> a - inverse (of_nat (Suc i))})"
+ by auto
+ next
+ fix i x assume [simp]: "x \<in> space M"
+ assume "f x \<le> a - inverse (of_nat (Suc i))"
+ also have "\<dots> < a" using `a \<noteq> 0` preal by auto
+ finally show "f x < a" .
+ qed
+ with a show ?thesis by auto
+ qed
+ next
+ case infinite
+ have "f -` {\<omega>} \<inter> space M = (\<Inter>n. {x\<in>space M. of_nat n < f x})"
+ proof (safe, simp_all, safe)
+ fix x assume *: "\<forall>n::nat. Real (real n) < f x"
+ show "f x = \<omega>" proof (rule ccontr)
+ assume "f x \<noteq> \<omega>"
+ with real_arch_lt[of "real (f x)"] obtain n where "f x < of_nat n"
+ by (auto simp: pinfreal_noteq_omega_Ex)
+ with *[THEN spec, of n] show False by auto
+ qed
+ qed
+ with a' have \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" by auto
+ moreover have "{x \<in> space M. f x < a} = space M - f -` {\<omega>} \<inter> space M"
+ using infinite by auto
+ ultimately show ?thesis by auto
+ qed
qed
qed
-lemma mono_convergent_limit:
- fixes f :: "'a \<Rightarrow> real"
- assumes "mono_convergent u f s" and "x \<in> s" and "0 < r"
- shows "\<exists>N. \<forall>n\<ge>N. f x - u n x < r"
-proof -
- from LIMSEQ_D[OF mono_convergentD(2)[OF assms(1,2)] `0 < r`]
- obtain N where "\<And>n. N \<le> n \<Longrightarrow> \<bar> u n x - f x \<bar> < r" by auto
- with mono_convergent_le[OF assms(1,2)] `0 < r`
- show ?thesis by (auto intro!: exI[of _ N])
-qed
-
-lemma mon_upclose_le_mono_convergent:
- assumes mc: "\<And>n. mono_convergent (\<lambda>m. u m n) (f n) s" and "x \<in> s"
- and "incseq (\<lambda>n. f n x)"
- shows "mon_upclose n (u n) x <= f n x"
-proof -
- obtain m where *: "mon_upclose n (u n) x = u n m x" and "m \<le> n"
- using mon_upclose_ex[of n "u n" x] by auto
- note this(1)
- also have "u n m x \<le> f m x" using mono_convergent_le[OF assms(1,2)] .
- also have "... \<le> f n x" using assms(3) `m \<le> n` unfolding incseq_def by auto
- finally show ?thesis .
-qed
-
-lemma mon_upclose_mono_convergent:
- assumes mc_u: "\<And>n. mono_convergent (\<lambda>m. u m n) (f n) s"
- and mc_f: "mono_convergent f h s"
- shows "mono_convergent (\<lambda>n. mon_upclose n (u n)) h s"
-proof (rule mono_convergentI)
- fix x assume "x \<in> s"
- show "incseq (\<lambda>n. mon_upclose n (u n) x)" unfolding incseq_def
- proof safe
- fix m n :: nat assume "m \<le> n"
- obtain i where mon: "mon_upclose m (u m) x = u m i x" and "i \<le> m"
- using mon_upclose_ex[of m "u m" x] by auto
- hence "i \<le> n" using `m \<le> n` by auto
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_greater:
+ "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a < f x} \<in> sets M)"
+proof safe
+ fix a assume f: "f \<in> borel_measurable M"
+ have "{x\<in>space M. a < f x} = f -` {a <..} \<inter> space M" by auto
+ with f show "{x\<in>space M. a < f x} \<in> sets M"
+ by (auto intro!: measurable_sets)
+next
+ assume *: "\<forall>a. {x\<in>space M. a < f x} \<in> sets M"
+ hence **: "\<forall>a. {x\<in>space M. f x < a} \<in> sets M"
+ unfolding less_eq_le_pinfreal_measurable
+ unfolding greater_eq_le_measurable .
- note mon
- also have "u m i x \<le> u n i x"
- using mono_convergentD(1)[OF mc_u `x \<in> s`] `m \<le> n`
- unfolding incseq_def by auto
- also have "u n i x \<le> mon_upclose n (u n) x"
- using mon_upclose_all[OF `i \<le> n`, of "u n" x] .
- finally show "mon_upclose m (u m) x \<le> mon_upclose n (u n) x" .
- qed
-
- show "(\<lambda>i. mon_upclose i (u i) x) ----> h x"
- proof (rule LIMSEQ_I)
- fix r :: real assume "0 < r"
- hence "0 < r / 2" by auto
- from mono_convergent_limit[OF mc_f `x \<in> s` this]
- obtain N where f_h: "\<And>n. N \<le> n \<Longrightarrow> h x - f n x < r / 2" by auto
-
- from mono_convergent_limit[OF mc_u `x \<in> s` `0 < r / 2`]
- obtain N' where u_f: "\<And>n. N' \<le> n \<Longrightarrow> f N x - u n N x < r / 2" by auto
+ show "f \<in> borel_measurable M" unfolding borel_measurable_pinfreal_eq_real borel_measurable_iff_greater
+ proof safe
+ have "f -` {\<omega>} \<inter> space M = space M - {x\<in>space M. f x < \<omega>}" by auto
+ then show \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" using ** by auto
- show "\<exists>N. \<forall>n\<ge>N. norm (mon_upclose n (u n) x - h x) < r"
- proof (rule exI[of _ "max N N'"], safe)
- fix n assume "max N N' \<le> n"
- hence "N \<le> n" and "N' \<le> n" by auto
- hence "u n N x \<le> mon_upclose n (u n) x"
- using mon_upclose_all[of N n "u n" x] by auto
- moreover
- from add_strict_mono[OF u_f[OF `N' \<le> n`] f_h[OF order_refl]]
- have "h x - u n N x < r" by auto
- ultimately have "h x - mon_upclose n (u n) x < r" by auto
- moreover
- obtain i where "mon_upclose n (u n) x = u n i x"
- using mon_upclose_ex[of n "u n"] by blast
- with mono_convergent_le[OF mc_u `x \<in> s`, of n i]
- mono_convergent_le[OF mc_f `x \<in> s`, of i]
- have "mon_upclose n (u n) x \<le> h x" by auto
- ultimately
- show "norm (mon_upclose n (u n) x - h x) < r" by auto
- qed
+ fix a
+ have "{w \<in> space M. a < real (f w)} =
+ (if 0 \<le> a then {w\<in>space M. Real a < f w} - (f -` {\<omega>} \<inter> space M) else space M)"
+ proof (split split_if, safe del: notI)
+ fix x assume "0 \<le> a"
+ { assume "a < real (f x)" then show "Real a < f x" "x \<notin> f -` {\<omega>} \<inter> space M"
+ using `0 \<le> a` by (cases "f x", auto) }
+ { assume "Real a < f x" "x \<notin> f -` {\<omega>}" then show "a < real (f x)"
+ using `0 \<le> a` by (cases "f x", auto) }
+ next
+ fix x assume "\<not> 0 \<le> a" then show "a < real (f x)" by (cases "f x") auto
+ qed
+ then show "{w \<in> space M. a < real (f w)} \<in> sets M"
+ using \<omega> * by (auto intro!: Diff)
qed
qed
-lemma mono_conv_outgrow:
- assumes "incseq x" "x ----> y" "z < y"
- shows "\<exists>b. \<forall> a \<ge> b. z < x a"
-using assms
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_less:
+ "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x < a} \<in> sets M)"
+ using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable greater_eq_le_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_le:
+ "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x \<le> a} \<in> sets M)"
+ using borel_measurable_pinfreal_iff_greater unfolding less_eq_ge_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_ge:
+ "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
+ using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq_const:
+ fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M"
+ shows "{x\<in>space M. f x = c} \<in> sets M"
+proof -
+ have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
+ then show ?thesis using assms by (auto intro!: measurable_sets)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_neq_const:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ assumes "f \<in> borel_measurable M"
+ shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
+proof -
+ have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
+ then show ?thesis using assms by (auto intro!: measurable_sets)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_less[intro,simp]:
+ fixes f g :: "'a \<Rightarrow> pinfreal"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "{x \<in> space M. f x < g x} \<in> sets M"
+proof -
+ have "(\<lambda>x. real (f x)) \<in> borel_measurable M"
+ "(\<lambda>x. real (g x)) \<in> borel_measurable M"
+ using assms by (auto intro!: borel_measurable_real)
+ from borel_measurable_less[OF this]
+ have "{x \<in> space M. real (f x) < real (g x)} \<in> sets M" .
+ moreover have "{x \<in> space M. f x \<noteq> \<omega>} \<in> sets M" using f by (rule borel_measurable_pinfreal_neq_const)
+ moreover have "{x \<in> space M. g x = \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_eq_const)
+ moreover have "{x \<in> space M. g x \<noteq> \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_neq_const)
+ moreover have "{x \<in> space M. f x < g x} = ({x \<in> space M. g x = \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>}) \<union>
+ ({x \<in> space M. g x \<noteq> \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>} \<inter> {x \<in> space M. real (f x) < real (g x)})"
+ by (auto simp: real_of_pinfreal_strict_mono_iff)
+ ultimately show ?thesis by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_le[intro,simp]:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
+proof -
+ have "{x \<in> space M. f x \<le> g x} = space M - {x \<in> space M. g x < f x}" by auto
+ then show ?thesis using g f by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq[intro,simp]:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "{w \<in> space M. f w = g w} \<in> sets M"
+proof -
+ have "{x \<in> space M. f x = g x} = {x \<in> space M. g x \<le> f x} \<inter> {x \<in> space M. f x \<le> g x}" by auto
+ then show ?thesis using g f by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_neq[intro,simp]:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
proof -
- from assms have "y - z > 0" by simp
- hence A: "\<exists>n. (\<forall> m \<ge> n. \<bar> x m + - y \<bar> < y - z)" using assms
- unfolding incseq_def LIMSEQ_def dist_real_def diff_minus
- by simp
- have "\<forall>m. x m \<le> y" using incseq_le assms by auto
- hence B: "\<forall>m. \<bar> x m + - y \<bar> = y - x m"
- by (metis abs_if abs_minus_add_cancel less_iff_diff_less_0 linorder_not_le diff_minus)
- from A B show ?thesis by auto
+ have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}" by auto
+ thus ?thesis using f g by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_add[intro, simp]:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ assumes measure: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
+proof -
+ have *: "(\<lambda>x. f x + g x) =
+ (\<lambda>x. if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else Real (real (f x) + real (g x)))"
+ by (auto simp: expand_fun_eq pinfreal_noteq_omega_Ex)
+ show ?thesis using assms unfolding *
+ by (auto intro!: measurable_If)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_times[intro, simp]:
+ fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
+proof -
+ have *: "(\<lambda>x. f x * g x) =
+ (\<lambda>x. if f x = 0 then 0 else if g x = 0 then 0 else if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else
+ Real (real (f x) * real (g x)))"
+ by (auto simp: expand_fun_eq pinfreal_noteq_omega_Ex)
+ show ?thesis using assms unfolding *
+ by (auto intro!: measurable_If)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_setsum[simp, intro]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pinfreal"
+ assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
+proof cases
+ assume "finite S"
+ thus ?thesis using assms
+ by induct auto
+qed (simp add: borel_measurable_const)
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_min[intro, simp]:
+ fixes f g :: "'a \<Rightarrow> pinfreal"
+ assumes "f \<in> borel_measurable M"
+ assumes "g \<in> borel_measurable M"
+ shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
+ using assms unfolding min_def by (auto intro!: measurable_If)
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_max[intro]:
+ fixes f g :: "'a \<Rightarrow> pinfreal"
+ assumes "f \<in> borel_measurable M"
+ and "g \<in> borel_measurable M"
+ shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
+ using assms unfolding max_def by (auto intro!: measurable_If)
+
+lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]:
+ fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
+ assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(SUP i : A. f i) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
+ unfolding borel_measurable_pinfreal_iff_greater
+proof safe
+ fix a
+ have "{x\<in>space M. a < ?sup x} = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
+ by (auto simp: less_Sup_iff SUPR_def[where 'a=pinfreal] SUPR_fun_expand[where 'b=pinfreal])
+ then show "{x\<in>space M. a < ?sup x} \<in> sets M"
+ using assms by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
+ fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
+ assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(INF i : A. f i) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
+ unfolding borel_measurable_pinfreal_iff_less
+proof safe
+ fix a
+ have "{x\<in>space M. ?inf x < a} = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
+ by (auto simp: Inf_less_iff INFI_def[where 'a=pinfreal] INFI_fun_expand)
+ then show "{x\<in>space M. ?inf x < a} \<in> sets M"
+ using assms by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_diff:
+ fixes f g :: "'a \<Rightarrow> pinfreal"
+ assumes "f \<in> borel_measurable M"
+ assumes "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+ unfolding borel_measurable_pinfreal_iff_greater
+proof safe
+ fix a
+ have "{x \<in> space M. a < f x - g x} = {x \<in> space M. g x + a < f x}"
+ by (simp add: pinfreal_less_minus_iff)
+ then show "{x \<in> space M. a < f x - g x} \<in> sets M"
+ using assms by auto
qed
end
--- a/src/HOL/Probability/Caratheodory.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Probability/Caratheodory.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1,43 +1,28 @@
header {*Caratheodory Extension Theorem*}
theory Caratheodory
- imports Sigma_Algebra SeriesPlus
+ imports Sigma_Algebra Positive_Infinite_Real
begin
text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
subsection {* Measure Spaces *}
-text {*A measure assigns a nonnegative real to every measurable set.
- It is countably additive for disjoint sets.*}
-
-record 'a measure_space = "'a algebra" +
- measure:: "'a set \<Rightarrow> real"
-
-definition
- disjoint_family_on where
- "disjoint_family_on A S \<longleftrightarrow> (\<forall>m\<in>S. \<forall>n\<in>S. m \<noteq> n \<longrightarrow> A m \<inter> A n = {})"
-
-abbreviation
- "disjoint_family A \<equiv> disjoint_family_on A UNIV"
-
-definition
- positive where
- "positive M f \<longleftrightarrow> f {} = (0::real) & (\<forall>x \<in> sets M. 0 \<le> f x)"
+definition "positive f \<longleftrightarrow> f {} = (0::pinfreal)" -- "Positive is enforced by the type"
definition
additive where
- "additive M f \<longleftrightarrow>
- (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {}
+ "additive M f \<longleftrightarrow>
+ (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {}
\<longrightarrow> f (x \<union> y) = f x + f y)"
definition
countably_additive where
- "countably_additive M f \<longleftrightarrow>
- (\<forall>A. range A \<subseteq> sets M \<longrightarrow>
+ "countably_additive M f \<longleftrightarrow>
+ (\<forall>A. range A \<subseteq> sets M \<longrightarrow>
disjoint_family A \<longrightarrow>
- (\<Union>i. A i) \<in> sets M \<longrightarrow>
- (\<lambda>n. f (A n)) sums f (\<Union>i. A i))"
+ (\<Union>i. A i) \<in> sets M \<longrightarrow>
+ (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i))"
definition
increasing where
@@ -45,90 +30,58 @@
definition
subadditive where
- "subadditive M f \<longleftrightarrow>
- (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {}
+ "subadditive M f \<longleftrightarrow>
+ (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {}
\<longrightarrow> f (x \<union> y) \<le> f x + f y)"
definition
countably_subadditive where
- "countably_subadditive M f \<longleftrightarrow>
- (\<forall>A. range A \<subseteq> sets M \<longrightarrow>
+ "countably_subadditive M f \<longleftrightarrow>
+ (\<forall>A. range A \<subseteq> sets M \<longrightarrow>
disjoint_family A \<longrightarrow>
- (\<Union>i. A i) \<in> sets M \<longrightarrow>
- summable (f o A) \<longrightarrow>
- f (\<Union>i. A i) \<le> suminf (\<lambda>n. f (A n)))"
+ (\<Union>i. A i) \<in> sets M \<longrightarrow>
+ f (\<Union>i. A i) \<le> psuminf (\<lambda>n. f (A n)))"
definition
lambda_system where
- "lambda_system M f =
+ "lambda_system M f =
{l. l \<in> sets M & (\<forall>x \<in> sets M. f (l \<inter> x) + f ((space M - l) \<inter> x) = f x)}"
definition
outer_measure_space where
- "outer_measure_space M f \<longleftrightarrow>
- positive M f & increasing M f & countably_subadditive M f"
+ "outer_measure_space M f \<longleftrightarrow>
+ positive f \<and> increasing M f \<and> countably_subadditive M f"
definition
measure_set where
"measure_set M f X =
- {r . \<exists>A. range A \<subseteq> sets M & disjoint_family A & X \<subseteq> (\<Union>i. A i) & (f \<circ> A) sums r}"
-
+ {r . \<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>\<^isub>\<infinity> i. f (A i)) = r}"
locale measure_space = sigma_algebra +
- assumes positive: "!!a. a \<in> sets M \<Longrightarrow> 0 \<le> measure M a"
- and empty_measure [simp]: "measure M {} = (0::real)"
- and ca: "countably_additive M (measure M)"
-
-subsection {* Basic Lemmas *}
-
-lemma positive_imp_0: "positive M f \<Longrightarrow> f {} = 0"
- by (simp add: positive_def)
-
-lemma positive_imp_pos: "positive M f \<Longrightarrow> x \<in> sets M \<Longrightarrow> 0 \<le> f x"
- by (simp add: positive_def)
+ fixes \<mu> :: "'a set \<Rightarrow> pinfreal"
+ assumes empty_measure [simp]: "\<mu> {} = 0"
+ and ca: "countably_additive M \<mu>"
lemma increasingD:
"increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M \<Longrightarrow> f x \<le> f y"
by (auto simp add: increasing_def)
lemma subadditiveD:
- "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M
+ "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M
\<Longrightarrow> f (x \<union> y) \<le> f x + f y"
by (auto simp add: subadditive_def)
lemma additiveD:
- "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M
+ "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M
\<Longrightarrow> f (x \<union> y) = f x + f y"
by (auto simp add: additive_def)
lemma countably_additiveD:
"countably_additive M f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A
- \<Longrightarrow> (\<Union>i. A i) \<in> sets M \<Longrightarrow> (\<lambda>n. f (A n)) sums f (\<Union>i. A i)"
+ \<Longrightarrow> (\<Union>i. A i) \<in> sets M \<Longrightarrow> (\<Sum>\<^isub>\<infinity> n. f (A n)) = f (\<Union>i. A i)"
by (simp add: countably_additive_def)
-lemma Int_Diff_disjoint: "A \<inter> B \<inter> (A - B) = {}"
- by blast
-
-lemma Int_Diff_Un: "A \<inter> B \<union> (A - B) = A"
- by blast
-
-lemma disjoint_family_subset:
- "disjoint_family A \<Longrightarrow> (!!x. B x \<subseteq> A x) \<Longrightarrow> disjoint_family B"
- by (force simp add: disjoint_family_on_def)
-
-subsection {* A Two-Element Series *}
-
-definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "
- where "binaryset A B = (\<lambda>\<^isup>x. {})(0 := A, Suc 0 := B)"
-
-lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}"
- apply (simp add: binaryset_def)
- apply (rule set_ext)
- apply (auto simp add: image_iff)
- done
-
-lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"
- by (simp add: UNION_eq_Union_image range_binaryset_eq)
+section "Extend binary sets"
lemma LIMSEQ_binaryset:
assumes f: "f {} = 0"
@@ -153,17 +106,31 @@
lemma binaryset_sums:
assumes f: "f {} = 0"
shows "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
- by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f])
+ by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f])
lemma suminf_binaryset_eq:
"f {} = 0 \<Longrightarrow> suminf (\<lambda>n. f (binaryset A B n)) = f A + f B"
by (metis binaryset_sums sums_unique)
+lemma binaryset_psuminf:
+ assumes "f {} = 0"
+ shows "(\<Sum>\<^isub>\<infinity> n. f (binaryset A B n)) = f A + f B" (is "?suminf = ?sum")
+proof -
+ have *: "{..<2} = {0, 1::nat}" by auto
+ have "\<forall>n\<ge>2. f (binaryset A B n) = 0"
+ unfolding binaryset_def
+ using assms by auto
+ hence "?suminf = (\<Sum>N<2. f (binaryset A B N))"
+ by (rule psuminf_finite)
+ also have "... = ?sum" unfolding * binaryset_def
+ by simp
+ finally show ?thesis .
+qed
subsection {* Lambda Systems *}
lemma (in algebra) lambda_system_eq:
- "lambda_system M f =
+ "lambda_system M f =
{l. l \<in> sets M & (\<forall>x \<in> sets M. f (x \<inter> l) + f (x - l) = f x)}"
proof -
have [simp]: "!!l x. l \<in> sets M \<Longrightarrow> x \<in> sets M \<Longrightarrow> (space M - l) \<inter> x = x - l"
@@ -173,28 +140,28 @@
qed
lemma (in algebra) lambda_system_empty:
- "positive M f \<Longrightarrow> {} \<in> lambda_system M f"
- by (auto simp add: positive_def lambda_system_eq)
+ "positive f \<Longrightarrow> {} \<in> lambda_system M f"
+ by (auto simp add: positive_def lambda_system_eq)
lemma lambda_system_sets:
"x \<in> lambda_system M f \<Longrightarrow> x \<in> sets M"
by (simp add: lambda_system_def)
lemma (in algebra) lambda_system_Compl:
- fixes f:: "'a set \<Rightarrow> real"
+ fixes f:: "'a set \<Rightarrow> pinfreal"
assumes x: "x \<in> lambda_system M f"
shows "space M - x \<in> lambda_system M f"
proof -
have "x \<subseteq> space M"
by (metis sets_into_space lambda_system_sets x)
hence "space M - (space M - x) = x"
- by (metis double_diff equalityE)
+ by (metis double_diff equalityE)
with x show ?thesis
- by (force simp add: lambda_system_def)
+ by (force simp add: lambda_system_def ac_simps)
qed
lemma (in algebra) lambda_system_Int:
- fixes f:: "'a set \<Rightarrow> real"
+ fixes f:: "'a set \<Rightarrow> pinfreal"
assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
shows "x \<inter> y \<in> lambda_system M f"
proof -
@@ -213,42 +180,42 @@
ultimately
have ey: "f (u - x \<inter> y) = f (u \<inter> y - x) + f (u - y)" using fy
by force
- have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y)
+ have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y)
= (f (u \<inter> (x \<inter> y)) + f (u \<inter> y - x)) + f (u - y)"
- by (simp add: ey)
+ by (simp add: ey ac_simps)
also have "... = (f ((u \<inter> y) \<inter> x) + f (u \<inter> y - x)) + f (u - y)"
- by (simp add: Int_ac)
+ by (simp add: Int_ac)
also have "... = f (u \<inter> y) + f (u - y)"
using fx [THEN bspec, of "u \<inter> y"] Int y u
by force
also have "... = f u"
- by (metis fy u)
+ by (metis fy u)
finally show "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) = f u" .
qed
qed
lemma (in algebra) lambda_system_Un:
- fixes f:: "'a set \<Rightarrow> real"
+ fixes f:: "'a set \<Rightarrow> pinfreal"
assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
shows "x \<union> y \<in> lambda_system M f"
proof -
have "(space M - x) \<inter> (space M - y) \<in> sets M"
- by (metis Diff_Un Un compl_sets lambda_system_sets xl yl)
+ by (metis Diff_Un Un compl_sets lambda_system_sets xl yl)
moreover
have "x \<union> y = space M - ((space M - x) \<inter> (space M - y))"
by auto (metis subsetD lambda_system_sets sets_into_space xl yl)+
ultimately show ?thesis
- by (metis lambda_system_Compl lambda_system_Int xl yl)
+ by (metis lambda_system_Compl lambda_system_Int xl yl)
qed
lemma (in algebra) lambda_system_algebra:
- "positive M f \<Longrightarrow> algebra (M (|sets := lambda_system M f|))"
- apply (auto simp add: algebra_def)
+ "positive f \<Longrightarrow> algebra (M (|sets := lambda_system M f|))"
+ apply (auto simp add: algebra_def)
apply (metis lambda_system_sets set_mp sets_into_space)
apply (metis lambda_system_empty)
apply (metis lambda_system_Compl)
- apply (metis lambda_system_Un)
+ apply (metis lambda_system_Un)
done
lemma (in algebra) lambda_system_strong_additive:
@@ -259,19 +226,13 @@
have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast
moreover
have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast
- moreover
+ moreover
have "(z \<inter> (x \<union> y)) \<in> sets M"
- by (metis Int Un lambda_system_sets xl yl z)
+ by (metis Int Un lambda_system_sets xl yl z)
ultimately show ?thesis using xl yl
by (simp add: lambda_system_eq)
qed
-lemma (in algebra) Int_space_eq1 [simp]: "x \<in> sets M \<Longrightarrow> space M \<inter> x = x"
- by (metis Int_absorb1 sets_into_space)
-
-lemma (in algebra) Int_space_eq2 [simp]: "x \<in> sets M \<Longrightarrow> x \<inter> space M = x"
- by (metis Int_absorb2 sets_into_space)
-
lemma (in algebra) lambda_system_additive:
"additive (M (|sets := lambda_system M f|)) f"
proof (auto simp add: additive_def)
@@ -279,14 +240,14 @@
assume disj: "x \<inter> y = {}"
and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
hence "x \<in> sets M" "y \<in> sets M" by (blast intro: lambda_system_sets)+
- thus "f (x \<union> y) = f x + f y"
+ thus "f (x \<union> y) = f x + f y"
using lambda_system_strong_additive [OF top disj xl yl]
by (simp add: Un)
qed
lemma (in algebra) countably_subadditive_subadditive:
- assumes f: "positive M f" and cs: "countably_subadditive M f"
+ assumes f: "positive f" and cs: "countably_subadditive M f"
shows "subadditive M f"
proof (auto simp add: subadditive_def)
fix x y
@@ -295,159 +256,80 @@
by (auto simp add: disjoint_family_on_def binaryset_def)
hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
(\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
- summable (f o (binaryset x y)) \<longrightarrow>
- f (\<Union>i. binaryset x y i) \<le> suminf (\<lambda>n. f (binaryset x y n))"
+ f (\<Union>i. binaryset x y i) \<le> (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
using cs by (simp add: countably_subadditive_def)
hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
- summable (f o (binaryset x y)) \<longrightarrow>
- f (x \<union> y) \<le> suminf (\<lambda>n. f (binaryset x y n))"
+ f (x \<union> y) \<le> (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
by (simp add: range_binaryset_eq UN_binaryset_eq)
- thus "f (x \<union> y) \<le> f x + f y" using f x y binaryset_sums
- by (auto simp add: Un sums_iff positive_def o_def)
-qed
-
-
-definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set "
- where "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
-
-lemma finite_UN_disjointed_eq: "(\<Union>i\<in>{0..<n}. disjointed A i) = (\<Union>i\<in>{0..<n}. A i)"
-proof (induct n)
- case 0 show ?case by simp
-next
- case (Suc n)
- thus ?case by (simp add: atLeastLessThanSuc disjointed_def)
+ thus "f (x \<union> y) \<le> f x + f y" using f x y
+ by (auto simp add: Un o_def binaryset_psuminf positive_def)
qed
-lemma UN_disjointed_eq: "(\<Union>i. disjointed A i) = (\<Union>i. A i)"
- apply (rule UN_finite2_eq [where k=0])
- apply (simp add: finite_UN_disjointed_eq)
- done
-
-lemma less_disjoint_disjointed: "m<n \<Longrightarrow> disjointed A m \<inter> disjointed A n = {}"
- by (auto simp add: disjointed_def)
-
-lemma disjoint_family_disjointed: "disjoint_family (disjointed A)"
- by (simp add: disjoint_family_on_def)
- (metis neq_iff Int_commute less_disjoint_disjointed)
-
-lemma disjointed_subset: "disjointed A n \<subseteq> A n"
- by (auto simp add: disjointed_def)
-
-
-lemma (in algebra) UNION_in_sets:
- fixes A:: "nat \<Rightarrow> 'a set"
- assumes A: "range A \<subseteq> sets M "
- shows "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
-proof (induct n)
- case 0 show ?case by simp
-next
- case (Suc n)
- thus ?case
- by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff)
-qed
-
-lemma (in algebra) range_disjointed_sets:
- assumes A: "range A \<subseteq> sets M "
- shows "range (disjointed A) \<subseteq> sets M"
-proof (auto simp add: disjointed_def)
- fix n
- show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> sets M" using UNION_in_sets
- by (metis A Diff UNIV_I image_subset_iff)
-qed
-
-lemma sigma_algebra_disjoint_iff:
- "sigma_algebra M \<longleftrightarrow>
- algebra M &
- (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow>
- (\<Union>i::nat. A i) \<in> sets M)"
-proof (auto simp add: sigma_algebra_iff)
- fix A :: "nat \<Rightarrow> 'a set"
- assume M: "algebra M"
- and A: "range A \<subseteq> sets M"
- and UnA: "\<forall>A. range A \<subseteq> sets M \<longrightarrow>
- disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
- hence "range (disjointed A) \<subseteq> sets M \<longrightarrow>
- disjoint_family (disjointed A) \<longrightarrow>
- (\<Union>i. disjointed A i) \<in> sets M" by blast
- hence "(\<Union>i. disjointed A i) \<in> sets M"
- by (simp add: algebra.range_disjointed_sets M A disjoint_family_disjointed)
- thus "(\<Union>i::nat. A i) \<in> sets M" by (simp add: UN_disjointed_eq)
-qed
-
-
lemma (in algebra) additive_sum:
fixes A:: "nat \<Rightarrow> 'a set"
- assumes f: "positive M f" and ad: "additive M f"
+ assumes f: "positive f" and ad: "additive M f"
and A: "range A \<subseteq> sets M"
and disj: "disjoint_family A"
- shows "setsum (f o A) {0..<n} = f (\<Union>i\<in>{0..<n}. A i)"
+ shows "setsum (f \<circ> A) {0..<n} = f (\<Union>i\<in>{0..<n}. A i)"
proof (induct n)
- case 0 show ?case using f by (simp add: positive_def)
+ case 0 show ?case using f by (simp add: positive_def)
next
- case (Suc n)
- have "A n \<inter> (\<Union>i\<in>{0..<n}. A i) = {}" using disj
+ case (Suc n)
+ have "A n \<inter> (\<Union>i\<in>{0..<n}. A i) = {}" using disj
by (auto simp add: disjoint_family_on_def neq_iff) blast
- moreover
- have "A n \<in> sets M" using A by blast
+ moreover
+ have "A n \<in> sets M" using A by blast
moreover have "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
by (metis A UNION_in_sets atLeast0LessThan)
- moreover
+ moreover
ultimately have "f (A n \<union> (\<Union>i\<in>{0..<n}. A i)) = f (A n) + f(\<Union>i\<in>{0..<n}. A i)"
- using ad UNION_in_sets A by (auto simp add: additive_def)
+ using ad UNION_in_sets A by (auto simp add: additive_def)
with Suc.hyps show ?case using ad
- by (auto simp add: atLeastLessThanSuc additive_def)
+ by (auto simp add: atLeastLessThanSuc additive_def)
qed
lemma countably_subadditiveD:
"countably_subadditive M f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow>
- (\<Union>i. A i) \<in> sets M \<Longrightarrow> summable (f o A) \<Longrightarrow> f (\<Union>i. A i) \<le> suminf (f o A)"
+ (\<Union>i. A i) \<in> sets M \<Longrightarrow> f (\<Union>i. A i) \<le> psuminf (f o A)"
by (auto simp add: countably_subadditive_def o_def)
-lemma (in algebra) increasing_additive_summable:
- fixes A:: "nat \<Rightarrow> 'a set"
- assumes f: "positive M f" and ad: "additive M f"
+lemma (in algebra) increasing_additive_bound:
+ fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pinfreal"
+ assumes f: "positive f" and ad: "additive M f"
and inc: "increasing M f"
and A: "range A \<subseteq> sets M"
and disj: "disjoint_family A"
- shows "summable (f o A)"
-proof (rule pos_summable)
- fix n
- show "0 \<le> (f \<circ> A) n" using f A
- by (force simp add: positive_def)
- next
- fix n
- have "setsum (f \<circ> A) {0..<n} = f (\<Union>i\<in>{0..<n}. A i)"
- by (rule additive_sum [OF f ad A disj])
+ shows "psuminf (f \<circ> A) \<le> f (space M)"
+proof (safe intro!: psuminf_bound)
+ fix N
+ have "setsum (f \<circ> A) {0..<N} = f (\<Union>i\<in>{0..<N}. A i)"
+ by (rule additive_sum [OF f ad A disj])
also have "... \<le> f (space M)" using space_closed A
- by (blast intro: increasingD [OF inc] UNION_in_sets top)
- finally show "setsum (f \<circ> A) {0..<n} \<le> f (space M)" .
+ by (blast intro: increasingD [OF inc] UNION_in_sets top)
+ finally show "setsum (f \<circ> A) {..<N} \<le> f (space M)" by (simp add: atLeast0LessThan)
qed
-lemma lambda_system_positive:
- "positive M f \<Longrightarrow> positive (M (|sets := lambda_system M f|)) f"
- by (simp add: positive_def lambda_system_def)
-
lemma lambda_system_increasing:
"increasing M f \<Longrightarrow> increasing (M (|sets := lambda_system M f|)) f"
- by (simp add: increasing_def lambda_system_def)
+ by (simp add: increasing_def lambda_system_def)
lemma (in algebra) lambda_system_strong_sum:
- fixes A:: "nat \<Rightarrow> 'a set"
- assumes f: "positive M f" and a: "a \<in> sets M"
+ fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> pinfreal"
+ assumes f: "positive f" and a: "a \<in> sets M"
and A: "range A \<subseteq> lambda_system M f"
and disj: "disjoint_family A"
shows "(\<Sum>i = 0..<n. f (a \<inter>A i)) = f (a \<inter> (\<Union>i\<in>{0..<n}. A i))"
proof (induct n)
- case 0 show ?case using f by (simp add: positive_def)
+ case 0 show ?case using f by (simp add: positive_def)
next
- case (Suc n)
+ case (Suc n)
have 2: "A n \<inter> UNION {0..<n} A = {}" using disj
- by (force simp add: disjoint_family_on_def neq_iff)
+ by (force simp add: disjoint_family_on_def neq_iff)
have 3: "A n \<in> lambda_system M f" using A
by blast
have 4: "UNION {0..<n} A \<in> lambda_system M f"
- using A algebra.UNION_in_sets [OF local.lambda_system_algebra [OF f]]
+ using A algebra.UNION_in_sets [OF local.lambda_system_algebra, of f, OF f]
by simp
from Suc.hyps show ?case
by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4])
@@ -458,89 +340,77 @@
assumes oms: "outer_measure_space M f"
and A: "range A \<subseteq> lambda_system M f"
and disj: "disjoint_family A"
- shows "(\<Union>i. A i) \<in> lambda_system M f & (f \<circ> A) sums f (\<Union>i. A i)"
+ shows "(\<Union>i. A i) \<in> lambda_system M f \<and> psuminf (f \<circ> A) = f (\<Union>i. A i)"
proof -
- have pos: "positive M f" and inc: "increasing M f"
- and csa: "countably_subadditive M f"
+ have pos: "positive f" and inc: "increasing M f"
+ and csa: "countably_subadditive M f"
by (metis oms outer_measure_space_def)+
have sa: "subadditive M f"
- by (metis countably_subadditive_subadditive csa pos)
- have A': "range A \<subseteq> sets (M(|sets := lambda_system M f|))" using A
+ by (metis countably_subadditive_subadditive csa pos)
+ have A': "range A \<subseteq> sets (M(|sets := lambda_system M f|))" using A
by simp
have alg_ls: "algebra (M(|sets := lambda_system M f|))"
- by (rule lambda_system_algebra [OF pos])
+ by (rule lambda_system_algebra) (rule pos)
have A'': "range A \<subseteq> sets M"
by (metis A image_subset_iff lambda_system_sets)
- have sumfA: "summable (f \<circ> A)"
- by (metis algebra.increasing_additive_summable [OF alg_ls]
- lambda_system_positive lambda_system_additive lambda_system_increasing
- A' oms outer_measure_space_def disj)
+
have U_in: "(\<Union>i. A i) \<in> sets M"
by (metis A'' countable_UN)
- have U_eq: "f (\<Union>i. A i) = suminf (f o A)"
+ have U_eq: "f (\<Union>i. A i) = psuminf (f o A)"
proof (rule antisym)
- show "f (\<Union>i. A i) \<le> suminf (f \<circ> A)"
- by (rule countably_subadditiveD [OF csa A'' disj U_in sumfA])
- show "suminf (f \<circ> A) \<le> f (\<Union>i. A i)"
- by (rule suminf_le [OF sumfA])
+ show "f (\<Union>i. A i) \<le> psuminf (f \<circ> A)"
+ by (rule countably_subadditiveD [OF csa A'' disj U_in])
+ show "psuminf (f \<circ> A) \<le> f (\<Union>i. A i)"
+ by (rule psuminf_bound, unfold atLeast0LessThan[symmetric])
(metis algebra.additive_sum [OF alg_ls] pos disj UN_Un Un_UNIV_right
- lambda_system_positive lambda_system_additive
- subset_Un_eq increasingD [OF inc] A' A'' UNION_in_sets U_in)
+ lambda_system_additive subset_Un_eq increasingD [OF inc]
+ A' A'' UNION_in_sets U_in)
qed
{
- fix a
- assume a [iff]: "a \<in> sets M"
+ fix a
+ assume a [iff]: "a \<in> sets M"
have "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) = f a"
proof -
- have summ: "summable (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A)" using pos A''
- apply -
- apply (rule summable_comparison_test [OF _ sumfA])
- apply (rule_tac x="0" in exI)
- apply (simp add: positive_def)
- apply (auto simp add: )
- apply (subst abs_of_nonneg)
- apply (metis A'' Int UNIV_I a image_subset_iff)
- apply (blast intro: increasingD [OF inc])
- done
show ?thesis
proof (rule antisym)
have "range (\<lambda>i. a \<inter> A i) \<subseteq> sets M" using A''
by blast
- moreover
+ moreover
have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj
- by (auto simp add: disjoint_family_on_def)
- moreover
+ by (auto simp add: disjoint_family_on_def)
+ moreover
have "a \<inter> (\<Union>i. A i) \<in> sets M"
by (metis Int U_in a)
- ultimately
- have "f (a \<inter> (\<Union>i. A i)) \<le> suminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A)"
- using countably_subadditiveD [OF csa, of "(\<lambda>i. a \<inter> A i)"] summ
- by (simp add: o_def)
- moreover
- have "suminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) \<le> f a - f (a - (\<Union>i. A i))"
- proof (rule suminf_le [OF summ])
+ ultimately
+ have "f (a \<inter> (\<Union>i. A i)) \<le> psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A)"
+ using countably_subadditiveD [OF csa, of "(\<lambda>i. a \<inter> A i)"]
+ by (simp add: o_def)
+ hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le>
+ psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) + f (a - (\<Union>i. A i))"
+ by (rule add_right_mono)
+ moreover
+ have "psuminf (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) + f (a - (\<Union>i. A i)) \<le> f a"
+ proof (safe intro!: psuminf_bound_add)
fix n
have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
- by (metis A'' UNION_in_sets)
+ by (metis A'' UNION_in_sets)
have le_fa: "f (UNION {0..<n} A \<inter> a) \<le> f a" using A''
by (blast intro: increasingD [OF inc] A'' UNION_in_sets)
have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system M f"
- using algebra.UNION_in_sets [OF lambda_system_algebra [OF pos]]
- by (simp add: A)
- hence eq_fa: "f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i)) = f a"
+ using algebra.UNION_in_sets [OF lambda_system_algebra [of f, OF pos]]
+ by (simp add: A)
+ hence eq_fa: "f a = f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i))"
by (simp add: lambda_system_eq UNION_in)
have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
- by (blast intro: increasingD [OF inc] UNION_eq_Union_image
+ by (blast intro: increasingD [OF inc] UNION_eq_Union_image
UNION_in U_in)
- thus "setsum (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) {0..<n} \<le> f a - f (a - (\<Union>i. A i))"
- using eq_fa
- by (simp add: suminf_le [OF summ] lambda_system_strong_sum pos
- a A disj)
+ thus "setsum (f \<circ> (\<lambda>i. a \<inter> i) \<circ> A) {..<n} + f (a - (\<Union>i. A i)) \<le> f a"
+ by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
qed
- ultimately show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a"
- by arith
+ ultimately show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a"
+ by (rule order_trans)
next
- have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))"
+ have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))"
by (blast intro: increasingD [OF inc] U_in)
also have "... \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))"
by (blast intro: subadditiveD [OF sa] U_in)
@@ -549,68 +419,54 @@
qed
}
thus ?thesis
- by (simp add: lambda_system_eq sums_iff U_eq U_in sumfA)
+ by (simp add: lambda_system_eq sums_iff U_eq U_in)
qed
lemma (in sigma_algebra) caratheodory_lemma:
assumes oms: "outer_measure_space M f"
- shows "measure_space (|space = space M, sets = lambda_system M f, measure = f|)"
+ shows "measure_space (|space = space M, sets = lambda_system M f|) f"
proof -
- have pos: "positive M f"
+ have pos: "positive f"
by (metis oms outer_measure_space_def)
- have alg: "algebra (|space = space M, sets = lambda_system M f, measure = f|)"
- using lambda_system_algebra [OF pos]
- by (simp add: algebra_def)
- then moreover
- have "sigma_algebra (|space = space M, sets = lambda_system M f, measure = f|)"
+ have alg: "algebra (|space = space M, sets = lambda_system M f|)"
+ using lambda_system_algebra [of f, OF pos]
+ by (simp add: algebra_def)
+ then moreover
+ have "sigma_algebra (|space = space M, sets = lambda_system M f|)"
using lambda_system_caratheodory [OF oms]
- by (simp add: sigma_algebra_disjoint_iff)
- moreover
- have "measure_space_axioms (|space = space M, sets = lambda_system M f, measure = f|)"
+ by (simp add: sigma_algebra_disjoint_iff)
+ moreover
+ have "measure_space_axioms (|space = space M, sets = lambda_system M f|) f"
using pos lambda_system_caratheodory [OF oms]
- by (simp add: measure_space_axioms_def positive_def lambda_system_sets
- countably_additive_def o_def)
- ultimately
+ by (simp add: measure_space_axioms_def positive_def lambda_system_sets
+ countably_additive_def o_def)
+ ultimately
show ?thesis
- by intro_locales (auto simp add: sigma_algebra_def)
+ by intro_locales (auto simp add: sigma_algebra_def)
qed
lemma (in algebra) inf_measure_nonempty:
- assumes f: "positive M f" and b: "b \<in> sets M" and a: "a \<subseteq> b"
+ assumes f: "positive f" and b: "b \<in> sets M" and a: "a \<subseteq> b"
shows "f b \<in> measure_set M f a"
proof -
- have "(f \<circ> (\<lambda>i. {})(0 := b)) sums setsum (f \<circ> (\<lambda>i. {})(0 := b)) {0..<1::nat}"
- by (rule series_zero) (simp add: positive_imp_0 [OF f])
- also have "... = f b"
+ have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = setsum (f \<circ> (\<lambda>i. {})(0 := b)) {..<1::nat}"
+ by (rule psuminf_finite) (simp add: f[unfolded positive_def])
+ also have "... = f b"
by simp
- finally have "(f \<circ> (\<lambda>i. {})(0 := b)) sums f b" .
- thus ?thesis using a
- by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"]
- simp add: measure_set_def disjoint_family_on_def b split_if_mem2)
-qed
-
-lemma (in algebra) inf_measure_pos0:
- "positive M f \<Longrightarrow> x \<in> measure_set M f a \<Longrightarrow> 0 \<le> x"
-apply (auto simp add: positive_def measure_set_def sums_iff intro!: suminf_ge_zero)
-apply blast
-done
-
-lemma (in algebra) inf_measure_pos:
- shows "positive M f \<Longrightarrow> x \<subseteq> space M \<Longrightarrow> 0 \<le> Inf (measure_set M f x)"
-apply (rule Inf_greatest)
-apply (metis emptyE inf_measure_nonempty top)
-apply (metis inf_measure_pos0)
-done
+ finally have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = f b" .
+ thus ?thesis using a b
+ by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"]
+ simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def)
+qed
lemma (in algebra) additive_increasing:
- assumes posf: "positive M f" and addf: "additive M f"
+ assumes posf: "positive f" and addf: "additive M f"
shows "increasing M f"
-proof (auto simp add: increasing_def)
+proof (auto simp add: increasing_def)
fix x y
assume xy: "x \<in> sets M" "y \<in> sets M" "x \<subseteq> y"
- have "f x \<le> f x + f (y-x)" using posf
- by (simp add: positive_def) (metis Diff xy(1,2))
+ have "f x \<le> f x + f (y-x)" ..
also have "... = f (x \<union> (y-x))" using addf
by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
also have "... = f y"
@@ -619,42 +475,42 @@
qed
lemma (in algebra) countably_additive_additive:
- assumes posf: "positive M f" and ca: "countably_additive M f"
+ assumes posf: "positive f" and ca: "countably_additive M f"
shows "additive M f"
-proof (auto simp add: additive_def)
+proof (auto simp add: additive_def)
fix x y
assume x: "x \<in> sets M" and y: "y \<in> sets M" and "x \<inter> y = {}"
hence "disjoint_family (binaryset x y)"
- by (auto simp add: disjoint_family_on_def binaryset_def)
- hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
- (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
- f (\<Union>i. binaryset x y i) = suminf (\<lambda>n. f (binaryset x y n))"
+ by (auto simp add: disjoint_family_on_def binaryset_def)
+ hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
+ (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
+ f (\<Union>i. binaryset x y i) = (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
using ca
- by (simp add: countably_additive_def) (metis UN_binaryset_eq sums_unique)
- hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
- f (x \<union> y) = suminf (\<lambda>n. f (binaryset x y n))"
+ by (simp add: countably_additive_def)
+ hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
+ f (x \<union> y) = (\<Sum>\<^isub>\<infinity> n. f (binaryset x y n))"
by (simp add: range_binaryset_eq UN_binaryset_eq)
thus "f (x \<union> y) = f x + f y" using posf x y
- by (simp add: Un suminf_binaryset_eq positive_def)
-qed
-
+ by (auto simp add: Un binaryset_psuminf positive_def)
+qed
+
lemma (in algebra) inf_measure_agrees:
- assumes posf: "positive M f" and ca: "countably_additive M f"
- and s: "s \<in> sets M"
+ assumes posf: "positive f" and ca: "countably_additive M f"
+ and s: "s \<in> sets M"
shows "Inf (measure_set M f s) = f s"
-proof (rule Inf_eq)
+ unfolding Inf_pinfreal_def
+proof (safe intro!: Greatest_equality)
fix z
assume z: "z \<in> measure_set M f s"
- from this obtain A where
+ from this obtain A where
A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
- and "s \<subseteq> (\<Union>x. A x)" and sm: "summable (f \<circ> A)"
- and si: "suminf (f \<circ> A) = z"
- by (auto simp add: measure_set_def sums_iff)
+ and "s \<subseteq> (\<Union>x. A x)" and si: "psuminf (f \<circ> A) = z"
+ by (auto simp add: measure_set_def comp_def)
hence seq: "s = (\<Union>i. A i \<inter> s)" by blast
have inc: "increasing M f"
by (metis additive_increasing ca countably_additive_additive posf)
- have sums: "(\<lambda>i. f (A i \<inter> s)) sums f (\<Union>i. A i \<inter> s)"
- proof (rule countably_additiveD [OF ca])
+ have sums: "psuminf (\<lambda>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)"
+ proof (rule countably_additiveD [OF ca])
show "range (\<lambda>n. A n \<inter> s) \<subseteq> sets M" using A s
by blast
show "disjoint_family (\<lambda>n. A n \<inter> s)" using disj
@@ -662,228 +518,184 @@
show "(\<Union>i. A i \<inter> s) \<in> sets M" using A s
by (metis UN_extend_simps(4) s seq)
qed
- hence "f s = suminf (\<lambda>i. f (A i \<inter> s))"
+ hence "f s = psuminf (\<lambda>i. f (A i \<inter> s))"
using seq [symmetric] by (simp add: sums_iff)
- also have "... \<le> suminf (f \<circ> A)"
- proof (rule summable_le [OF _ _ sm])
- show "\<forall>n. f (A n \<inter> s) \<le> (f \<circ> A) n" using A s
- by (force intro: increasingD [OF inc])
- show "summable (\<lambda>i. f (A i \<inter> s))" using sums
- by (simp add: sums_iff)
+ also have "... \<le> psuminf (f \<circ> A)"
+ proof (rule psuminf_le)
+ fix n show "f (A n \<inter> s) \<le> (f \<circ> A) n" using A s
+ by (force intro: increasingD [OF inc])
qed
- also have "... = z" by (rule si)
+ also have "... = z" by (rule si)
finally show "f s \<le> z" .
next
fix y
- assume y: "!!u. u \<in> measure_set M f s \<Longrightarrow> y \<le> u"
+ assume y: "\<forall>u \<in> measure_set M f s. y \<le> u"
thus "y \<le> f s"
- by (blast intro: inf_measure_nonempty [OF posf s subset_refl])
+ by (blast intro: inf_measure_nonempty [of f, OF posf s subset_refl])
qed
lemma (in algebra) inf_measure_empty:
- assumes posf: "positive M f"
+ assumes posf: "positive f"
shows "Inf (measure_set M f {}) = 0"
proof (rule antisym)
- show "0 \<le> Inf (measure_set M f {})"
- by (metis empty_subsetI inf_measure_pos posf)
show "Inf (measure_set M f {}) \<le> 0"
- by (metis Inf_lower empty_sets inf_measure_pos0 inf_measure_nonempty posf
- positive_imp_0 subset_refl)
-qed
+ by (metis complete_lattice_class.Inf_lower empty_sets inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
+qed simp
lemma (in algebra) inf_measure_positive:
- "positive M f \<Longrightarrow>
- positive (| space = space M, sets = Pow (space M) |)
- (\<lambda>x. Inf (measure_set M f x))"
- by (simp add: positive_def inf_measure_empty inf_measure_pos)
+ "positive f \<Longrightarrow>
+ positive (\<lambda>x. Inf (measure_set M f x))"
+ by (simp add: positive_def inf_measure_empty)
lemma (in algebra) inf_measure_increasing:
- assumes posf: "positive M f"
+ assumes posf: "positive f"
shows "increasing (| space = space M, sets = Pow (space M) |)
(\<lambda>x. Inf (measure_set M f x))"
-apply (auto simp add: increasing_def)
-apply (rule Inf_greatest, metis emptyE inf_measure_nonempty top posf)
-apply (rule Inf_lower)
+apply (auto simp add: increasing_def)
+apply (rule complete_lattice_class.Inf_greatest)
+apply (rule complete_lattice_class.Inf_lower)
apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast)
-apply (blast intro: inf_measure_pos0 posf)
done
lemma (in algebra) inf_measure_le:
- assumes posf: "positive M f" and inc: "increasing M f"
- and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M & s \<subseteq> (\<Union>i. A i) & (f \<circ> A) sums r}"
+ assumes posf: "positive f" and inc: "increasing M f"
+ and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M \<and> s \<subseteq> (\<Union>i. A i) \<and> psuminf (f \<circ> A) = r}"
shows "Inf (measure_set M f s) \<le> x"
proof -
from x
- obtain A where A: "range A \<subseteq> sets M" and ss: "s \<subseteq> (\<Union>i. A i)"
- and sm: "summable (f \<circ> A)" and xeq: "suminf (f \<circ> A) = x"
- by (auto simp add: sums_iff)
+ obtain A where A: "range A \<subseteq> sets M" and ss: "s \<subseteq> (\<Union>i. A i)"
+ and xeq: "psuminf (f \<circ> A) = x"
+ by auto
have dA: "range (disjointed A) \<subseteq> sets M"
by (metis A range_disjointed_sets)
- have "\<forall>n. \<bar>(f o disjointed A) n\<bar> \<le> (f \<circ> A) n"
- proof (auto)
- fix n
- have "\<bar>f (disjointed A n)\<bar> = f (disjointed A n)" using posf dA
- by (auto simp add: positive_def image_subset_iff)
- also have "... \<le> f (A n)"
- by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A)
- finally show "\<bar>f (disjointed A n)\<bar> \<le> f (A n)" .
- qed
- from Series.summable_le2 [OF this sm]
- have sda: "summable (f o disjointed A)"
- "suminf (f o disjointed A) \<le> suminf (f \<circ> A)"
- by blast+
- hence ley: "suminf (f o disjointed A) \<le> x"
- by (metis xeq)
- from sda have "(f \<circ> disjointed A) sums suminf (f \<circ> disjointed A)"
- by (simp add: sums_iff)
- hence y: "suminf (f o disjointed A) \<in> measure_set M f s"
+ have "\<forall>n.(f o disjointed A) n \<le> (f \<circ> A) n" unfolding comp_def
+ by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def)
+ hence sda: "psuminf (f o disjointed A) \<le> psuminf (f \<circ> A)"
+ by (blast intro: psuminf_le)
+ hence ley: "psuminf (f o disjointed A) \<le> x"
+ by (metis xeq)
+ hence y: "psuminf (f o disjointed A) \<in> measure_set M f s"
apply (auto simp add: measure_set_def)
- apply (rule_tac x="disjointed A" in exI)
- apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA)
+ apply (rule_tac x="disjointed A" in exI)
+ apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA comp_def)
done
show ?thesis
- by (blast intro: y order_trans [OF _ ley] inf_measure_pos0 posf)
+ by (blast intro: y order_trans [OF _ ley] posf complete_lattice_class.Inf_lower)
qed
lemma (in algebra) inf_measure_close:
- assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (space M)"
- shows "\<exists>A l. range A \<subseteq> sets M & disjoint_family A & s \<subseteq> (\<Union>i. A i) &
- (f \<circ> A) sums l & l \<le> Inf (measure_set M f s) + e"
-proof -
- have " measure_set M f s \<noteq> {}"
- by (metis emptyE ss inf_measure_nonempty [OF posf top])
- hence "\<exists>l \<in> measure_set M f s. l < Inf (measure_set M f s) + e"
- by (rule Inf_close [OF _ e])
- thus ?thesis
- by (auto simp add: measure_set_def, rule_tac x=" A" in exI, auto)
+ assumes posf: "positive f" and e: "0 < e" and ss: "s \<subseteq> (space M)"
+ shows "\<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
+ psuminf (f \<circ> A) \<le> Inf (measure_set M f s) + e"
+proof (cases "Inf (measure_set M f s) = \<omega>")
+ case False
+ obtain l where "l \<in> measure_set M f s" "l \<le> Inf (measure_set M f s) + e"
+ using Inf_close[OF False e] by auto
+ thus ?thesis
+ by (auto intro!: exI[of _ l] simp: measure_set_def comp_def)
+next
+ case True
+ have "measure_set M f s \<noteq> {}"
+ by (metis emptyE ss inf_measure_nonempty [of f, OF posf top])
+ then obtain l where "l \<in> measure_set M f s" by auto
+ moreover from True have "l \<le> Inf (measure_set M f s) + e" by simp
+ ultimately show ?thesis
+ by (auto intro!: exI[of _ l] simp: measure_set_def comp_def)
qed
lemma (in algebra) inf_measure_countably_subadditive:
- assumes posf: "positive M f" and inc: "increasing M f"
+ assumes posf: "positive f" and inc: "increasing M f"
shows "countably_subadditive (| space = space M, sets = Pow (space M) |)
(\<lambda>x. Inf (measure_set M f x))"
-proof (auto simp add: countably_subadditive_def o_def, rule field_le_epsilon)
- fix A :: "nat \<Rightarrow> 'a set" and e :: real
- assume A: "range A \<subseteq> Pow (space M)"
- and disj: "disjoint_family A"
- and sb: "(\<Union>i. A i) \<subseteq> space M"
- and sum1: "summable (\<lambda>n. Inf (measure_set M f (A n)))"
- and e: "0 < e"
- have "!!n. \<exists>B l. range B \<subseteq> sets M \<and> disjoint_family B \<and> A n \<subseteq> (\<Union>i. B i) \<and>
- (f o B) sums l \<and>
- l \<le> Inf (measure_set M f (A n)) + e * (1/2)^(Suc n)"
- apply (rule inf_measure_close [OF posf])
- apply (metis e half mult_pos_pos zero_less_power)
- apply (metis UNIV_I UN_subset_iff sb)
- done
- hence "\<exists>BB ll. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
- A n \<subseteq> (\<Union>i. BB n i) \<and> (f o BB n) sums ll n \<and>
- ll n \<le> Inf (measure_set M f (A n)) + e * (1/2)^(Suc n)"
- by (rule choice2)
- then obtain BB ll
- where BB: "!!n. (range (BB n) \<subseteq> sets M)"
- and disjBB: "!!n. disjoint_family (BB n)"
- and sbBB: "!!n. A n \<subseteq> (\<Union>i. BB n i)"
- and BBsums: "!!n. (f o BB n) sums ll n"
- and ll: "!!n. ll n \<le> Inf (measure_set M f (A n)) + e * (1/2)^(Suc n)"
- by auto blast
- have llpos: "!!n. 0 \<le> ll n"
- by (metis BBsums sums_iff o_apply posf positive_imp_pos suminf_ge_zero
- range_subsetD BB)
- have sll: "summable ll &
- suminf ll \<le> suminf (\<lambda>n. Inf (measure_set M f (A n))) + e"
- proof -
- have "(\<lambda>n. e * (1/2)^(Suc n)) sums (e*1)"
- by (rule sums_mult [OF power_half_series])
- hence sum0: "summable (\<lambda>n. e * (1 / 2) ^ Suc n)"
- and eqe: "(\<Sum>n. e * (1 / 2) ^ n / 2) = e"
- by (auto simp add: sums_iff)
- have 0: "suminf (\<lambda>n. Inf (measure_set M f (A n))) +
- suminf (\<lambda>n. e * (1/2)^(Suc n)) =
- suminf (\<lambda>n. Inf (measure_set M f (A n)) + e * (1/2)^(Suc n))"
- by (rule suminf_add [OF sum1 sum0])
- have 1: "\<forall>n. \<bar>ll n\<bar> \<le> Inf (measure_set M f (A n)) + e * (1/2) ^ Suc n"
- by (metis ll llpos abs_of_nonneg)
- have 2: "summable (\<lambda>n. Inf (measure_set M f (A n)) + e*(1/2)^(Suc n))"
- by (rule summable_add [OF sum1 sum0])
- have "suminf ll \<le> (\<Sum>n. Inf (measure_set M f (A n)) + e*(1/2) ^ Suc n)"
- using Series.summable_le2 [OF 1 2] by auto
- also have "... = (\<Sum>n. Inf (measure_set M f (A n))) +
- (\<Sum>n. e * (1 / 2) ^ Suc n)"
- by (metis 0)
- also have "... = (\<Sum>n. Inf (measure_set M f (A n))) + e"
- by (simp add: eqe)
- finally show ?thesis using Series.summable_le2 [OF 1 2] by auto
- qed
- def C \<equiv> "(split BB) o prod_decode"
- have C: "!!n. C n \<in> sets M"
- apply (rule_tac p="prod_decode n" in PairE)
- apply (simp add: C_def)
- apply (metis BB subsetD rangeI)
- done
- have sbC: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)"
- proof (auto simp add: C_def)
- fix x i
- assume x: "x \<in> A i"
- with sbBB [of i] obtain j where "x \<in> BB i j"
- by blast
- thus "\<exists>i. x \<in> split BB (prod_decode i)"
- by (metis prod_encode_inverse prod.cases)
- qed
- have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> prod_decode"
- by (rule ext) (auto simp add: C_def)
- also have "... sums suminf ll"
- proof (rule suminf_2dimen)
- show "\<And>m n. 0 \<le> (f \<circ> (\<lambda>(x, y). BB x y)) (m, n)" using posf BB
- by (force simp add: positive_def)
- show "\<And>m. (\<lambda>n. (f \<circ> (\<lambda>(x, y). BB x y)) (m, n)) sums ll m"using BBsums BB
- by (force simp add: o_def)
- show "summable ll" using sll
- by auto
- qed
- finally have Csums: "(f \<circ> C) sums suminf ll" .
- have "Inf (measure_set M f (\<Union>i. A i)) \<le> suminf ll"
- apply (rule inf_measure_le [OF posf inc], auto)
- apply (rule_tac x="C" in exI)
- apply (auto simp add: C sbC Csums)
- done
- also have "... \<le> (\<Sum>n. Inf (measure_set M f (A n))) + e" using sll
- by blast
- finally show "Inf (measure_set M f (\<Union>i. A i)) \<le>
- (\<Sum>n. Inf (measure_set M f (A n))) + e" .
+ unfolding countably_subadditive_def o_def
+proof (safe, simp, rule pinfreal_le_epsilon)
+ fix A :: "nat \<Rightarrow> 'a set" and e :: pinfreal
+
+ let "?outer n" = "Inf (measure_set M f (A n))"
+ assume A: "range A \<subseteq> Pow (space M)"
+ and disj: "disjoint_family A"
+ and sb: "(\<Union>i. A i) \<subseteq> space M"
+ and e: "0 < e"
+ hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
+ A n \<subseteq> (\<Union>i. BB n i) \<and>
+ psuminf (f o BB n) \<le> ?outer n + e * (1/2)^(Suc n)"
+ apply (safe intro!: choice inf_measure_close [of f, OF posf _])
+ using e sb by (cases e, auto simp add: not_le mult_pos_pos)
+ then obtain BB
+ where BB: "\<And>n. (range (BB n) \<subseteq> sets M)"
+ and disjBB: "\<And>n. disjoint_family (BB n)"
+ and sbBB: "\<And>n. A n \<subseteq> (\<Union>i. BB n i)"
+ and BBle: "\<And>n. psuminf (f o BB n) \<le> ?outer n + e * (1/2)^(Suc n)"
+ by auto blast
+ have sll: "(\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n)) \<le> psuminf ?outer + e"
+ proof -
+ have "(\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n)) \<le> (\<Sum>\<^isub>\<infinity> n. ?outer n + e*(1/2) ^ Suc n)"
+ by (rule psuminf_le[OF BBle])
+ also have "... = psuminf ?outer + e"
+ using psuminf_half_series by simp
+ finally show ?thesis .
+ qed
+ def C \<equiv> "(split BB) o prod_decode"
+ have C: "!!n. C n \<in> sets M"
+ apply (rule_tac p="prod_decode n" in PairE)
+ apply (simp add: C_def)
+ apply (metis BB subsetD rangeI)
+ done
+ have sbC: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)"
+ proof (auto simp add: C_def)
+ fix x i
+ assume x: "x \<in> A i"
+ with sbBB [of i] obtain j where "x \<in> BB i j"
+ by blast
+ thus "\<exists>i. x \<in> split BB (prod_decode i)"
+ by (metis prod_encode_inverse prod.cases)
+ qed
+ have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> prod_decode"
+ by (rule ext) (auto simp add: C_def)
+ moreover have "psuminf ... = (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))" using BBle
+ by (force intro!: psuminf_2dimen simp: o_def)
+ ultimately have Csums: "psuminf (f \<circ> C) = (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))" by simp
+ have "Inf (measure_set M f (\<Union>i. A i)) \<le> (\<Sum>\<^isub>\<infinity> n. psuminf (f o BB n))"
+ apply (rule inf_measure_le [OF posf(1) inc], auto)
+ apply (rule_tac x="C" in exI)
+ apply (auto simp add: C sbC Csums)
+ done
+ also have "... \<le> (\<Sum>\<^isub>\<infinity>n. Inf (measure_set M f (A n))) + e" using sll
+ by blast
+ finally show "Inf (measure_set M f (\<Union>i. A i)) \<le> psuminf ?outer + e" .
qed
lemma (in algebra) inf_measure_outer:
- "positive M f \<Longrightarrow> increasing M f
+ "\<lbrakk> positive f ; increasing M f \<rbrakk>
\<Longrightarrow> outer_measure_space (| space = space M, sets = Pow (space M) |)
(\<lambda>x. Inf (measure_set M f x))"
- by (simp add: outer_measure_space_def inf_measure_positive
- inf_measure_increasing inf_measure_countably_subadditive)
+ by (simp add: outer_measure_space_def inf_measure_empty
+ inf_measure_increasing inf_measure_countably_subadditive positive_def)
(*MOVE UP*)
lemma (in algebra) algebra_subset_lambda_system:
- assumes posf: "positive M f" and inc: "increasing M f"
+ assumes posf: "positive f" and inc: "increasing M f"
and add: "additive M f"
shows "sets M \<subseteq> lambda_system (| space = space M, sets = Pow (space M) |)
(\<lambda>x. Inf (measure_set M f x))"
-proof (auto dest: sets_into_space
- simp add: algebra.lambda_system_eq [OF algebra_Pow])
+proof (auto dest: sets_into_space
+ simp add: algebra.lambda_system_eq [OF algebra_Pow])
fix x s
assume x: "x \<in> sets M"
and s: "s \<subseteq> space M"
- have [simp]: "!!x. x \<in> sets M \<Longrightarrow> s \<inter> (space M - x) = s-x" using s
+ have [simp]: "!!x. x \<in> sets M \<Longrightarrow> s \<inter> (space M - x) = s-x" using s
by blast
have "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
\<le> Inf (measure_set M f s)"
- proof (rule field_le_epsilon)
- fix e :: real
+ proof (rule pinfreal_le_epsilon)
+ fix e :: pinfreal
assume e: "0 < e"
- from inf_measure_close [OF posf e s]
- obtain A l where A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
- and sUN: "s \<subseteq> (\<Union>i. A i)" and fsums: "(f \<circ> A) sums l"
- and l: "l \<le> Inf (measure_set M f s) + e"
+ from inf_measure_close [of f, OF posf e s]
+ obtain A where A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
+ and sUN: "s \<subseteq> (\<Union>i. A i)"
+ and l: "psuminf (f \<circ> A) \<le> Inf (measure_set M f s) + e"
by auto
have [simp]: "!!x. x \<in> sets M \<Longrightarrow>
(f o (\<lambda>z. z \<inter> (space M - x)) o A) = (f o (\<lambda>z. z - x) o A)"
@@ -891,104 +703,87 @@
have [simp]: "!!n. f (A n \<inter> x) + f (A n - x) = f (A n)"
by (subst additiveD [OF add, symmetric])
(auto simp add: x range_subsetD [OF A] Int_Diff_Un Int_Diff_disjoint)
- have fsumb: "summable (f \<circ> A)"
- by (metis fsums sums_iff)
{ fix u
assume u: "u \<in> sets M"
- have [simp]: "\<And>n. \<bar>f (A n \<inter> u)\<bar> \<le> f (A n)"
- by (simp add: positive_imp_pos [OF posf] increasingD [OF inc]
- u Int range_subsetD [OF A])
- have 1: "summable (f o (\<lambda>z. z\<inter>u) o A)"
- by (rule summable_comparison_test [OF _ fsumb]) simp
- have 2: "Inf (measure_set M f (s\<inter>u)) \<le> suminf (f o (\<lambda>z. z\<inter>u) o A)"
- proof (rule Inf_lower)
- show "suminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A) \<in> measure_set M f (s \<inter> u)"
- apply (simp add: measure_set_def)
- apply (rule_tac x="(\<lambda>z. z \<inter> u) o A" in exI)
- apply (auto simp add: disjoint_family_subset [OF disj])
- apply (blast intro: u range_subsetD [OF A])
+ have [simp]: "\<And>n. f (A n \<inter> u) \<le> f (A n)"
+ by (simp add: increasingD [OF inc] u Int range_subsetD [OF A])
+ have 2: "Inf (measure_set M f (s \<inter> u)) \<le> psuminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A)"
+ proof (rule complete_lattice_class.Inf_lower)
+ show "psuminf (f \<circ> (\<lambda>z. z \<inter> u) \<circ> A) \<in> measure_set M f (s \<inter> u)"
+ apply (simp add: measure_set_def)
+ apply (rule_tac x="(\<lambda>z. z \<inter> u) o A" in exI)
+ apply (auto simp add: disjoint_family_subset [OF disj] o_def)
+ apply (blast intro: u range_subsetD [OF A])
apply (blast dest: subsetD [OF sUN])
- apply (metis 1 o_assoc sums_iff)
done
- next
- show "\<And>x. x \<in> measure_set M f (s \<inter> u) \<Longrightarrow> 0 \<le> x"
- by (blast intro: inf_measure_pos0 [OF posf])
- qed
- note 1 2
+ qed
} note lesum = this
- have sum1: "summable (f o (\<lambda>z. z\<inter>x) o A)"
- and inf1: "Inf (measure_set M f (s\<inter>x)) \<le> suminf (f o (\<lambda>z. z\<inter>x) o A)"
- and sum2: "summable (f o (\<lambda>z. z \<inter> (space M - x)) o A)"
- and inf2: "Inf (measure_set M f (s \<inter> (space M - x)))
- \<le> suminf (f o (\<lambda>z. z \<inter> (space M - x)) o A)"
+ have inf1: "Inf (measure_set M f (s\<inter>x)) \<le> psuminf (f o (\<lambda>z. z\<inter>x) o A)"
+ and inf2: "Inf (measure_set M f (s \<inter> (space M - x)))
+ \<le> psuminf (f o (\<lambda>z. z \<inter> (space M - x)) o A)"
by (metis Diff lesum top x)+
hence "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
- \<le> suminf (f o (\<lambda>s. s\<inter>x) o A) + suminf (f o (\<lambda>s. s-x) o A)"
- by (simp add: x)
- also have "... \<le> suminf (f o A)" using suminf_add [OF sum1 sum2]
- by (simp add: x) (simp add: o_def)
+ \<le> psuminf (f o (\<lambda>s. s\<inter>x) o A) + psuminf (f o (\<lambda>s. s-x) o A)"
+ by (simp add: x add_mono)
+ also have "... \<le> psuminf (f o A)"
+ by (simp add: x psuminf_add[symmetric] o_def)
also have "... \<le> Inf (measure_set M f s) + e"
- by (metis fsums l sums_unique)
+ by (rule l)
finally show "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
\<le> Inf (measure_set M f s) + e" .
qed
- moreover
+ moreover
have "Inf (measure_set M f s)
\<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
proof -
have "Inf (measure_set M f s) = Inf (measure_set M f ((s\<inter>x) \<union> (s-x)))"
by (metis Un_Diff_Int Un_commute)
- also have "... \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
- apply (rule subadditiveD)
- apply (iprover intro: algebra.countably_subadditive_subadditive algebra_Pow
+ also have "... \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))"
+ apply (rule subadditiveD)
+ apply (iprover intro: algebra.countably_subadditive_subadditive algebra_Pow
inf_measure_positive inf_measure_countably_subadditive posf inc)
- apply (auto simp add: subsetD [OF s])
+ apply (auto simp add: subsetD [OF s])
done
finally show ?thesis .
qed
- ultimately
+ ultimately
show "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
= Inf (measure_set M f s)"
by (rule order_antisym)
qed
lemma measure_down:
- "measure_space N \<Longrightarrow> sigma_algebra M \<Longrightarrow> sets M \<subseteq> sets N \<Longrightarrow>
- (measure M = measure N) \<Longrightarrow> measure_space M"
- by (simp add: measure_space_def measure_space_axioms_def positive_def
- countably_additive_def)
+ "measure_space N \<mu> \<Longrightarrow> sigma_algebra M \<Longrightarrow> sets M \<subseteq> sets N \<Longrightarrow>
+ (\<nu> = \<mu>) \<Longrightarrow> measure_space M \<nu>"
+ by (simp add: measure_space_def measure_space_axioms_def positive_def
+ countably_additive_def)
blast
theorem (in algebra) caratheodory:
- assumes posf: "positive M f" and ca: "countably_additive M f"
- shows "\<exists>MS :: 'a measure_space.
- (\<forall>s \<in> sets M. measure MS s = f s) \<and>
- ((|space = space MS, sets = sets MS|) = sigma (space M) (sets M)) \<and>
- measure_space MS"
+ assumes posf: "positive f" and ca: "countably_additive M f"
+ shows "\<exists>\<mu> :: 'a set \<Rightarrow> pinfreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and> measure_space (sigma (space M) (sets M)) \<mu>"
proof -
have inc: "increasing M f"
- by (metis additive_increasing ca countably_additive_additive posf)
+ by (metis additive_increasing ca countably_additive_additive posf)
let ?infm = "(\<lambda>x. Inf (measure_set M f x))"
def ls \<equiv> "lambda_system (|space = space M, sets = Pow (space M)|) ?infm"
- have mls: "measure_space (|space = space M, sets = ls, measure = ?infm|)"
+ have mls: "measure_space \<lparr>space = space M, sets = ls\<rparr> ?infm"
using sigma_algebra.caratheodory_lemma
[OF sigma_algebra_Pow inf_measure_outer [OF posf inc]]
by (simp add: ls_def)
- hence sls: "sigma_algebra (|space = space M, sets = ls, measure = ?infm|)"
- by (simp add: measure_space_def)
- have "sets M \<subseteq> ls"
+ hence sls: "sigma_algebra (|space = space M, sets = ls|)"
+ by (simp add: measure_space_def)
+ have "sets M \<subseteq> ls"
by (simp add: ls_def)
(metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
- hence sgs_sb: "sigma_sets (space M) (sets M) \<subseteq> ls"
+ hence sgs_sb: "sigma_sets (space M) (sets M) \<subseteq> ls"
using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"]
by simp
- have "measure_space (|space = space M,
- sets = sigma_sets (space M) (sets M),
- measure = ?infm|)"
- by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
+ have "measure_space (sigma (space M) (sets M)) ?infm"
+ unfolding sigma_def
+ by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
(simp_all add: sgs_sb space_closed)
- thus ?thesis
- by (force simp add: sigma_def inf_measure_agrees [OF posf ca])
-qed
+ thus ?thesis using inf_measure_agrees [OF posf ca] by (auto intro!: exI[of _ ?infm])
+ qed
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Euclidean_Lebesgue.thy Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,213 @@
+theory Euclidean_Lebesgue
+ imports Lebesgue_Integration Lebesgue_Measure
+begin
+
+lemma simple_function_has_integral:
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+ assumes f:"lebesgue.simple_function f"
+ and f':"\<forall>x. f x \<noteq> \<omega>"
+ and om:"\<forall>x\<in>range f. lmeasure (f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
+ shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
+ unfolding lebesgue.simple_integral_def
+ apply(subst lebesgue_simple_function_indicator[OF f])
+proof- case goal1
+ have *:"\<And>x. \<forall>y\<in>range f. y * indicator (f -` {y}) x \<noteq> \<omega>"
+ "\<forall>x\<in>range f. x * lmeasure (f -` {x} \<inter> UNIV) \<noteq> \<omega>"
+ using f' om unfolding indicator_def by auto
+ show ?case unfolding space_lebesgue_space real_of_pinfreal_setsum'[OF *(1),THEN sym]
+ unfolding real_of_pinfreal_setsum'[OF *(2),THEN sym]
+ unfolding real_of_pinfreal_setsum space_lebesgue_space
+ apply(rule has_integral_setsum)
+ proof safe show "finite (range f)" using f by (auto dest: lebesgue.simple_functionD)
+ fix y::'a show "((\<lambda>x. real (f y * indicator (f -` {f y}) x)) has_integral
+ real (f y * lmeasure (f -` {f y} \<inter> UNIV))) UNIV"
+ proof(cases "f y = 0") case False
+ have mea:"gmeasurable (f -` {f y})" apply(rule glmeasurable_finite)
+ using assms unfolding lebesgue.simple_function_def using False by auto
+ have *:"\<And>x. real (indicator (f -` {f y}) x::pinfreal) = (if x \<in> f -` {f y} then 1 else 0)" by auto
+ show ?thesis unfolding real_of_pinfreal_mult[THEN sym]
+ apply(rule has_integral_cmul[where 'b=real, unfolded real_scaleR_def])
+ unfolding Int_UNIV_right lmeasure_gmeasure[OF mea,THEN sym]
+ unfolding measure_integral_univ[OF mea] * apply(rule integrable_integral)
+ unfolding gmeasurable_integrable[THEN sym] using mea .
+ qed auto
+ qed qed
+
+lemma (in measure_space) positive_integral_omega:
+ assumes "f \<in> borel_measurable M"
+ and "positive_integral f \<noteq> \<omega>"
+ shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
+proof -
+ have "\<omega> * \<mu> (f -` {\<omega>} \<inter> space M) = positive_integral (\<lambda>x. \<omega> * indicator (f -` {\<omega>} \<inter> space M) x)"
+ using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \<omega> \<omega>] by simp
+ also have "\<dots> \<le> positive_integral f"
+ by (auto intro!: positive_integral_mono simp: indicator_def)
+ finally show ?thesis
+ using assms(2) by (cases ?thesis) auto
+qed
+
+lemma (in measure_space) simple_integral_omega:
+ assumes "simple_function f"
+ and "simple_integral f \<noteq> \<omega>"
+ shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
+proof (rule positive_integral_omega)
+ show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
+ show "positive_integral f \<noteq> \<omega>"
+ using assms by (simp add: positive_integral_eq_simple_integral)
+qed
+
+lemma bounded_realI: assumes "\<forall>x\<in>s. abs (x::real) \<le> B" shows "bounded s"
+ unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI)
+ using assms by auto
+
+lemma simple_function_has_integral':
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+ assumes f:"lebesgue.simple_function f"
+ and i: "lebesgue.simple_integral f \<noteq> \<omega>"
+ shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
+proof- let ?f = "\<lambda>x. if f x = \<omega> then 0 else f x"
+ { fix x have "real (f x) = real (?f x)" by (cases "f x") auto } note * = this
+ have **:"{x. f x \<noteq> ?f x} = f -` {\<omega>}" by auto
+ have **:"lmeasure {x\<in>space lebesgue_space. f x \<noteq> ?f x} = 0"
+ using lebesgue.simple_integral_omega[OF assms] by(auto simp add:**)
+ show ?thesis apply(subst lebesgue.simple_integral_cong'[OF f _ **])
+ apply(rule lebesgue.simple_function_compose1[OF f])
+ unfolding * defer apply(rule simple_function_has_integral)
+ proof-
+ show "lebesgue.simple_function ?f"
+ using lebesgue.simple_function_compose1[OF f] .
+ show "\<forall>x. ?f x \<noteq> \<omega>" by auto
+ show "\<forall>x\<in>range ?f. lmeasure (?f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
+ proof (safe, simp, safe, rule ccontr)
+ fix y assume "f y \<noteq> \<omega>" "f y \<noteq> 0"
+ hence "(\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y} = f -` {f y}"
+ by (auto split: split_if_asm)
+ moreover assume "lmeasure ((\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y}) = \<omega>"
+ ultimately have "lmeasure (f -` {f y}) = \<omega>" by simp
+ moreover
+ have "f y * lmeasure (f -` {f y}) \<noteq> \<omega>" using i f
+ unfolding lebesgue.simple_integral_def setsum_\<omega> lebesgue.simple_function_def
+ by auto
+ ultimately have "f y = 0" by (auto split: split_if_asm)
+ then show False using `f y \<noteq> 0` by simp
+ qed
+ qed
+qed
+
+lemma (in measure_space) positive_integral_monotone_convergence:
+ fixes f::"nat \<Rightarrow> 'a \<Rightarrow> pinfreal"
+ assumes i: "\<And>i. f i \<in> borel_measurable M" and mono: "\<And>x. mono (\<lambda>n. f n x)"
+ and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
+ shows "u \<in> borel_measurable M"
+ and "(\<lambda>i. positive_integral (f i)) ----> positive_integral u" (is ?ilim)
+proof -
+ from positive_integral_isoton[unfolded isoton_fun_expand isoton_iff_Lim_mono, of f u]
+ show ?ilim using mono lim i by auto
+ have "(SUP i. f i) = u" using mono lim SUP_Lim_pinfreal
+ unfolding expand_fun_eq SUPR_fun_expand mono_def by auto
+ moreover have "(SUP i. f i) \<in> borel_measurable M"
+ using i by (rule borel_measurable_SUP)
+ ultimately show "u \<in> borel_measurable M" by simp
+qed
+
+lemma positive_integral_has_integral:
+ fixes f::"'a::ordered_euclidean_space => pinfreal"
+ assumes f:"f \<in> borel_measurable lebesgue_space"
+ and int_om:"lebesgue.positive_integral f \<noteq> \<omega>"
+ and f_om:"\<forall>x. f x \<noteq> \<omega>" (* TODO: remove this *)
+ shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.positive_integral f))) UNIV"
+proof- let ?i = "lebesgue.positive_integral f"
+ from lebesgue.borel_measurable_implies_simple_function_sequence[OF f]
+ guess u .. note conjunctD2[OF this,rule_format] note u = conjunctD2[OF this(1)] this(2)
+ let ?u = "\<lambda>i x. real (u i x)" and ?f = "\<lambda>x. real (f x)"
+ have u_simple:"\<And>k. lebesgue.simple_integral (u k) = lebesgue.positive_integral (u k)"
+ apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) ..
+ (*unfolding image_iff defer apply(rule) using u(2) by smt*)
+ have int_u_le:"\<And>k. lebesgue.simple_integral (u k) \<le> lebesgue.positive_integral f"
+ unfolding u_simple apply(rule lebesgue.positive_integral_mono)
+ using isoton_Sup[OF u(3)] unfolding le_fun_def by auto
+ have u_int_om:"\<And>i. lebesgue.simple_integral (u i) \<noteq> \<omega>"
+ proof- case goal1 thus ?case using int_u_le[of i] int_om by auto qed
+
+ note u_int = simple_function_has_integral'[OF u(1) this]
+ have "(\<lambda>x. real (f x)) integrable_on UNIV \<and>
+ (\<lambda>k. gintegral UNIV (\<lambda>x. real (u k x))) ----> gintegral UNIV (\<lambda>x. real (f x))"
+ apply(rule monotone_convergence_increasing) apply(rule,rule,rule u_int)
+ proof safe case goal1 show ?case apply(rule real_of_pinfreal_mono) using u(2,3) by auto
+ next case goal2 show ?case using u(3) apply(subst lim_Real[THEN sym])
+ prefer 3 apply(subst Real_real') defer apply(subst Real_real')
+ using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] using f_om u by auto
+ next case goal3
+ show ?case apply(rule bounded_realI[where B="real (lebesgue.positive_integral f)"])
+ apply safe apply(subst abs_of_nonneg) apply(rule integral_nonneg,rule) apply(rule u_int)
+ unfolding integral_unique[OF u_int] defer apply(rule real_of_pinfreal_mono[OF _ int_u_le])
+ using u int_om by auto
+ qed note int = conjunctD2[OF this]
+
+ have "(\<lambda>i. lebesgue.simple_integral (u i)) ----> ?i" unfolding u_simple
+ apply(rule lebesgue.positive_integral_monotone_convergence(2))
+ apply(rule lebesgue.borel_measurable_simple_function[OF u(1)])
+ using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] by auto
+ hence "(\<lambda>i. real (lebesgue.simple_integral (u i))) ----> real ?i" apply-
+ apply(subst lim_Real[THEN sym]) prefer 3
+ apply(subst Real_real') defer apply(subst Real_real')
+ using u f_om int_om u_int_om by auto
+ note * = LIMSEQ_unique[OF this int(2)[unfolded integral_unique[OF u_int]]]
+ show ?thesis unfolding * by(rule integrable_integral[OF int(1)])
+qed
+
+lemma lebesgue_integral_has_integral:
+ fixes f::"'a::ordered_euclidean_space => real"
+ assumes f:"lebesgue.integrable f"
+ shows "(f has_integral (lebesgue.integral f)) UNIV"
+proof- let ?n = "\<lambda>x. - min (f x) 0" and ?p = "\<lambda>x. max (f x) 0"
+ have *:"f = (\<lambda>x. ?p x - ?n x)" apply rule by auto
+ note f = lebesgue.integrableD[OF f]
+ show ?thesis unfolding lebesgue.integral_def apply(subst *)
+ proof(rule has_integral_sub) case goal1
+ have *:"\<forall>x. Real (f x) \<noteq> \<omega>" by auto
+ note lebesgue.borel_measurable_Real[OF f(1)]
+ from positive_integral_has_integral[OF this f(2) *]
+ show ?case unfolding real_Real_max .
+ next case goal2
+ have *:"\<forall>x. Real (- f x) \<noteq> \<omega>" by auto
+ note lebesgue.borel_measurable_uminus[OF f(1)]
+ note lebesgue.borel_measurable_Real[OF this]
+ from positive_integral_has_integral[OF this f(3) *]
+ show ?case unfolding real_Real_max minus_min_eq_max by auto
+ qed
+qed
+
+lemma lmeasurable_imp_borel[dest]: fixes s::"'a::ordered_euclidean_space set"
+ assumes "s \<in> sets borel_space" shows "lmeasurable s"
+proof- let ?S = "range (\<lambda>(a, b). {a .. (b :: 'a\<Colon>ordered_euclidean_space)})"
+ have *:"?S \<subseteq> sets lebesgue_space" by auto
+ have "s \<in> sigma_sets UNIV ?S" using assms
+ unfolding borel_space_eq_atLeastAtMost by (simp add: sigma_def)
+ thus ?thesis using lebesgue.sigma_subset[of ?S,unfolded sets_sigma,OF *]
+ by auto
+qed
+
+lemma lmeasurable_open[dest]:
+ assumes "open s" shows "lmeasurable s"
+proof- have "s \<in> sets borel_space" using assms by auto
+ thus ?thesis by auto qed
+
+lemma continuous_on_imp_borel_measurable:
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
+ assumes "continuous_on UNIV f"
+ shows "f \<in> borel_measurable lebesgue_space"
+ apply(rule lebesgue.borel_measurableI)
+ unfolding lebesgue_measurable apply(rule lmeasurable_open)
+ using continuous_open_preimage[OF assms] unfolding vimage_def by auto
+
+
+lemma (in measure_space) integral_monotone_convergence_pos':
+ assumes i: "\<And>i. integrable (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
+ and pos: "\<And>x i. 0 \<le> f i x"
+ and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
+ and ilim: "(\<lambda>i. integral (f i)) ----> x"
+ shows "integrable u \<and> integral u = x"
+ using integral_monotone_convergence_pos[OF assms] by auto
+
+end
--- a/src/HOL/Probability/Information.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Probability/Information.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1,7 +1,12 @@
theory Information
-imports Probability_Space Product_Measure Convex
+imports Probability_Space Product_Measure Convex Radon_Nikodym
begin
+lemma real_of_pinfreal_inverse[simp]:
+ fixes X :: pinfreal
+ shows "real (inverse X) = 1 / real X"
+ by (cases X) (auto simp: inverse_eq_divide)
+
section "Convex theory"
lemma log_setsum:
@@ -41,7 +46,7 @@
assume *: "s - {i. a i = 0} = {}"
hence "setsum a (s - {i. a i = 0}) = 0" by (simp add: * setsum_empty)
with sum_1 show False by simp
-qed
+ qed
fix i assume "i \<in> s - {i. a i = 0}"
hence "i \<in> s" "a i \<noteq> 0" by simp_all
@@ -50,133 +55,6 @@
ultimately show ?thesis by simp
qed
-section "Information theory"
-
-lemma (in finite_prob_space) sum_over_space_distrib:
- "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
- unfolding distribution_def prob_space[symmetric] using finite_space
- by (subst measure_finitely_additive'')
- (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob])
-
-locale finite_information_space = finite_prob_space +
- fixes b :: real assumes b_gt_1: "1 < b"
-
-definition
- "KL_divergence b M X Y =
- measure_space.integral (M\<lparr>measure := X\<rparr>)
- (\<lambda>x. log b ((measure_space.RN_deriv (M \<lparr>measure := Y\<rparr> ) X) x))"
-
-lemma (in finite_prob_space) distribution_mono:
- assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
- shows "distribution X x \<le> distribution Y y"
- unfolding distribution_def
- using assms by (auto simp: sets_eq_Pow intro!: measure_mono)
-
-lemma (in prob_space) distribution_remove_const:
- shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
- and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
- and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
- and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
- and "distribution (\<lambda>x. ()) {()} = 1"
- unfolding prob_space[symmetric]
- by (auto intro!: arg_cong[where f=prob] simp: distribution_def)
-
-
-context finite_information_space
-begin
-
-lemma distribution_mono_gt_0:
- assumes gt_0: "0 < distribution X x"
- assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
- shows "0 < distribution Y y"
- by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
-
-lemma
- assumes "0 \<le> A" and pos: "0 < A \<Longrightarrow> 0 < B" "0 < A \<Longrightarrow> 0 < C"
- shows mult_log_mult: "A * log b (B * C) = A * log b B + A * log b C" (is "?mult")
- and mult_log_divide: "A * log b (B / C) = A * log b B - A * log b C" (is "?div")
-proof -
- have "?mult \<and> ?div"
-proof (cases "A = 0")
- case False
- hence "0 < A" using `0 \<le> A` by auto
- with pos[OF this] show "?mult \<and> ?div" using b_gt_1
- by (auto simp: log_divide log_mult field_simps)
-qed simp
- thus ?mult and ?div by auto
-qed
-
-lemma split_pairs:
- shows
- "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and
- "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto
-
-ML {*
-
- (* tactic to solve equations of the form @{term "W * log b (X / (Y * Z)) = W * log b X - W * log b (Y * Z)"}
- where @{term W} is a joint distribution of @{term X}, @{term Y}, and @{term Z}. *)
-
- val mult_log_intros = [@{thm mult_log_divide}, @{thm mult_log_mult}]
- val intros = [@{thm divide_pos_pos}, @{thm mult_pos_pos}, @{thm positive_distribution}]
-
- val distribution_gt_0_tac = (rtac @{thm distribution_mono_gt_0}
- THEN' assume_tac
- THEN' clarsimp_tac (clasimpset_of @{context} addsimps2 @{thms split_pairs}))
-
- val distr_mult_log_eq_tac = REPEAT_ALL_NEW (CHANGED o TRY o
- (resolve_tac (mult_log_intros @ intros)
- ORELSE' distribution_gt_0_tac
- ORELSE' clarsimp_tac (clasimpset_of @{context})))
-
- fun instanciate_term thy redex intro =
- let
- val intro_concl = Thm.concl_of intro
-
- val lhs = intro_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst
-
- val m = SOME (Pattern.match thy (lhs, redex) (Vartab.empty, Vartab.empty))
- handle Pattern.MATCH => NONE
-
- in
- Option.map (fn m => Envir.subst_term m intro_concl) m
- end
-
- fun mult_log_simproc simpset redex =
- let
- val ctxt = Simplifier.the_context simpset
- val thy = ProofContext.theory_of ctxt
- fun prove (SOME thm) = (SOME
- (Goal.prove ctxt [] [] thm (K (distr_mult_log_eq_tac 1))
- |> mk_meta_eq)
- handle THM _ => NONE)
- | prove NONE = NONE
- in
- get_first (instanciate_term thy (term_of redex) #> prove) mult_log_intros
- end
-*}
-
-simproc_setup mult_log ("distribution X x * log b (A * B)" |
- "distribution X x * log b (A / B)") = {* K mult_log_simproc *}
-
-end
-
-lemma KL_divergence_eq_finite:
- assumes u: "finite_measure_space (M\<lparr>measure := u\<rparr>)"
- assumes v: "finite_measure_space (M\<lparr>measure := v\<rparr>)"
- assumes u_0: "\<And>x. \<lbrakk> x \<in> space M ; v {x} = 0 \<rbrakk> \<Longrightarrow> u {x} = 0"
- shows "KL_divergence b M u v = (\<Sum>x\<in>space M. u {x} * log b (u {x} / v {x}))" (is "_ = ?sum")
-proof (simp add: KL_divergence_def, subst finite_measure_space.integral_finite_singleton, simp_all add: u)
- have ms_u: "measure_space (M\<lparr>measure := u\<rparr>)"
- using u unfolding finite_measure_space_def by simp
-
- show "(\<Sum>x \<in> space M. log b (measure_space.RN_deriv (M\<lparr>measure := v\<rparr>) u x) * u {x}) = ?sum"
- apply (rule setsum_cong[OF refl])
- apply simp
- apply (safe intro!: arg_cong[where f="log b"] )
- apply (subst finite_measure_space.RN_deriv_finite_singleton)
- using assms ms_u by auto
-qed
-
lemma log_setsum_divide:
assumes "finite S" and "S \<noteq> {}" and "1 < b"
assumes "(\<Sum>x\<in>S. g x) = 1"
@@ -227,47 +105,235 @@
finally show ?thesis .
qed
-lemma KL_divergence_positive_finite:
- assumes u: "finite_prob_space (M\<lparr>measure := u\<rparr>)"
- assumes v: "finite_prob_space (M\<lparr>measure := v\<rparr>)"
- assumes u_0: "\<And>x. \<lbrakk> x \<in> space M ; v {x} = 0 \<rbrakk> \<Longrightarrow> u {x} = 0"
- and "1 < b"
- shows "0 \<le> KL_divergence b M u v"
+lemma (in finite_prob_space) sum_over_space_distrib:
+ "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
+ unfolding distribution_def measure_space_1[symmetric] using finite_space
+ by (subst measure_finitely_additive'')
+ (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\<mu>])
+
+lemma (in finite_prob_space) sum_over_space_real_distribution:
+ "(\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
+ unfolding distribution_def prob_space[symmetric] using finite_space
+ by (subst real_finite_measure_finite_Union[symmetric])
+ (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob])
+
+section "Information theory"
+
+definition
+ "KL_divergence b M \<mu> \<nu> =
+ measure_space.integral M \<mu> (\<lambda>x. log b (real (sigma_finite_measure.RN_deriv M \<nu> \<mu> x)))"
+
+locale finite_information_space = finite_prob_space +
+ fixes b :: real assumes b_gt_1: "1 < b"
+
+lemma (in finite_prob_space) distribution_mono:
+ assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
+ shows "distribution X x \<le> distribution Y y"
+ unfolding distribution_def
+ using assms by (auto simp: sets_eq_Pow intro!: measure_mono)
+
+lemma (in prob_space) distribution_remove_const:
+ shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
+ and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
+ and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
+ and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
+ and "distribution (\<lambda>x. ()) {()} = 1"
+ unfolding measure_space_1[symmetric]
+ by (auto intro!: arg_cong[where f="\<mu>"] simp: distribution_def)
+
+context finite_information_space
+begin
+
+lemma distribution_mono_gt_0:
+ assumes gt_0: "0 < distribution X x"
+ assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
+ shows "0 < distribution Y y"
+ by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
+
+lemma
+ assumes "0 \<le> A" and pos: "0 < A \<Longrightarrow> 0 < B" "0 < A \<Longrightarrow> 0 < C"
+ shows mult_log_mult: "A * log b (B * C) = A * log b B + A * log b C" (is "?mult")
+ and mult_log_divide: "A * log b (B / C) = A * log b B - A * log b C" (is "?div")
proof -
- interpret u: finite_prob_space "M\<lparr>measure := u\<rparr>" using u .
- interpret v: finite_prob_space "M\<lparr>measure := v\<rparr>" using v .
+ have "?mult \<and> ?div"
+ proof (cases "A = 0")
+ case False
+ hence "0 < A" using `0 \<le> A` by auto
+ with pos[OF this] show "?mult \<and> ?div" using b_gt_1
+ by (auto simp: log_divide log_mult field_simps)
+ qed simp
+ thus ?mult and ?div by auto
+qed
+
+lemma split_pairs:
+ shows
+ "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and
+ "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto
+
+lemma (in finite_information_space) distribution_finite:
+ "distribution X A \<noteq> \<omega>"
+ using measure_finite[of "X -` A \<inter> space M"]
+ unfolding distribution_def sets_eq_Pow by auto
+
+lemma (in finite_information_space) real_distribution_gt_0[simp]:
+ "0 < real (distribution Y y) \<longleftrightarrow> 0 < distribution Y y"
+ using assms by (auto intro!: real_pinfreal_pos distribution_finite)
- have *: "space M \<noteq> {}" using u.not_empty by simp
+lemma real_distribution_mult_pos_pos:
+ assumes "0 < distribution Y y"
+ and "0 < distribution X x"
+ shows "0 < real (distribution Y y * distribution X x)"
+ unfolding real_of_pinfreal_mult[symmetric]
+ using assms by (auto intro!: mult_pos_pos)
+
+lemma real_distribution_divide_pos_pos:
+ assumes "0 < distribution Y y"
+ and "0 < distribution X x"
+ shows "0 < real (distribution Y y / distribution X x)"
+ unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric]
+ using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
+
+lemma real_distribution_mult_inverse_pos_pos:
+ assumes "0 < distribution Y y"
+ and "0 < distribution X x"
+ shows "0 < real (distribution Y y * inverse (distribution X x))"
+ unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric]
+ using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
+
+ML {*
+
+ (* tactic to solve equations of the form @{term "W * log b (X / (Y * Z)) = W * log b X - W * log b (Y * Z)"}
+ where @{term W} is a joint distribution of @{term X}, @{term Y}, and @{term Z}. *)
+
+ val mult_log_intros = [@{thm mult_log_divide}, @{thm mult_log_mult}]
+ val intros = [@{thm divide_pos_pos}, @{thm mult_pos_pos}, @{thm real_pinfreal_nonneg},
+ @{thm real_distribution_divide_pos_pos},
+ @{thm real_distribution_mult_inverse_pos_pos},
+ @{thm real_distribution_mult_pos_pos}]
+
+ val distribution_gt_0_tac = (rtac @{thm distribution_mono_gt_0}
+ THEN' assume_tac
+ THEN' clarsimp_tac (clasimpset_of @{context} addsimps2 @{thms split_pairs}))
- have "- (KL_divergence b M u v) \<le> log b (\<Sum>x\<in>space M. v {x})"
- proof (subst KL_divergence_eq_finite, safe intro!: log_setsum_divide *)
- show "finite_measure_space (M\<lparr>measure := u\<rparr>)"
- "finite_measure_space (M\<lparr>measure := v\<rparr>)"
- using u v unfolding finite_prob_space_eq by simp_all
+ val distr_mult_log_eq_tac = REPEAT_ALL_NEW (CHANGED o TRY o
+ (resolve_tac (mult_log_intros @ intros)
+ ORELSE' distribution_gt_0_tac
+ ORELSE' clarsimp_tac (clasimpset_of @{context})))
+
+ fun instanciate_term thy redex intro =
+ let
+ val intro_concl = Thm.concl_of intro
+
+ val lhs = intro_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst
+
+ val m = SOME (Pattern.match thy (lhs, redex) (Vartab.empty, Vartab.empty))
+ handle Pattern.MATCH => NONE
+
+ in
+ Option.map (fn m => Envir.subst_term m intro_concl) m
+ end
- show "finite (space M)" using u.finite_space by simp
- show "1 < b" by fact
- show "(\<Sum>x\<in>space M. u {x}) = 1" using u.sum_over_space_eq_1 by simp
+ fun mult_log_simproc simpset redex =
+ let
+ val ctxt = Simplifier.the_context simpset
+ val thy = ProofContext.theory_of ctxt
+ fun prove (SOME thm) = (SOME
+ (Goal.prove ctxt [] [] thm (K (distr_mult_log_eq_tac 1))
+ |> mk_meta_eq)
+ handle THM _ => NONE)
+ | prove NONE = NONE
+ in
+ get_first (instanciate_term thy (term_of redex) #> prove) mult_log_intros
+ end
+*}
+
+simproc_setup mult_log ("real (distribution X x) * log b (A * B)" |
+ "real (distribution X x) * log b (A / B)") = {* K mult_log_simproc *}
+
+end
+
+lemma (in finite_measure_space) absolutely_continuousI:
+ assumes "finite_measure_space M \<nu>"
+ assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
+ shows "absolutely_continuous \<nu>"
+proof (unfold absolutely_continuous_def sets_eq_Pow, safe)
+ fix N assume "\<mu> N = 0" "N \<subseteq> space M"
+
+ interpret v: finite_measure_space M \<nu> by fact
- fix x assume x: "x \<in> space M"
- thus pos: "0 \<le> u {x}" "0 \<le> v {x}"
- using u.positive u.sets_eq_Pow v.positive v.sets_eq_Pow by simp_all
+ have "\<nu> N = \<nu> (\<Union>x\<in>N. {x})" by simp
+ also have "\<dots> = (\<Sum>x\<in>N. \<nu> {x})"
+ proof (rule v.measure_finitely_additive''[symmetric])
+ show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset)
+ show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto
+ fix x assume "x \<in> N" thus "{x} \<in> sets M" using `N \<subseteq> space M` sets_eq_Pow by auto
+ qed
+ also have "\<dots> = 0"
+ proof (safe intro!: setsum_0')
+ fix x assume "x \<in> N"
+ hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono)
+ hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp
+ thus "\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
+ qed
+ finally show "\<nu> N = 0" .
+qed
+
+lemma (in finite_measure_space) KL_divergence_eq_finite:
+ assumes v: "finite_measure_space M \<nu>"
+ assumes ac: "\<forall>x\<in>space M. \<mu> {x} = 0 \<longrightarrow> \<nu> {x} = 0"
+ shows "KL_divergence b M \<nu> \<mu> = (\<Sum>x\<in>space M. real (\<nu> {x}) * log b (real (\<nu> {x}) / real (\<mu> {x})))" (is "_ = ?sum")
+proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v])
+ interpret v: finite_measure_space M \<nu> by fact
+ have ms: "measure_space M \<nu>" by fact
+
+ have ac: "absolutely_continuous \<nu>"
+ using ac by (auto intro!: absolutely_continuousI[OF v])
+
+ show "(\<Sum>x \<in> space M. log b (real (RN_deriv \<nu> x)) * real (\<nu> {x})) = ?sum"
+ using RN_deriv_finite_measure[OF ms ac]
+ by (auto intro!: setsum_cong simp: field_simps real_of_pinfreal_mult[symmetric])
+qed
- { assume "v {x} = 0" from u_0[OF x this] show "u {x} = 0" . }
- { assume "0 < u {x}"
- hence "v {x} \<noteq> 0" using u_0[OF x] by auto
- with pos show "0 < v {x}" by simp }
+lemma (in finite_prob_space) finite_sum_over_space_eq_1:
+ "(\<Sum>x\<in>space M. real (\<mu> {x})) = 1"
+ using sum_over_space_eq_1 finite_measure by (simp add: real_of_pinfreal_setsum sets_eq_Pow)
+
+lemma (in finite_prob_space) KL_divergence_positive_finite:
+ assumes v: "finite_prob_space M \<nu>"
+ assumes ac: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
+ and "1 < b"
+ shows "0 \<le> KL_divergence b M \<nu> \<mu>"
+proof -
+ interpret v: finite_prob_space M \<nu> using v .
+
+ have *: "space M \<noteq> {}" using not_empty by simp
+
+ hence "- (KL_divergence b M \<nu> \<mu>) \<le> log b (\<Sum>x\<in>space M. real (\<mu> {x}))"
+ proof (subst KL_divergence_eq_finite)
+ show "finite_measure_space M \<nu>" by fact
+
+ show "\<forall>x\<in>space M. \<mu> {x} = 0 \<longrightarrow> \<nu> {x} = 0" using ac by auto
+ show "- (\<Sum>x\<in>space M. real (\<nu> {x}) * log b (real (\<nu> {x}) / real (\<mu> {x}))) \<le> log b (\<Sum>x\<in>space M. real (\<mu> {x}))"
+ proof (safe intro!: log_setsum_divide *)
+ show "finite (space M)" using finite_space by simp
+ show "1 < b" by fact
+ show "(\<Sum>x\<in>space M. real (\<nu> {x})) = 1" using v.finite_sum_over_space_eq_1 by simp
+
+ fix x assume x: "x \<in> space M"
+ { assume "0 < real (\<nu> {x})"
+ hence "\<mu> {x} \<noteq> 0" using ac[OF x] by auto
+ thus "0 < prob {x}" using measure_finite[of "{x}"] sets_eq_Pow x
+ by (cases "\<mu> {x}") simp_all }
+ qed auto
qed
- thus "0 \<le> KL_divergence b M u v" using v.sum_over_space_eq_1 by simp
+ thus "0 \<le> KL_divergence b M \<nu> \<mu>" using finite_sum_over_space_eq_1 by simp
qed
definition (in prob_space)
- "mutual_information b s1 s2 X Y \<equiv>
- let prod_space =
- prod_measure_space (\<lparr>space = space s1, sets = sets s1, measure = distribution X\<rparr>)
- (\<lparr>space = space s2, sets = sets s2, measure = distribution Y\<rparr>)
- in
- KL_divergence b prod_space (joint_distribution X Y) (measure prod_space)"
+ "mutual_information b S T X Y =
+ KL_divergence b (prod_measure_space S T)
+ (joint_distribution X Y)
+ (prod_measure S (distribution X) T (distribution Y))"
abbreviation (in finite_information_space)
finite_mutual_information ("\<I>'(_ ; _')") where
@@ -275,20 +341,18 @@
\<lparr> space = X`space M, sets = Pow (X`space M) \<rparr>
\<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr> X Y"
-lemma (in finite_measure_space) measure_spaceI: "measure_space M"
- by unfold_locales
-
lemma prod_measure_times_finite:
- assumes fms: "finite_measure_space M" "finite_measure_space M'" and a: "a \<in> space M \<times> space M'"
- shows "prod_measure M M' {a} = measure M {fst a} * measure M' {snd a}"
+ assumes fms: "finite_measure_space M \<mu>" "finite_measure_space N \<nu>" and a: "a \<in> space M \<times> space N"
+ shows "prod_measure M \<mu> N \<nu> {a} = \<mu> {fst a} * \<nu> {snd a}"
proof (cases a)
case (Pair b c)
hence a_eq: "{a} = {b} \<times> {c}" by simp
- with fms[THEN finite_measure_space.measure_spaceI]
- fms[THEN finite_measure_space.sets_eq_Pow] a Pair
- show ?thesis unfolding a_eq
- by (subst prod_measure_times) simp_all
+ interpret M: finite_measure_space M \<mu> by fact
+ interpret N: finite_measure_space N \<nu> by fact
+
+ from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair
+ show ?thesis unfolding a_eq by simp
qed
lemma setsum_cartesian_product':
@@ -296,44 +360,44 @@
unfolding setsum_cartesian_product by simp
lemma (in finite_information_space)
- assumes MX: "finite_prob_space \<lparr> space = space MX, sets = sets MX, measure = distribution X\<rparr>"
- (is "finite_prob_space ?MX")
- assumes MY: "finite_prob_space \<lparr> space = space MY, sets = sets MY, measure = distribution Y\<rparr>"
- (is "finite_prob_space ?MY")
+ assumes MX: "finite_prob_space MX (distribution X)"
+ assumes MY: "finite_prob_space MY (distribution Y)"
and X_space: "X ` space M \<subseteq> space MX" and Y_space: "Y ` space M \<subseteq> space MY"
shows mutual_information_eq_generic:
"mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY.
- joint_distribution X Y {(x,y)} *
- log b (joint_distribution X Y {(x,y)} /
- (distribution X {x} * distribution Y {y})))"
+ real (joint_distribution X Y {(x,y)}) *
+ log b (real (joint_distribution X Y {(x,y)}) /
+ (real (distribution X {x}) * real (distribution Y {y}))))"
(is "?equality")
and mutual_information_positive_generic:
"0 \<le> mutual_information b MX MY X Y" (is "?positive")
proof -
- let ?P = "prod_measure_space ?MX ?MY"
- let ?measure = "joint_distribution X Y"
- let ?P' = "measure_update (\<lambda>_. ?measure) ?P"
+ let ?P = "prod_measure_space MX MY"
+ let ?\<mu> = "prod_measure MX (distribution X) MY (distribution Y)"
+ let ?\<nu> = "joint_distribution X Y"
- interpret X: finite_prob_space "?MX" using MX .
- moreover interpret Y: finite_prob_space "?MY" using MY .
- ultimately have ms_X: "measure_space ?MX"
- and ms_Y: "measure_space ?MY" by unfold_locales
+ interpret X: finite_prob_space MX "distribution X" by fact
+ moreover interpret Y: finite_prob_space MY "distribution Y" by fact
+ have ms_X: "measure_space MX (distribution X)"
+ and ms_Y: "measure_space MY (distribution Y)"
+ and fms: "finite_measure_space MX (distribution X)" "finite_measure_space MY (distribution Y)" by fact+
+ have fms_P: "finite_measure_space ?P ?\<mu>"
+ by (rule X.finite_measure_space_finite_prod_measure) fact
+ then interpret P: finite_measure_space ?P ?\<mu> .
- have fms_P: "finite_measure_space ?P"
- by (rule finite_measure_space_finite_prod_measure) fact+
-
- have fms_P': "finite_measure_space ?P'"
+ have fms_P': "finite_measure_space ?P ?\<nu>"
using finite_product_measure_space[of "space MX" "space MY"]
X.finite_space Y.finite_space sigma_prod_sets_finite[OF X.finite_space Y.finite_space]
X.sets_eq_Pow Y.sets_eq_Pow
- by (simp add: prod_measure_space_def)
+ by (simp add: prod_measure_space_def sigma_def)
+ then interpret P': finite_measure_space ?P ?\<nu> .
{ fix x assume "x \<in> space ?P"
- hence x_in_MX: "{fst x} \<in> sets MX" using X.sets_eq_Pow
+ hence in_MX: "{fst x} \<in> sets MX" "{snd x} \<in> sets MY" using X.sets_eq_Pow Y.sets_eq_Pow
by (auto simp: prod_measure_space_def)
- assume "measure ?P {x} = 0"
- with prod_measure_times[OF ms_X ms_Y, of "{fst x}" "{snd x}"] x_in_MX
+ assume "?\<mu> {x} = 0"
+ with X.finite_prod_measure_times[OF fms(2), of "{fst x}" "{snd x}"] in_MX
have "distribution X {fst x} = 0 \<or> distribution Y {snd x} = 0"
by (simp add: prod_measure_space_def)
@@ -342,33 +406,34 @@
note measure_0 = this
show ?equality
- unfolding Let_def mutual_information_def using fms_P fms_P' measure_0 MX MY
- by (subst KL_divergence_eq_finite)
- (simp_all add: prod_measure_space_def prod_measure_times_finite
- finite_prob_space_eq setsum_cartesian_product')
+ unfolding Let_def mutual_information_def
+ using measure_0 fms_P fms_P' MX MY P.absolutely_continuous_def
+ by (subst P.KL_divergence_eq_finite)
+ (auto simp add: prod_measure_space_def prod_measure_times_finite
+ finite_prob_space_eq setsum_cartesian_product' real_of_pinfreal_mult[symmetric])
show ?positive
unfolding Let_def mutual_information_def using measure_0 b_gt_1
- proof (safe intro!: KL_divergence_positive_finite, simp_all)
- from ms_X ms_Y X.top Y.top X.prob_space Y.prob_space
- have "measure ?P (space ?P) = 1"
- by (simp add: prod_measure_space_def, subst prod_measure_times, simp_all)
- with fms_P show "finite_prob_space ?P"
+ proof (safe intro!: finite_prob_space.KL_divergence_positive_finite, simp_all)
+ have "?\<mu> (space ?P) = 1"
+ using X.top Y.top X.measure_space_1 Y.measure_space_1 fms
+ by (simp add: prod_measure_space_def X.finite_prod_measure_times)
+ with fms_P show "finite_prob_space ?P ?\<mu>"
by (simp add: finite_prob_space_eq)
- from ms_X ms_Y X.top Y.top X.prob_space Y.prob_space Y.not_empty X_space Y_space
- have "measure ?P' (space ?P') = 1" unfolding prob_space[symmetric]
- by (auto simp add: prod_measure_space_def distribution_def vimage_Times comp_def
- intro!: arg_cong[where f=prob])
- with fms_P' show "finite_prob_space ?P'"
+ from ms_X ms_Y X.top Y.top X.measure_space_1 Y.measure_space_1 Y.not_empty X_space Y_space
+ have "?\<nu> (space ?P) = 1" unfolding measure_space_1[symmetric]
+ by (auto intro!: arg_cong[where f="\<mu>"]
+ simp add: prod_measure_space_def distribution_def vimage_Times comp_def)
+ with fms_P' show "finite_prob_space ?P ?\<nu>"
by (simp add: finite_prob_space_eq)
qed
qed
lemma (in finite_information_space) mutual_information_eq:
"\<I>(X;Y) = (\<Sum> (x,y) \<in> X ` space M \<times> Y ` space M.
- distribution (\<lambda>x. (X x, Y x)) {(x,y)} * log b (distribution (\<lambda>x. (X x, Y x)) {(x,y)} /
- (distribution X {x} * distribution Y {y})))"
+ real (distribution (\<lambda>x. (X x, Y x)) {(x,y)}) * log b (real (distribution (\<lambda>x. (X x, Y x)) {(x,y)}) /
+ (real (distribution X {x}) * real (distribution Y {y}))))"
by (subst mutual_information_eq_generic) (simp_all add: finite_prob_space_of_images)
lemma (in finite_information_space) mutual_information_positive: "0 \<le> \<I>(X;Y)"
@@ -383,18 +448,19 @@
lemma (in finite_information_space) joint_distribution_remove[simp]:
"joint_distribution X X {(x, x)} = distribution X {x}"
- unfolding distribution_def by (auto intro!: arg_cong[where f=prob])
+ unfolding distribution_def by (auto intro!: arg_cong[where f="\<mu>"])
lemma (in finite_information_space) entropy_eq:
- "\<H>(X) = -(\<Sum> x \<in> X ` space M. distribution X {x} * log b (distribution X {x}))"
+ "\<H>(X) = -(\<Sum> x \<in> X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))"
proof -
{ fix f
- { fix x y
- have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto
- hence "distribution (\<lambda>x. (X x, X x)) {(x,y)} * f x y = (if x = y then distribution X {x} * f x y else 0)"
- unfolding distribution_def by auto }
- hence "(\<Sum>(x, y) \<in> X ` space M \<times> X ` space M. joint_distribution X X {(x, y)} * f x y) =
- (\<Sum>x \<in> X ` space M. distribution X {x} * f x x)"
+ { fix x y
+ have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto
+ hence "real (distribution (\<lambda>x. (X x, X x)) {(x,y)}) * f x y =
+ (if x = y then real (distribution X {x}) * f x y else 0)"
+ unfolding distribution_def by auto }
+ hence "(\<Sum>(x, y) \<in> X ` space M \<times> X ` space M. real (joint_distribution X X {(x, y)}) * f x y) =
+ (\<Sum>x \<in> X ` space M. real (distribution X {x}) * f x x)"
unfolding setsum_cartesian_product' by (simp add: setsum_cases finite_space) }
note remove_cartesian_product = this
@@ -407,13 +473,9 @@
unfolding entropy_def using mutual_information_positive .
definition (in prob_space)
- "conditional_mutual_information b s1 s2 s3 X Y Z \<equiv>
- let prod_space =
- prod_measure_space \<lparr>space = space s2, sets = sets s2, measure = distribution Y\<rparr>
- \<lparr>space = space s3, sets = sets s3, measure = distribution Z\<rparr>
- in
- mutual_information b s1 prod_space X (\<lambda>x. (Y x, Z x)) -
- mutual_information b s1 s3 X Z"
+ "conditional_mutual_information b M1 M2 M3 X Y Z \<equiv>
+ mutual_information b M1 (prod_measure_space M2 M3) X (\<lambda>x. (Y x, Z x)) -
+ mutual_information b M1 M3 X Z"
abbreviation (in finite_information_space)
finite_conditional_mutual_information ("\<I>'( _ ; _ | _ ')") where
@@ -441,19 +503,37 @@
"(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
by (auto intro!: inj_onI setsum_distribution_gen)
+lemma (in finite_information_space) setsum_real_distribution_gen:
+ assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
+ and "inj_on f (X`space M)"
+ shows "(\<Sum>x \<in> X`space M. real (distribution Y {f x})) = real (distribution Z {c})"
+ unfolding distribution_def assms
+ using finite_space assms
+ by (subst real_finite_measure_finite_Union[symmetric])
+ (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
+ intro!: arg_cong[where f=prob])
+
+lemma (in finite_information_space) setsum_real_distribution:
+ "(\<Sum>x \<in> X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})"
+ "(\<Sum>y \<in> Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})"
+ "(\<Sum>x \<in> X`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})"
+ "(\<Sum>y \<in> Y`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})"
+ "(\<Sum>z \<in> Z`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})"
+ by (auto intro!: inj_onI setsum_real_distribution_gen)
+
lemma (in finite_information_space) conditional_mutual_information_eq_sum:
"\<I>(X ; Y | Z) =
(\<Sum>(x, y, z)\<in>X ` space M \<times> (\<lambda>x. (Y x, Z x)) ` space M.
- distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
- log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}/
- distribution (\<lambda>x. (Y x, Z x)) {(y, z)})) -
+ real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
+ log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)})/
+ real (distribution (\<lambda>x. (Y x, Z x)) {(y, z)}))) -
(\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M.
- distribution (\<lambda>x. (X x, Z x)) {(x,z)} * log b (distribution (\<lambda>x. (X x, Z x)) {(x,z)} / distribution Z {z}))"
+ real (distribution (\<lambda>x. (X x, Z x)) {(x,z)}) * log b (real (distribution (\<lambda>x. (X x, Z x)) {(x,z)}) / real (distribution Z {z})))"
(is "_ = ?rhs")
proof -
have setsum_product:
- "\<And>f x. (\<Sum>v\<in>(\<lambda>x. (Y x, Z x)) ` space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x,v)} * f v)
- = (\<Sum>v\<in>Y ` space M \<times> Z ` space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x,v)} * f v)"
+ "\<And>f x. (\<Sum>v\<in>(\<lambda>x. (Y x, Z x)) ` space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x,v)}) * f v)
+ = (\<Sum>v\<in>Y ` space M \<times> Z ` space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x,v)}) * f v)"
proof (safe intro!: setsum_mono_zero_cong_left imageI)
fix x y z f
assume *: "(Y y, Z z) \<notin> (\<lambda>x. (Y x, Z x)) ` space M" and "y \<in> space M" "z \<in> space M"
@@ -463,31 +543,32 @@
have "(Y y, Z z) \<in> (\<lambda>x. (Y x, Z x)) ` space M" using eq[symmetric] x' by auto
thus "x' \<in> {}" using * by auto
qed
- thus "joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, Y y, Z z)} * f (Y y) (Z z) = 0"
+ thus "real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, Y y, Z z)}) * f (Y y) (Z z) = 0"
unfolding distribution_def by simp
qed (simp add: finite_space)
thus ?thesis
unfolding conditional_mutual_information_def Let_def mutual_information_eq
- apply (subst mutual_information_eq_generic)
- by (auto simp add: prod_measure_space_def sigma_prod_sets_finite finite_space
+ by (subst mutual_information_eq_generic)
+ (auto simp: prod_measure_space_def sigma_prod_sets_finite finite_space sigma_def
finite_prob_space_of_images finite_product_prob_space_of_images
setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf
- setsum_left_distrib[symmetric] setsum_distribution
+ setsum_left_distrib[symmetric] setsum_real_distribution
cong: setsum_cong)
qed
lemma (in finite_information_space) conditional_mutual_information_eq:
"\<I>(X ; Y | Z) = (\<Sum>(x, y, z) \<in> X ` space M \<times> Y ` space M \<times> Z ` space M.
- distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
- log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}/
- (joint_distribution X Z {(x, z)} * joint_distribution Y Z {(y,z)} / distribution Z {z})))"
+ real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
+ log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) /
+ (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
unfolding conditional_mutual_information_def Let_def mutual_information_eq
- apply (subst mutual_information_eq_generic)
- by (auto simp add: prod_measure_space_def sigma_prod_sets_finite finite_space
- finite_prob_space_of_images finite_product_prob_space_of_images
+ by (subst mutual_information_eq_generic)
+ (auto simp add: prod_measure_space_def sigma_prod_sets_finite finite_space
+ finite_prob_space_of_images finite_product_prob_space_of_images sigma_def
setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf
- setsum_left_distrib[symmetric] setsum_distribution setsum_commute[where A="Y`space M"]
+ setsum_left_distrib[symmetric] setsum_real_distribution setsum_commute[where A="Y`space M"]
+ real_of_pinfreal_mult[symmetric]
cong: setsum_cong)
lemma (in finite_information_space) conditional_mutual_information_eq_mutual_information:
@@ -500,14 +581,30 @@
by (simp add: setsum_cartesian_product' distribution_remove_const)
qed
+lemma (in finite_prob_space) distribution_finite:
+ "distribution X A \<noteq> \<omega>"
+ by (auto simp: sets_eq_Pow distribution_def intro!: measure_finite)
+
+lemma (in finite_prob_space) real_distribution_order:
+ shows "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution X {x})"
+ and "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution Y {y})"
+ and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution X {x})"
+ and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution Y {y})"
+ and "distribution X {x} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
+ and "distribution Y {y} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
+ using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
+ using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
+ using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"]
+ by auto
+
lemma (in finite_information_space) conditional_mutual_information_positive:
"0 \<le> \<I>(X ; Y | Z)"
proof -
- let ?dXYZ = "distribution (\<lambda>x. (X x, Y x, Z x))"
- let ?dXZ = "joint_distribution X Z"
- let ?dYZ = "joint_distribution Y Z"
- let ?dX = "distribution X"
- let ?dZ = "distribution Z"
+ let "?dXYZ A" = "real (distribution (\<lambda>x. (X x, Y x, Z x)) A)"
+ let "?dXZ A" = "real (joint_distribution X Z A)"
+ let "?dYZ A" = "real (joint_distribution Y Z A)"
+ let "?dX A" = "real (distribution X A)"
+ let "?dZ A" = "real (distribution Z A)"
let ?M = "X ` space M \<times> Y ` space M \<times> Z ` space M"
have split_beta: "\<And>f. split f = (\<lambda>x. f (fst x) (snd x))" by (simp add: expand_fun_eq)
@@ -521,26 +618,28 @@
show "1 < b" using b_gt_1 .
fix x assume "x \<in> ?M"
- show "0 \<le> ?dXYZ {(fst x, fst (snd x), snd (snd x))}" using positive_distribution .
+ let ?x = "(fst x, fst (snd x), snd (snd x))"
+
+ show "0 \<le> ?dXYZ {?x}" using real_pinfreal_nonneg .
show "0 \<le> ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
- by (auto intro!: mult_nonneg_nonneg positive_distribution simp: zero_le_divide_iff)
+ by (simp add: real_pinfreal_nonneg mult_nonneg_nonneg divide_nonneg_nonneg)
- assume *: "0 < ?dXYZ {(fst x, fst (snd x), snd (snd x))}"
+ assume *: "0 < ?dXYZ {?x}"
thus "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
- by (auto intro!: divide_pos_pos mult_pos_pos
- intro: distribution_order(6) distribution_mono_gt_0)
- qed (simp_all add: setsum_cartesian_product' sum_over_space_distrib setsum_distribution finite_space)
+ apply (rule_tac divide_pos_pos mult_pos_pos)+
+ by (auto simp add: real_distribution_gt_0 intro: distribution_order(6) distribution_mono_gt_0)
+ qed (simp_all add: setsum_cartesian_product' sum_over_space_real_distribution setsum_real_distribution finite_space)
also have "(\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\<Sum>z\<in>Z`space M. ?dZ {z})"
apply (simp add: setsum_cartesian_product')
apply (subst setsum_commute)
apply (subst (2) setsum_commute)
- by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] setsum_distribution
+ by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] setsum_real_distribution
intro!: setsum_cong)
finally show ?thesis
- unfolding conditional_mutual_information_eq sum_over_space_distrib by simp
+ unfolding conditional_mutual_information_eq sum_over_space_real_distribution
+ by (simp add: real_of_pinfreal_mult[symmetric])
qed
-
definition (in prob_space)
"conditional_entropy b S T X Y = conditional_mutual_information b S S T X X Y"
@@ -556,8 +655,8 @@
lemma (in finite_information_space) conditional_entropy_eq:
"\<H>(X | Z) =
- (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M.
- joint_distribution X Z {(x, z)} *
- log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
+ real (joint_distribution X Z {(x, z)}) *
+ log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
proof -
have *: "\<And>x y z. (\<lambda>x. (X x, X x, Z x)) -` {(x, y, z)} = (if x = y then (\<lambda>x. (X x, Z x)) -` {(x, z)} else {})" by auto
show ?thesis
@@ -571,7 +670,7 @@
unfolding mutual_information_eq entropy_eq conditional_entropy_eq
using finite_space
by (auto simp add: setsum_addf setsum_subtractf setsum_cartesian_product'
- setsum_left_distrib[symmetric] setsum_addf setsum_distribution)
+ setsum_left_distrib[symmetric] setsum_addf setsum_real_distribution)
lemma (in finite_information_space) conditional_entropy_less_eq_entropy:
"\<H>(X | Z) \<le> \<H>(X)"
@@ -587,9 +686,8 @@
assumes "x \<in> X ` space M" and "distribution X {x} = 1"
shows "\<H>(X) = 0"
proof -
- interpret X: finite_prob_space "\<lparr> space = X ` space M,
- sets = Pow (X ` space M),
- measure = distribution X\<rparr>" by (rule finite_prob_space_of_images)
+ interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X"
+ by (rule finite_prob_space_of_images)
have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
using X.measure_compl[of "{x}"] assms by auto
@@ -598,10 +696,10 @@
{ fix y assume asm: "y \<noteq> x" "y \<in> X ` space M"
hence "{y} \<subseteq> X ` space M - {x}" by auto
- from X.measure_mono[OF this] X0 X.positive[of "{y}"] asm
+ from X.measure_mono[OF this] X0 asm
have "distribution X {y} = 0" by auto }
- hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> distribution X {y} = (if x = y then 1 else 0)"
+ hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> real (distribution X {y}) = (if x = y then 1 else 0)"
using assms by auto
have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp
@@ -610,71 +708,32 @@
qed
(* --------------- upper bound on entropy for a rv ------------------------- *)
+lemma (in finite_prob_space) distribution_1:
+ "distribution X A \<le> 1"
+ unfolding distribution_def measure_space_1[symmetric]
+ by (auto intro!: measure_mono simp: sets_eq_Pow)
+
+lemma (in finite_prob_space) real_distribution_1:
+ "real (distribution X A) \<le> 1"
+ unfolding real_pinfreal_1[symmetric]
+ by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp
+
lemma (in finite_information_space) finite_entropy_le_card:
"\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))"
proof -
- interpret X: finite_prob_space "\<lparr>space = X ` space M,
- sets = Pow (X ` space M),
- measure = distribution X\<rparr>"
- using finite_prob_space_of_images by auto
-
- have triv: "\<And> x. (if distribution X {x} \<noteq> 0 then distribution X {x} else 0) = distribution X {x}"
+ let "?d x" = "distribution X {x}"
+ let "?p x" = "real (?d x)"
+ have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))"
+ by (auto intro!: setsum_cong simp: entropy_eq setsum_negf[symmetric])
+ also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))"
+ apply (rule log_setsum')
+ using not_empty b_gt_1 finite_space sum_over_space_real_distribution
by auto
- hence sum1: "(\<Sum> x \<in> X ` space M \<inter> {y. distribution X {y} \<noteq> 0}. distribution X {x}) = 1"
- using X.measure_finitely_additive''[of "X ` space M" "\<lambda> x. {x}", simplified]
- sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
- unfolding disjoint_family_on_def X.prob_space[symmetric]
- using finite_imageI[OF finite_space, of X] by (auto simp add:triv setsum_restrict_set)
- have pos: "\<And> x. x \<in> X ` space M \<inter> {y. distribution X {y} \<noteq> 0} \<Longrightarrow> inverse (distribution X {x}) > 0"
- using X.positive sets_eq_Pow unfolding inverse_positive_iff_positive less_le by auto
- { assume asm: "X ` space M \<inter> {y. distribution X {y} \<noteq> 0} = {}"
- { fix x assume "x \<in> X ` space M"
- hence "distribution X {x} = 0" using asm by blast }
- hence A: "(\<Sum> x \<in> X ` space M. distribution X {x}) = 0" by auto
- have B: "(\<Sum> x \<in> X ` space M. distribution X {x})
- \<ge> (\<Sum> x \<in> X ` space M \<inter> {y. distribution X {y} \<noteq> 0}. distribution X {x})"
- using finite_imageI[OF finite_space, of X]
- by (subst setsum_mono2) auto
- from A B have "False" using sum1 by auto } note not_empty = this
- { fix x assume asm: "x \<in> X ` space M"
- have "- distribution X {x} * log b (distribution X {x})
- = - (if distribution X {x} \<noteq> 0
- then distribution X {x} * log b (distribution X {x})
- else 0)"
- by auto
- also have "\<dots> = (if distribution X {x} \<noteq> 0
- then distribution X {x} * - log b (distribution X {x})
- else 0)"
- by auto
- also have "\<dots> = (if distribution X {x} \<noteq> 0
- then distribution X {x} * log b (inverse (distribution X {x}))
- else 0)"
- using log_inverse b_gt_1 X.positive[of "{x}"] asm by auto
- finally have "- distribution X {x} * log b (distribution X {x})
- = (if distribution X {x} \<noteq> 0
- then distribution X {x} * log b (inverse (distribution X {x}))
- else 0)"
- by auto } note log_inv = this
- have "- (\<Sum> x \<in> X ` space M. distribution X {x} * log b (distribution X {x}))
- = (\<Sum> x \<in> X ` space M. (if distribution X {x} \<noteq> 0
- then distribution X {x} * log b (inverse (distribution X {x}))
- else 0))"
- unfolding setsum_negf[symmetric] using log_inv by auto
- also have "\<dots> = (\<Sum> x \<in> X ` space M \<inter> {y. distribution X {y} \<noteq> 0}.
- distribution X {x} * log b (inverse (distribution X {x})))"
- unfolding setsum_restrict_set[OF finite_imageI[OF finite_space, of X]] by auto
- also have "\<dots> \<le> log b (\<Sum> x \<in> X ` space M \<inter> {y. distribution X {y} \<noteq> 0}.
- distribution X {x} * (inverse (distribution X {x})))"
- apply (subst log_setsum[OF _ _ b_gt_1 sum1,
- unfolded greaterThan_iff, OF _ _ _]) using pos sets_eq_Pow
- X.finite_space assms X.positive not_empty by auto
- also have "\<dots> = log b (\<Sum> x \<in> X ` space M \<inter> {y. distribution X {y} \<noteq> 0}. 1)"
- by auto
- also have "\<dots> \<le> log b (real_of_nat (card (X ` space M \<inter> {y. distribution X {y} \<noteq> 0})))"
- by auto
- finally have "- (\<Sum>x\<in>X ` space M. distribution X {x} * log b (distribution X {x}))
- \<le> log b (real_of_nat (card (X ` space M \<inter> {y. distribution X {y} \<noteq> 0})))" by simp
- thus ?thesis unfolding entropy_eq real_eq_of_nat by auto
+ also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?d x \<noteq> 0 then 1 else 0)"
+ apply (rule arg_cong[where f="\<lambda>f. log b (\<Sum>x\<in>X`space M. f x)"])
+ using distribution_finite[of X] by (auto simp: expand_fun_eq real_of_pinfreal_eq_0)
+ finally show ?thesis
+ using finite_space by (auto simp: setsum_cases real_eq_of_nat)
qed
(* --------------- entropy is maximal for a uniform rv --------------------- *)
@@ -689,7 +748,7 @@
have "1 = prob (space M)"
using prob_space by auto
also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
- using measure_finitely_additive''[of "space M" "\<lambda> x. {x}", simplified]
+ using real_finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
finite_space unfolding disjoint_family_on_def prob_space[symmetric]
by (auto simp add:setsum_restrict_set)
@@ -708,33 +767,21 @@
assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
shows "\<H>(X) = log b (real (card (X ` space M)))"
proof -
- interpret X: finite_prob_space "\<lparr>space = X ` space M,
- sets = Pow (X ` space M),
- measure = distribution X\<rparr>"
- using finite_prob_space_of_images by auto
+ note uniform =
+ finite_prob_space_of_images[of X, THEN finite_prob_space.uniform_prob, simplified]
+
+ have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
+ using finite_space not_empty by auto
- { fix x assume xasm: "x \<in> X ` space M"
- hence card_gt0: "real (card (X ` space M)) > 0"
- using card_gt_0_iff X.finite_space by auto
- from xasm have "\<And> y. y \<in> X ` space M \<Longrightarrow> distribution X {y} = distribution X {x}"
- using assms by blast
- hence "- (\<Sum>x\<in>X ` space M. distribution X {x} * log b (distribution X {x}))
- = - real (card (X ` space M)) * distribution X {x} * log b (distribution X {x})"
- unfolding real_eq_of_nat by auto
- also have "\<dots> = - real (card (X ` space M)) * (1 / real (card (X ` space M))) * log b (1 / real (card (X ` space M)))"
- by (auto simp: X.uniform_prob[simplified, OF xasm assms])
- also have "\<dots> = log b (real (card (X ` space M)))"
- unfolding inverse_eq_divide[symmetric]
- using card_gt0 log_inverse b_gt_1
- by (auto simp add:field_simps card_gt0)
- finally have ?thesis
- unfolding entropy_eq by auto }
- moreover
- { assume "X ` space M = {}"
- hence "distribution X (X ` space M) = 0"
- using X.empty_measure by simp
- hence "False" using X.prob_space by auto }
- ultimately show ?thesis by auto
+ { fix x assume "x \<in> X ` space M"
+ hence "real (distribution X {x}) = 1 / real (card (X ` space M))"
+ proof (rule uniform)
+ fix x y assume "x \<in> X`space M" "y \<in> X`space M"
+ from assms[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp
+ qed }
+ thus ?thesis
+ using not_empty finite_space b_gt_1 card_gt0
+ by (simp add: entropy_eq real_eq_of_nat[symmetric] log_divide)
qed
definition "subvimage A f g \<longleftrightarrow> (\<forall>x \<in> A. f -` {f x} \<inter> A \<subseteq> g -` {g x} \<inter> A)"
@@ -854,13 +901,13 @@
assumes svi: "subvimage (space M) X P"
shows "\<H>(X) = \<H>(P) + \<H>(X|P)"
proof -
- have "(\<Sum>x\<in>X ` space M. distribution X {x} * log b (distribution X {x})) =
+ have "(\<Sum>x\<in>X ` space M. real (distribution X {x}) * log b (real (distribution X {x}))) =
(\<Sum>y\<in>P `space M. \<Sum>x\<in>X ` space M.
- joint_distribution X P {(x, y)} * log b (joint_distribution X P {(x, y)}))"
+ real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})))"
proof (subst setsum_image_split[OF svi],
safe intro!: finite_imageI finite_space setsum_mono_zero_cong_left imageI)
fix p x assume in_space: "p \<in> space M" "x \<in> space M"
- assume "joint_distribution X P {(X x, P p)} * log b (joint_distribution X P {(X x, P p)}) \<noteq> 0"
+ assume "real (joint_distribution X P {(X x, P p)}) * log b (real (joint_distribution X P {(X x, P p)})) \<noteq> 0"
hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M \<noteq> {}" by (auto simp: distribution_def)
with svi[unfolded subvimage_def, rule_format, OF `x \<in> space M`]
show "x \<in> P -` {P p}" by auto
@@ -872,14 +919,14 @@
by auto
hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M = X -` {X x} \<inter> space M"
by auto
- thus "distribution X {X x} * log b (distribution X {X x}) =
- joint_distribution X P {(X x, P p)} *
- log b (joint_distribution X P {(X x, P p)})"
+ thus "real (distribution X {X x}) * log b (real (distribution X {X x})) =
+ real (joint_distribution X P {(X x, P p)}) *
+ log b (real (joint_distribution X P {(X x, P p)}))"
by (auto simp: distribution_def)
qed
thus ?thesis
unfolding entropy_eq conditional_entropy_eq
- by (simp add: setsum_cartesian_product' setsum_subtractf setsum_distribution
+ by (simp add: setsum_cartesian_product' setsum_subtractf setsum_real_distribution
setsum_left_distrib[symmetric] setsum_commute[where B="P`space M"])
qed
@@ -891,14 +938,14 @@
assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
shows "distribution X = distribution Y"
unfolding distribution_def expand_fun_eq
- using assms by (auto intro!: arg_cong[where f=prob])
+ using assms by (auto intro!: arg_cong[where f="\<mu>"])
lemma (in prob_space) joint_distribution_cong:
assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
shows "joint_distribution X Y = joint_distribution X' Y'"
unfolding distribution_def expand_fun_eq
- using assms by (auto intro!: arg_cong[where f=prob])
+ using assms by (auto intro!: arg_cong[where f="\<mu>"])
lemma image_cong:
"\<lbrakk> \<And>x. x \<in> S \<Longrightarrow> X x = X' x \<rbrakk> \<Longrightarrow> X ` S = X' ` S"
--- a/src/HOL/Probability/Lebesgue.thy Mon Aug 23 16:47:57 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1733 +0,0 @@
-header {*Lebesgue Integration*}
-
-theory Lebesgue
-imports Measure Borel
-begin
-
-text{*From the HOL4 Hurd/Coble Lebesgue integration, translated by Armin Heller and Johannes Hoelzl.*}
-
-definition
- "pos_part f = (\<lambda>x. max 0 (f x))"
-
-definition
- "neg_part f = (\<lambda>x. - min 0 (f x))"
-
-definition
- "nonneg f = (\<forall>x. 0 \<le> f x)"
-
-lemma nonneg_pos_part[intro!]:
- fixes f :: "'c \<Rightarrow> 'd\<Colon>{linorder,zero}"
- shows "nonneg (pos_part f)"
- unfolding nonneg_def pos_part_def by auto
-
-lemma nonneg_neg_part[intro!]:
- fixes f :: "'c \<Rightarrow> 'd\<Colon>{linorder,ordered_ab_group_add}"
- shows "nonneg (neg_part f)"
- unfolding nonneg_def neg_part_def min_def by auto
-
-lemma pos_neg_part_abs:
- fixes f :: "'a \<Rightarrow> real"
- shows "pos_part f x + neg_part f x = \<bar>f x\<bar>"
-unfolding abs_if pos_part_def neg_part_def by auto
-
-lemma pos_part_abs:
- fixes f :: "'a \<Rightarrow> real"
- shows "pos_part (\<lambda> x. \<bar>f x\<bar>) y = \<bar>f y\<bar>"
-unfolding pos_part_def abs_if by auto
-
-lemma neg_part_abs:
- fixes f :: "'a \<Rightarrow> real"
- shows "neg_part (\<lambda> x. \<bar>f x\<bar>) y = 0"
-unfolding neg_part_def abs_if by auto
-
-lemma (in measure_space)
- assumes "f \<in> borel_measurable M"
- shows pos_part_borel_measurable: "pos_part f \<in> borel_measurable M"
- and neg_part_borel_measurable: "neg_part f \<in> borel_measurable M"
-using assms
-proof -
- { fix a :: real
- { assume asm: "0 \<le> a"
- from asm have pp: "\<And> w. (pos_part f w \<le> a) = (f w \<le> a)" unfolding pos_part_def by auto
- have "{w | w. w \<in> space M \<and> f w \<le> a} \<in> sets M"
- unfolding pos_part_def using assms borel_measurable_le_iff by auto
- hence "{w . w \<in> space M \<and> pos_part f w \<le> a} \<in> sets M" using pp by auto }
- moreover have "a < 0 \<Longrightarrow> {w \<in> space M. pos_part f w \<le> a} \<in> sets M"
- unfolding pos_part_def using empty_sets by auto
- ultimately have "{w . w \<in> space M \<and> pos_part f w \<le> a} \<in> sets M"
- using le_less_linear by auto
- } hence pos: "pos_part f \<in> borel_measurable M" using borel_measurable_le_iff by auto
- { fix a :: real
- { assume asm: "0 \<le> a"
- from asm have pp: "\<And> w. (neg_part f w \<le> a) = (f w \<ge> - a)" unfolding neg_part_def by auto
- have "{w | w. w \<in> space M \<and> f w \<ge> - a} \<in> sets M"
- unfolding neg_part_def using assms borel_measurable_ge_iff by auto
- hence "{w . w \<in> space M \<and> neg_part f w \<le> a} \<in> sets M" using pp by auto }
- moreover have "a < 0 \<Longrightarrow> {w \<in> space M. neg_part f w \<le> a} = {}" unfolding neg_part_def by auto
- moreover hence "a < 0 \<Longrightarrow> {w \<in> space M. neg_part f w \<le> a} \<in> sets M" by (simp only: empty_sets)
- ultimately have "{w . w \<in> space M \<and> neg_part f w \<le> a} \<in> sets M"
- using le_less_linear by auto
- } hence neg: "neg_part f \<in> borel_measurable M" using borel_measurable_le_iff by auto
- from pos neg show "pos_part f \<in> borel_measurable M" and "neg_part f \<in> borel_measurable M" by auto
-qed
-
-lemma (in measure_space) pos_part_neg_part_borel_measurable_iff:
- "f \<in> borel_measurable M \<longleftrightarrow>
- pos_part f \<in> borel_measurable M \<and> neg_part f \<in> borel_measurable M"
-proof -
- { fix x
- have "pos_part f x - neg_part f x = f x"
- unfolding pos_part_def neg_part_def unfolding max_def min_def
- by auto }
- hence "(\<lambda> x. pos_part f x - neg_part f x) = (\<lambda>x. f x)" by auto
- hence f: "(\<lambda> x. pos_part f x - neg_part f x) = f" by blast
- from pos_part_borel_measurable[of f] neg_part_borel_measurable[of f]
- borel_measurable_diff_borel_measurable[of "pos_part f" "neg_part f"]
- show ?thesis unfolding f by fast
-qed
-
-context measure_space
-begin
-
-section "Simple discrete step function"
-
-definition
- "pos_simple f =
- { (s :: nat set, a, x).
- finite s \<and> nonneg f \<and> nonneg x \<and> a ` s \<subseteq> sets M \<and>
- (\<forall>t \<in> space M. (\<exists>!i\<in>s. t\<in>a i) \<and> (\<forall>i\<in>s. t \<in> a i \<longrightarrow> f t = x i)) }"
-
-definition
- "pos_simple_integral \<equiv> \<lambda>(s, a, x). \<Sum> i \<in> s. x i * measure M (a i)"
-
-definition
- "psfis f = pos_simple_integral ` (pos_simple f)"
-
-lemma pos_simpleE[consumes 1]:
- assumes ps: "(s, a, x) \<in> pos_simple f"
- obtains "finite s" and "nonneg f" and "nonneg x"
- and "a ` s \<subseteq> sets M" and "(\<Union>i\<in>s. a i) = space M"
- and "disjoint_family_on a s"
- and "\<And>t. t \<in> space M \<Longrightarrow> (\<exists>!i. i \<in> s \<and> t \<in> a i)"
- and "\<And>t i. \<lbrakk> t \<in> space M ; i \<in> s ; t \<in> a i \<rbrakk> \<Longrightarrow> f t = x i"
-proof
- show "finite s" and "nonneg f" and "nonneg x"
- and as_in_M: "a ` s \<subseteq> sets M"
- and *: "\<And>t. t \<in> space M \<Longrightarrow> (\<exists>!i. i \<in> s \<and> t \<in> a i)"
- and **: "\<And>t i. \<lbrakk> t \<in> space M ; i \<in> s ; t \<in> a i \<rbrakk> \<Longrightarrow> f t = x i"
- using ps unfolding pos_simple_def by auto
-
- thus t: "(\<Union>i\<in>s. a i) = space M"
- proof safe
- fix x assume "x \<in> space M"
- from *[OF this] show "x \<in> (\<Union>i\<in>s. a i)" by auto
- next
- fix t i assume "i \<in> s"
- hence "a i \<in> sets M" using as_in_M by auto
- moreover assume "t \<in> a i"
- ultimately show "t \<in> space M" using sets_into_space by blast
- qed
-
- show "disjoint_family_on a s" unfolding disjoint_family_on_def
- proof safe
- fix i j and t assume "i \<in> s" "t \<in> a i" and "j \<in> s" "t \<in> a j" and "i \<noteq> j"
- with t * show "t \<in> {}" by auto
- qed
-qed
-
-lemma pos_simple_cong:
- assumes "nonneg f" and "nonneg g" and "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
- shows "pos_simple f = pos_simple g"
- unfolding pos_simple_def using assms by auto
-
-lemma psfis_cong:
- assumes "nonneg f" and "nonneg g" and "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
- shows "psfis f = psfis g"
- unfolding psfis_def using pos_simple_cong[OF assms] by simp
-
-lemma psfis_0: "0 \<in> psfis (\<lambda>x. 0)"
- unfolding psfis_def pos_simple_def pos_simple_integral_def
- by (auto simp: nonneg_def
- intro: image_eqI[where x="({0}, (\<lambda>i. space M), (\<lambda>i. 0))"])
-
-lemma pos_simple_setsum_indicator_fn:
- assumes ps: "(s, a, x) \<in> pos_simple f"
- and "t \<in> space M"
- shows "(\<Sum>i\<in>s. x i * indicator_fn (a i) t) = f t"
-proof -
- from assms obtain i where *: "i \<in> s" "t \<in> a i"
- and "finite s" and xi: "x i = f t" by (auto elim!: pos_simpleE)
-
- have **: "(\<Sum>i\<in>s. x i * indicator_fn (a i) t) =
- (\<Sum>j\<in>s. if j \<in> {i} then x i else 0)"
- unfolding indicator_fn_def using assms *
- by (auto intro!: setsum_cong elim!: pos_simpleE)
- show ?thesis unfolding ** setsum_cases[OF `finite s`] xi
- using `i \<in> s` by simp
-qed
-
-lemma pos_simple_common_partition:
- assumes psf: "(s, a, x) \<in> pos_simple f"
- assumes psg: "(s', b, y) \<in> pos_simple g"
- obtains z z' c k where "(k, c, z) \<in> pos_simple f" "(k, c, z') \<in> pos_simple g"
-proof -
- (* definitions *)
-
- def k == "{0 ..< card (s \<times> s')}"
- have fs: "finite s" "finite s'" "finite k" unfolding k_def
- using psf psg unfolding pos_simple_def by auto
- hence "finite (s \<times> s')" by simp
- then obtain p where p: "p ` k = s \<times> s'" "inj_on p k"
- using ex_bij_betw_nat_finite[of "s \<times> s'"] unfolding bij_betw_def k_def by blast
- def c == "\<lambda> i. a (fst (p i)) \<inter> b (snd (p i))"
- def z == "\<lambda> i. x (fst (p i))"
- def z' == "\<lambda> i. y (snd (p i))"
-
- have "finite k" unfolding k_def by simp
-
- have "nonneg z" and "nonneg z'"
- using psf psg unfolding z_def z'_def pos_simple_def nonneg_def by auto
-
- have ck_subset_M: "c ` k \<subseteq> sets M"
- proof
- fix x assume "x \<in> c ` k"
- then obtain j where "j \<in> k" and "x = c j" by auto
- hence "p j \<in> s \<times> s'" using p(1) by auto
- hence "a (fst (p j)) \<in> sets M" (is "?a \<in> _")
- and "b (snd (p j)) \<in> sets M" (is "?b \<in> _")
- using psf psg unfolding pos_simple_def by auto
- thus "x \<in> sets M" unfolding `x = c j` c_def using Int by simp
- qed
-
- { fix t assume "t \<in> space M"
- hence ex1s: "\<exists>!i\<in>s. t \<in> a i" and ex1s': "\<exists>!i\<in>s'. t \<in> b i"
- using psf psg unfolding pos_simple_def by auto
- then obtain j j' where j: "j \<in> s" "t \<in> a j" and j': "j' \<in> s'" "t \<in> b j'"
- by auto
- then obtain i :: nat where i: "(j,j') = p i" and "i \<in> k" using p by auto
- have "\<exists>!i\<in>k. t \<in> c i"
- proof (rule ex1I[of _ i])
- show "\<And>x. x \<in> k \<Longrightarrow> t \<in> c x \<Longrightarrow> x = i"
- proof -
- fix l assume "l \<in> k" "t \<in> c l"
- hence "p l \<in> s \<times> s'" and t_in: "t \<in> a (fst (p l))" "t \<in> b (snd (p l))"
- using p unfolding c_def by auto
- hence "fst (p l) \<in> s" and "snd (p l) \<in> s'" by auto
- with t_in j j' ex1s ex1s' have "p l = (j, j')" by (cases "p l", auto)
- thus "l = i"
- using `(j, j') = p i` p(2)[THEN inj_onD] `l \<in> k` `i \<in> k` by auto
- qed
-
- show "i \<in> k \<and> t \<in> c i"
- using `i \<in> k` `t \<in> a j` `t \<in> b j'` c_def i[symmetric] by auto
- qed auto
- } note ex1 = this
-
- show thesis
- proof (rule that)
- { fix t i assume "t \<in> space M" and "i \<in> k"
- hence "p i \<in> s \<times> s'" using p(1) by auto
- hence "fst (p i) \<in> s" by auto
- moreover
- assume "t \<in> c i"
- hence "t \<in> a (fst (p i))" unfolding c_def by auto
- ultimately have "f t = z i" using psf `t \<in> space M`
- unfolding z_def pos_simple_def by auto
- }
- thus "(k, c, z) \<in> pos_simple f"
- using psf `finite k` `nonneg z` ck_subset_M ex1
- unfolding pos_simple_def by auto
-
- { fix t i assume "t \<in> space M" and "i \<in> k"
- hence "p i \<in> s \<times> s'" using p(1) by auto
- hence "snd (p i) \<in> s'" by auto
- moreover
- assume "t \<in> c i"
- hence "t \<in> b (snd (p i))" unfolding c_def by auto
- ultimately have "g t = z' i" using psg `t \<in> space M`
- unfolding z'_def pos_simple_def by auto
- }
- thus "(k, c, z') \<in> pos_simple g"
- using psg `finite k` `nonneg z'` ck_subset_M ex1
- unfolding pos_simple_def by auto
- qed
-qed
-
-lemma pos_simple_integral_equal:
- assumes psx: "(s, a, x) \<in> pos_simple f"
- assumes psy: "(s', b, y) \<in> pos_simple f"
- shows "pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)"
- unfolding pos_simple_integral_def
-proof simp
- have "(\<Sum>i\<in>s. x i * measure M (a i)) =
- (\<Sum>i\<in>s. (\<Sum>j \<in> s'. x i * measure M (a i \<inter> b j)))" (is "?left = _")
- using psy psx unfolding setsum_right_distrib[symmetric]
- by (auto intro!: setsum_cong measure_setsum_split elim!: pos_simpleE)
- also have "... = (\<Sum>i\<in>s. (\<Sum>j \<in> s'. y j * measure M (a i \<inter> b j)))"
- proof (rule setsum_cong, simp, rule setsum_cong, simp_all)
- fix i j assume i: "i \<in> s" and j: "j \<in> s'"
- hence "a i \<in> sets M" using psx by (auto elim!: pos_simpleE)
-
- show "measure M (a i \<inter> b j) = 0 \<or> x i = y j"
- proof (cases "a i \<inter> b j = {}")
- case True thus ?thesis using empty_measure by simp
- next
- case False then obtain t where t: "t \<in> a i" "t \<in> b j" by auto
- hence "t \<in> space M" using `a i \<in> sets M` sets_into_space by auto
- with psx psy t i j have "x i = f t" and "y j = f t"
- unfolding pos_simple_def by auto
- thus ?thesis by simp
- qed
- qed
- also have "... = (\<Sum>j\<in>s'. (\<Sum>i\<in>s. y j * measure M (a i \<inter> b j)))"
- by (subst setsum_commute) simp
- also have "... = (\<Sum>i\<in>s'. y i * measure M (b i))" (is "?sum_sum = ?right")
- proof (rule setsum_cong)
- fix j assume "j \<in> s'"
- have "y j * measure M (b j) = (\<Sum>i\<in>s. y j * measure M (b j \<inter> a i))"
- using psx psy `j \<in> s'` unfolding setsum_right_distrib[symmetric]
- by (auto intro!: measure_setsum_split elim!: pos_simpleE)
- thus "(\<Sum>i\<in>s. y j * measure M (a i \<inter> b j)) = y j * measure M (b j)"
- by (auto intro!: setsum_cong arg_cong[where f="measure M"])
- qed simp
- finally show "?left = ?right" .
-qed
-
-lemma psfis_present:
- assumes "A \<in> psfis f"
- assumes "B \<in> psfis g"
- obtains z z' c k where
- "A = pos_simple_integral (k, c, z)"
- "B = pos_simple_integral (k, c, z')"
- "(k, c, z) \<in> pos_simple f"
- "(k, c, z') \<in> pos_simple g"
-using assms
-proof -
- from assms obtain s a x s' b y where
- ps: "(s, a, x) \<in> pos_simple f" "(s', b, y) \<in> pos_simple g" and
- A: "A = pos_simple_integral (s, a, x)" and
- B: "B = pos_simple_integral (s', b, y)"
- unfolding psfis_def pos_simple_integral_def by auto
-
- guess z z' c k using pos_simple_common_partition[OF ps] . note part = this
- show thesis
- proof (rule that[of k c z z', OF _ _ part])
- show "A = pos_simple_integral (k, c, z)"
- unfolding A by (rule pos_simple_integral_equal[OF ps(1) part(1)])
- show "B = pos_simple_integral (k, c, z')"
- unfolding B by (rule pos_simple_integral_equal[OF ps(2) part(2)])
- qed
-qed
-
-lemma pos_simple_integral_add:
- assumes "(s, a, x) \<in> pos_simple f"
- assumes "(s', b, y) \<in> pos_simple g"
- obtains s'' c z where
- "(s'', c, z) \<in> pos_simple (\<lambda>x. f x + g x)"
- "(pos_simple_integral (s, a, x) +
- pos_simple_integral (s', b, y) =
- pos_simple_integral (s'', c, z))"
-using assms
-proof -
- guess z z' c k
- by (rule pos_simple_common_partition[OF `(s, a, x) \<in> pos_simple f ` `(s', b, y) \<in> pos_simple g`])
- note kczz' = this
- have x: "pos_simple_integral (s, a, x) = pos_simple_integral (k, c, z)"
- by (rule pos_simple_integral_equal) (fact, fact)
- have y: "pos_simple_integral (s', b, y) = pos_simple_integral (k, c, z')"
- by (rule pos_simple_integral_equal) (fact, fact)
-
- have "pos_simple_integral (k, c, (\<lambda> x. z x + z' x))
- = (\<Sum> x \<in> k. (z x + z' x) * measure M (c x))"
- unfolding pos_simple_integral_def by auto
- also have "\<dots> = (\<Sum> x \<in> k. z x * measure M (c x) + z' x * measure M (c x))"
- by (simp add: left_distrib)
- also have "\<dots> = (\<Sum> x \<in> k. z x * measure M (c x)) + (\<Sum> x \<in> k. z' x * measure M (c x))" using setsum_addf by auto
- also have "\<dots> = pos_simple_integral (k, c, z) + pos_simple_integral (k, c, z')" unfolding pos_simple_integral_def by auto
- finally have ths: "pos_simple_integral (s, a, x) + pos_simple_integral (s', b, y) =
- pos_simple_integral (k, c, (\<lambda> x. z x + z' x))" using x y by auto
- show ?thesis
- apply (rule that[of k c "(\<lambda> x. z x + z' x)", OF _ ths])
- using kczz' unfolding pos_simple_def nonneg_def by (auto intro!: add_nonneg_nonneg)
-qed
-
-lemma psfis_add:
- assumes "a \<in> psfis f" "b \<in> psfis g"
- shows "a + b \<in> psfis (\<lambda>x. f x + g x)"
-using assms pos_simple_integral_add
-unfolding psfis_def by auto blast
-
-lemma pos_simple_integral_mono_on_mspace:
- assumes f: "(s, a, x) \<in> pos_simple f"
- assumes g: "(s', b, y) \<in> pos_simple g"
- assumes mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
- shows "pos_simple_integral (s, a, x) \<le> pos_simple_integral (s', b, y)"
-using assms
-proof -
- guess z z' c k by (rule pos_simple_common_partition[OF f g])
- note kczz' = this
- (* w = z and w' = z' except where c _ = {} or undef *)
- def w == "\<lambda> i. if i \<notin> k \<or> c i = {} then 0 else z i"
- def w' == "\<lambda> i. if i \<notin> k \<or> c i = {} then 0 else z' i"
- { fix i
- have "w i \<le> w' i"
- proof (cases "i \<notin> k \<or> c i = {}")
- case False hence "i \<in> k" "c i \<noteq> {}" by auto
- then obtain v where v: "v \<in> c i" and "c i \<in> sets M"
- using kczz'(1) unfolding pos_simple_def by blast
- hence "v \<in> space M" using sets_into_space by blast
- with mono[OF `v \<in> space M`] kczz' `i \<in> k` `v \<in> c i`
- have "z i \<le> z' i" unfolding pos_simple_def by simp
- thus ?thesis unfolding w_def w'_def using False by auto
- next
- case True thus ?thesis unfolding w_def w'_def by auto
- qed
- } note w_mn = this
-
- (* some technical stuff for the calculation*)
- have "\<And> i. i \<in> k \<Longrightarrow> c i \<in> sets M" using kczz' unfolding pos_simple_def by auto
- hence "\<And> i. i \<in> k \<Longrightarrow> measure M (c i) \<ge> 0" using positive by auto
- hence w_mono: "\<And> i. i \<in> k \<Longrightarrow> w i * measure M (c i) \<le> w' i * measure M (c i)"
- using mult_right_mono w_mn by blast
-
- { fix i have "\<lbrakk>i \<in> k ; z i \<noteq> w i\<rbrakk> \<Longrightarrow> measure M (c i) = 0"
- unfolding w_def by (cases "c i = {}") auto }
- hence zw: "\<And> i. i \<in> k \<Longrightarrow> z i * measure M (c i) = w i * measure M (c i)" by auto
-
- { fix i have "i \<in> k \<Longrightarrow> z i * measure M (c i) = w i * measure M (c i)"
- unfolding w_def by (cases "c i = {}") simp_all }
- note zw = this
-
- { fix i have "i \<in> k \<Longrightarrow> z' i * measure M (c i) = w' i * measure M (c i)"
- unfolding w'_def by (cases "c i = {}") simp_all }
- note z'w' = this
-
- (* the calculation *)
- have "pos_simple_integral (s, a, x) = pos_simple_integral (k, c, z)"
- using f kczz'(1) by (rule pos_simple_integral_equal)
- also have "\<dots> = (\<Sum> i \<in> k. z i * measure M (c i))"
- unfolding pos_simple_integral_def by auto
- also have "\<dots> = (\<Sum> i \<in> k. w i * measure M (c i))"
- using setsum_cong2[of k, OF zw] by auto
- also have "\<dots> \<le> (\<Sum> i \<in> k. w' i * measure M (c i))"
- using setsum_mono[OF w_mono, of k] by auto
- also have "\<dots> = (\<Sum> i \<in> k. z' i * measure M (c i))"
- using setsum_cong2[of k, OF z'w'] by auto
- also have "\<dots> = pos_simple_integral (k, c, z')"
- unfolding pos_simple_integral_def by auto
- also have "\<dots> = pos_simple_integral (s', b, y)"
- using kczz'(2) g by (rule pos_simple_integral_equal)
- finally show "pos_simple_integral (s, a, x) \<le> pos_simple_integral (s', b, y)"
- by simp
-qed
-
-lemma pos_simple_integral_mono:
- assumes a: "(s, a, x) \<in> pos_simple f" "(s', b, y) \<in> pos_simple g"
- assumes "\<And> z. f z \<le> g z"
- shows "pos_simple_integral (s, a, x) \<le> pos_simple_integral (s', b, y)"
-using assms pos_simple_integral_mono_on_mspace by auto
-
-lemma psfis_mono:
- assumes "a \<in> psfis f" "b \<in> psfis g"
- assumes "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
- shows "a \<le> b"
-using assms pos_simple_integral_mono_on_mspace unfolding psfis_def by auto
-
-lemma pos_simple_fn_integral_unique:
- assumes "(s, a, x) \<in> pos_simple f" "(s', b, y) \<in> pos_simple f"
- shows "pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)"
-using assms by (rule pos_simple_integral_equal) (* FIXME: redundant lemma *)
-
-lemma psfis_unique:
- assumes "a \<in> psfis f" "b \<in> psfis f"
- shows "a = b"
-using assms by (intro order_antisym psfis_mono [OF _ _ order_refl])
-
-lemma pos_simple_integral_indicator:
- assumes "A \<in> sets M"
- obtains s a x where
- "(s, a, x) \<in> pos_simple (indicator_fn A)"
- "measure M A = pos_simple_integral (s, a, x)"
-using assms
-proof -
- def s == "{0, 1} :: nat set"
- def a == "\<lambda> i :: nat. if i = 0 then A else space M - A"
- def x == "\<lambda> i :: nat. if i = 0 then 1 else (0 :: real)"
- have eq: "pos_simple_integral (s, a, x) = measure M A"
- unfolding s_def a_def x_def pos_simple_integral_def by auto
- have "(s, a, x) \<in> pos_simple (indicator_fn A)"
- unfolding pos_simple_def indicator_fn_def s_def a_def x_def nonneg_def
- using assms sets_into_space by auto
- from that[OF this] eq show thesis by auto
-qed
-
-lemma psfis_indicator:
- assumes "A \<in> sets M"
- shows "measure M A \<in> psfis (indicator_fn A)"
-using pos_simple_integral_indicator[OF assms]
- unfolding psfis_def image_def by auto
-
-lemma pos_simple_integral_mult:
- assumes f: "(s, a, x) \<in> pos_simple f"
- assumes "0 \<le> z"
- obtains s' b y where
- "(s', b, y) \<in> pos_simple (\<lambda>x. z * f x)"
- "pos_simple_integral (s', b, y) = z * pos_simple_integral (s, a, x)"
- using assms that[of s a "\<lambda>i. z * x i"]
- by (simp add: pos_simple_def pos_simple_integral_def setsum_right_distrib algebra_simps nonneg_def mult_nonneg_nonneg)
-
-lemma psfis_mult:
- assumes "r \<in> psfis f"
- assumes "0 \<le> z"
- shows "z * r \<in> psfis (\<lambda>x. z * f x)"
-using assms
-proof -
- from assms obtain s a x
- where sax: "(s, a, x) \<in> pos_simple f"
- and r: "r = pos_simple_integral (s, a, x)"
- unfolding psfis_def image_def by auto
- obtain s' b y where
- "(s', b, y) \<in> pos_simple (\<lambda>x. z * f x)"
- "z * pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)"
- using pos_simple_integral_mult[OF sax `0 \<le> z`, of thesis] by auto
- thus ?thesis using r unfolding psfis_def image_def by auto
-qed
-
-lemma psfis_setsum_image:
- assumes "finite P"
- assumes "\<And>i. i \<in> P \<Longrightarrow> a i \<in> psfis (f i)"
- shows "(setsum a P) \<in> psfis (\<lambda>t. \<Sum>i \<in> P. f i t)"
-using assms
-proof (induct P)
- case empty
- let ?s = "{0 :: nat}"
- let ?a = "\<lambda> i. if i = (0 :: nat) then space M else {}"
- let ?x = "\<lambda> (i :: nat). (0 :: real)"
- have "(?s, ?a, ?x) \<in> pos_simple (\<lambda> t. (0 :: real))"
- unfolding pos_simple_def image_def nonneg_def by auto
- moreover have "(\<Sum> i \<in> ?s. ?x i * measure M (?a i)) = 0" by auto
- ultimately have "0 \<in> psfis (\<lambda> t. 0)"
- unfolding psfis_def image_def pos_simple_integral_def nonneg_def
- by (auto intro!:bexI[of _ "(?s, ?a, ?x)"])
- thus ?case by auto
-next
- case (insert x P) note asms = this
- have "finite P" by fact
- have "x \<notin> P" by fact
- have "(\<And>i. i \<in> P \<Longrightarrow> a i \<in> psfis (f i)) \<Longrightarrow>
- setsum a P \<in> psfis (\<lambda>t. \<Sum>i\<in>P. f i t)" by fact
- have "setsum a (insert x P) = a x + setsum a P" using `finite P` `x \<notin> P` by auto
- also have "\<dots> \<in> psfis (\<lambda> t. f x t + (\<Sum> i \<in> P. f i t))"
- using asms psfis_add by auto
- also have "\<dots> = psfis (\<lambda> t. \<Sum> i \<in> insert x P. f i t)"
- using `x \<notin> P` `finite P` by auto
- finally show ?case by simp
-qed
-
-lemma psfis_intro:
- assumes "a ` P \<subseteq> sets M" and "nonneg x" and "finite P"
- shows "(\<Sum>i\<in>P. x i * measure M (a i)) \<in> psfis (\<lambda>t. \<Sum>i\<in>P. x i * indicator_fn (a i) t)"
-using assms psfis_mult psfis_indicator
-unfolding image_def nonneg_def
-by (auto intro!:psfis_setsum_image)
-
-lemma psfis_nonneg: "a \<in> psfis f \<Longrightarrow> nonneg f"
-unfolding psfis_def pos_simple_def by auto
-
-lemma pos_psfis: "r \<in> psfis f \<Longrightarrow> 0 \<le> r"
-unfolding psfis_def pos_simple_integral_def image_def pos_simple_def nonneg_def
-using positive[unfolded positive_def] by (auto intro!:setsum_nonneg mult_nonneg_nonneg)
-
-lemma psfis_borel_measurable:
- assumes "a \<in> psfis f"
- shows "f \<in> borel_measurable M"
-using assms
-proof -
- from assms obtain s a' x where sa'x: "(s, a', x) \<in> pos_simple f" and sa'xa: "pos_simple_integral (s, a', x) = a"
- and fs: "finite s"
- unfolding psfis_def pos_simple_integral_def image_def
- by (auto elim:pos_simpleE)
- { fix w assume "w \<in> space M"
- hence "(f w \<le> a) = ((\<Sum> i \<in> s. x i * indicator_fn (a' i) w) \<le> a)"
- using pos_simple_setsum_indicator_fn[OF sa'x, of w] by simp
- } hence "\<And> w. (w \<in> space M \<and> f w \<le> a) = (w \<in> space M \<and> (\<Sum> i \<in> s. x i * indicator_fn (a' i) w) \<le> a)"
- by auto
- { fix i assume "i \<in> s"
- hence "indicator_fn (a' i) \<in> borel_measurable M"
- using borel_measurable_indicator using sa'x[unfolded pos_simple_def] by auto
- hence "(\<lambda> w. x i * indicator_fn (a' i) w) \<in> borel_measurable M"
- using affine_borel_measurable[of "\<lambda> w. indicator_fn (a' i) w" 0 "x i"]
- by (simp add: mult_commute) }
- from borel_measurable_setsum_borel_measurable[OF fs this] affine_borel_measurable
- have "(\<lambda> w. (\<Sum> i \<in> s. x i * indicator_fn (a' i) w)) \<in> borel_measurable M" by auto
- from borel_measurable_cong[OF pos_simple_setsum_indicator_fn[OF sa'x]] this
- show ?thesis by simp
-qed
-
-lemma psfis_mono_conv_mono:
- assumes mono: "mono_convergent u f (space M)"
- and ps_u: "\<And>n. x n \<in> psfis (u n)"
- and "x ----> y"
- and "r \<in> psfis s"
- and f_upper_bound: "\<And>t. t \<in> space M \<Longrightarrow> s t \<le> f t"
- shows "r <= y"
-proof (rule field_le_mult_one_interval)
- fix z :: real assume "0 < z" and "z < 1"
- hence "0 \<le> z" by auto
- let "?B' n" = "{w \<in> space M. z * s w <= u n w}"
-
- have "incseq x" unfolding incseq_def
- proof safe
- fix m n :: nat assume "m \<le> n"
- show "x m \<le> x n"
- proof (rule psfis_mono[OF `x m \<in> psfis (u m)` `x n \<in> psfis (u n)`])
- fix t assume "t \<in> space M"
- with mono_convergentD[OF mono this] `m \<le> n` show "u m t \<le> u n t"
- unfolding incseq_def by auto
- qed
- qed
-
- from `r \<in> psfis s`
- obtain s' a x' where r: "r = pos_simple_integral (s', a, x')"
- and ps_s: "(s', a, x') \<in> pos_simple s"
- unfolding psfis_def by auto
-
- { fix t i assume "i \<in> s'" "t \<in> a i"
- hence "t \<in> space M" using ps_s by (auto elim!: pos_simpleE) }
- note t_in_space = this
-
- { fix n
- from psfis_borel_measurable[OF `r \<in> psfis s`]
- psfis_borel_measurable[OF ps_u]
- have "?B' n \<in> sets M"
- by (auto intro!:
- borel_measurable_leq_borel_measurable
- borel_measurable_times_borel_measurable
- borel_measurable_const) }
- note B'_in_M = this
-
- { fix n have "(\<lambda>i. a i \<inter> ?B' n) ` s' \<subseteq> sets M" using B'_in_M ps_s
- by (auto intro!: Int elim!: pos_simpleE) }
- note B'_inter_a_in_M = this
-
- let "?sum n" = "(\<Sum>i\<in>s'. x' i * measure M (a i \<inter> ?B' n))"
- { fix n
- have "z * ?sum n \<le> x n"
- proof (rule psfis_mono[OF _ ps_u])
- have *: "\<And>i t. indicator_fn (?B' n) t * (x' i * indicator_fn (a i) t) =
- x' i * indicator_fn (a i \<inter> ?B' n) t" unfolding indicator_fn_def by auto
- have ps': "?sum n \<in> psfis (\<lambda>t. indicator_fn (?B' n) t * (\<Sum>i\<in>s'. x' i * indicator_fn (a i) t))"
- unfolding setsum_right_distrib * using B'_in_M ps_s
- by (auto intro!: psfis_intro Int elim!: pos_simpleE)
- also have "... = psfis (\<lambda>t. indicator_fn (?B' n) t * s t)" (is "psfis ?l = psfis ?r")
- proof (rule psfis_cong)
- show "nonneg ?l" using psfis_nonneg[OF ps'] .
- show "nonneg ?r" using psfis_nonneg[OF `r \<in> psfis s`] unfolding nonneg_def indicator_fn_def by auto
- fix t assume "t \<in> space M"
- show "?l t = ?r t" unfolding pos_simple_setsum_indicator_fn[OF ps_s `t \<in> space M`] ..
- qed
- finally show "z * ?sum n \<in> psfis (\<lambda>t. z * ?r t)" using psfis_mult[OF _ `0 \<le> z`] by simp
- next
- fix t assume "t \<in> space M"
- show "z * (indicator_fn (?B' n) t * s t) \<le> u n t"
- using psfis_nonneg[OF ps_u] unfolding nonneg_def indicator_fn_def by auto
- qed }
- hence *: "\<exists>N. \<forall>n\<ge>N. z * ?sum n \<le> x n" by (auto intro!: exI[of _ 0])
-
- show "z * r \<le> y" unfolding r pos_simple_integral_def
- proof (rule LIMSEQ_le[OF mult_right.LIMSEQ `x ----> y` *],
- simp, rule LIMSEQ_setsum, rule mult_right.LIMSEQ)
- fix i assume "i \<in> s'"
- from psfis_nonneg[OF `r \<in> psfis s`, unfolded nonneg_def]
- have "\<And>t. 0 \<le> s t" by simp
-
- have *: "(\<Union>j. a i \<inter> ?B' j) = a i"
- proof (safe, simp, safe)
- fix t assume "t \<in> a i"
- thus "t \<in> space M" using t_in_space[OF `i \<in> s'`] by auto
- show "\<exists>j. z * s t \<le> u j t"
- proof (cases "s t = 0")
- case True thus ?thesis
- using psfis_nonneg[OF ps_u] unfolding nonneg_def by auto
- next
- case False with `0 \<le> s t`
- have "0 < s t" by auto
- hence "z * s t < 1 * s t" using `0 < z` `z < 1`
- by (auto intro!: mult_strict_right_mono simp del: mult_1_left)
- also have "... \<le> f t" using f_upper_bound `t \<in> space M` by auto
- finally obtain b where "\<And>j. b \<le> j \<Longrightarrow> z * s t < u j t" using `t \<in> space M`
- using mono_conv_outgrow[of "\<lambda>n. u n t" "f t" "z * s t"]
- using mono_convergentD[OF mono] by auto
- from this[of b] show ?thesis by (auto intro!: exI[of _ b])
- qed
- qed
-
- show "(\<lambda>n. measure M (a i \<inter> ?B' n)) ----> measure M (a i)"
- proof (safe intro!:
- monotone_convergence[of "\<lambda>n. a i \<inter> ?B' n", unfolded comp_def *])
- fix n show "a i \<inter> ?B' n \<in> sets M"
- using B'_inter_a_in_M[of n] `i \<in> s'` by auto
- next
- fix j t assume "t \<in> space M" and "z * s t \<le> u j t"
- thus "z * s t \<le> u (Suc j) t"
- using mono_convergentD(1)[OF mono, unfolded incseq_def,
- rule_format, of t j "Suc j"]
- by auto
- qed
- qed
-qed
-
-section "Continuous posititve integration"
-
-definition
- "nnfis f = { y. \<exists>u x. mono_convergent u f (space M) \<and>
- (\<forall>n. x n \<in> psfis (u n)) \<and> x ----> y }"
-
-lemma psfis_nnfis:
- "a \<in> psfis f \<Longrightarrow> a \<in> nnfis f"
- unfolding nnfis_def mono_convergent_def incseq_def
- by (auto intro!: exI[of _ "\<lambda>n. f"] exI[of _ "\<lambda>n. a"] LIMSEQ_const)
-
-lemma nnfis_0: "0 \<in> nnfis (\<lambda> x. 0)"
- by (rule psfis_nnfis[OF psfis_0])
-
-lemma nnfis_times:
- assumes "a \<in> nnfis f" and "0 \<le> z"
- shows "z * a \<in> nnfis (\<lambda>t. z * f t)"
-proof -
- obtain u x where "mono_convergent u f (space M)" and
- "\<And>n. x n \<in> psfis (u n)" "x ----> a"
- using `a \<in> nnfis f` unfolding nnfis_def by auto
- with `0 \<le> z`show ?thesis unfolding nnfis_def mono_convergent_def incseq_def
- by (auto intro!: exI[of _ "\<lambda>n w. z * u n w"] exI[of _ "\<lambda>n. z * x n"]
- LIMSEQ_mult LIMSEQ_const psfis_mult mult_left_mono)
-qed
-
-lemma nnfis_add:
- assumes "a \<in> nnfis f" and "b \<in> nnfis g"
- shows "a + b \<in> nnfis (\<lambda>t. f t + g t)"
-proof -
- obtain u x w y
- where "mono_convergent u f (space M)" and
- "\<And>n. x n \<in> psfis (u n)" "x ----> a" and
- "mono_convergent w g (space M)" and
- "\<And>n. y n \<in> psfis (w n)" "y ----> b"
- using `a \<in> nnfis f` `b \<in> nnfis g` unfolding nnfis_def by auto
- thus ?thesis unfolding nnfis_def mono_convergent_def incseq_def
- by (auto intro!: exI[of _ "\<lambda>n a. u n a + w n a"] exI[of _ "\<lambda>n. x n + y n"]
- LIMSEQ_add LIMSEQ_const psfis_add add_mono)
-qed
-
-lemma nnfis_mono:
- assumes nnfis: "a \<in> nnfis f" "b \<in> nnfis g"
- and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
- shows "a \<le> b"
-proof -
- obtain u x w y where
- mc: "mono_convergent u f (space M)" "mono_convergent w g (space M)" and
- psfis: "\<And>n. x n \<in> psfis (u n)" "\<And>n. y n \<in> psfis (w n)" and
- limseq: "x ----> a" "y ----> b" using nnfis unfolding nnfis_def by auto
- show ?thesis
- proof (rule LIMSEQ_le_const2[OF limseq(1)], rule exI[of _ 0], safe)
- fix n
- show "x n \<le> b"
- proof (rule psfis_mono_conv_mono[OF mc(2) psfis(2) limseq(2) psfis(1)])
- fix t assume "t \<in> space M"
- from mono_convergent_le[OF mc(1) this] mono[OF this]
- show "u n t \<le> g t" by (rule order_trans)
- qed
- qed
-qed
-
-lemma nnfis_unique:
- assumes a: "a \<in> nnfis f" and b: "b \<in> nnfis f"
- shows "a = b"
- using nnfis_mono[OF a b] nnfis_mono[OF b a]
- by (auto intro!: order_antisym[of a b])
-
-lemma psfis_equiv:
- assumes "a \<in> psfis f" and "nonneg g"
- and "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
- shows "a \<in> psfis g"
- using assms unfolding psfis_def pos_simple_def by auto
-
-lemma psfis_mon_upclose:
- assumes "\<And>m. a m \<in> psfis (u m)"
- shows "\<exists>c. c \<in> psfis (mon_upclose n u)"
-proof (induct n)
- case 0 thus ?case unfolding mon_upclose.simps using assms ..
-next
- case (Suc n)
- then obtain sn an xn where ps: "(sn, an, xn) \<in> pos_simple (mon_upclose n u)"
- unfolding psfis_def by auto
- obtain ss as xs where ps': "(ss, as, xs) \<in> pos_simple (u (Suc n))"
- using assms[of "Suc n"] unfolding psfis_def by auto
- from pos_simple_common_partition[OF ps ps'] guess x x' a s .
- hence "(s, a, upclose x x') \<in> pos_simple (mon_upclose (Suc n) u)"
- by (simp add: upclose_def pos_simple_def nonneg_def max_def)
- thus ?case unfolding psfis_def by auto
-qed
-
-text {* Beppo-Levi monotone convergence theorem *}
-lemma nnfis_mon_conv:
- assumes mc: "mono_convergent f h (space M)"
- and nnfis: "\<And>n. x n \<in> nnfis (f n)"
- and "x ----> z"
- shows "z \<in> nnfis h"
-proof -
- have "\<forall>n. \<exists>u y. mono_convergent u (f n) (space M) \<and> (\<forall>m. y m \<in> psfis (u m)) \<and>
- y ----> x n"
- using nnfis unfolding nnfis_def by auto
- from choice[OF this] guess u ..
- from choice[OF this] guess y ..
- hence mc_u: "\<And>n. mono_convergent (u n) (f n) (space M)"
- and psfis: "\<And>n m. y n m \<in> psfis (u n m)" and "\<And>n. y n ----> x n"
- by auto
-
- let "?upclose n" = "mon_upclose n (\<lambda>m. u m n)"
-
- have "\<exists>c. \<forall>n. c n \<in> psfis (?upclose n)"
- by (safe intro!: choice psfis_mon_upclose) (rule psfis)
- then guess c .. note c = this[rule_format]
-
- show ?thesis unfolding nnfis_def
- proof (safe intro!: exI)
- show mc_upclose: "mono_convergent ?upclose h (space M)"
- by (rule mon_upclose_mono_convergent[OF mc_u mc])
- show psfis_upclose: "\<And>n. c n \<in> psfis (?upclose n)"
- using c .
-
- { fix n m :: nat assume "n \<le> m"
- hence "c n \<le> c m"
- using psfis_mono[OF c c]
- using mono_convergentD(1)[OF mc_upclose, unfolded incseq_def]
- by auto }
- hence "incseq c" unfolding incseq_def by auto
-
- { fix n
- have c_nnfis: "c n \<in> nnfis (?upclose n)" using c by (rule psfis_nnfis)
- from nnfis_mono[OF c_nnfis nnfis]
- mon_upclose_le_mono_convergent[OF mc_u]
- mono_convergentD(1)[OF mc]
- have "c n \<le> x n" by fast }
- note c_less_x = this
-
- { fix n
- note c_less_x[of n]
- also have "x n \<le> z"
- proof (rule incseq_le)
- show "x ----> z" by fact
- from mono_convergentD(1)[OF mc]
- show "incseq x" unfolding incseq_def
- by (auto intro!: nnfis_mono[OF nnfis nnfis])
- qed
- finally have "c n \<le> z" . }
- note c_less_z = this
-
- have "convergent c"
- proof (rule Bseq_mono_convergent[unfolded incseq_def[symmetric]])
- show "Bseq c"
- using pos_psfis[OF c] c_less_z
- by (auto intro!: BseqI'[of _ z])
- show "incseq c" by fact
- qed
- then obtain l where l: "c ----> l" unfolding convergent_def by auto
-
- have "l \<le> z" using c_less_x l
- by (auto intro!: LIMSEQ_le[OF _ `x ----> z`])
- moreover have "z \<le> l"
- proof (rule LIMSEQ_le_const2[OF `x ----> z`], safe intro!: exI[of _ 0])
- fix n
- have "l \<in> nnfis h"
- unfolding nnfis_def using l mc_upclose psfis_upclose by auto
- from nnfis this mono_convergent_le[OF mc]
- show "x n \<le> l" by (rule nnfis_mono)
- qed
- ultimately have "l = z" by (rule order_antisym)
- thus "c ----> z" using `c ----> l` by simp
- qed
-qed
-
-lemma nnfis_pos_on_mspace:
- assumes "a \<in> nnfis f" and "x \<in>space M"
- shows "0 \<le> f x"
-using assms
-proof -
- from assms[unfolded nnfis_def] guess u y by auto note uy = this
- hence "\<And> n. 0 \<le> u n x"
- unfolding nnfis_def psfis_def pos_simple_def nonneg_def mono_convergent_def
- by auto
- thus "0 \<le> f x" using uy[rule_format]
- unfolding nnfis_def psfis_def pos_simple_def nonneg_def mono_convergent_def
- using incseq_le[of "\<lambda> n. u n x" "f x"] order_trans
- by fast
-qed
-
-lemma nnfis_borel_measurable:
- assumes "a \<in> nnfis f"
- shows "f \<in> borel_measurable M"
-using assms unfolding nnfis_def
-apply auto
-apply (rule mono_convergent_borel_measurable)
-using psfis_borel_measurable
-by auto
-
-lemma borel_measurable_mon_conv_psfis:
- assumes f_borel: "f \<in> borel_measurable M"
- and nonneg: "\<And>t. t \<in> space M \<Longrightarrow> 0 \<le> f t"
- shows"\<exists>u x. mono_convergent u f (space M) \<and> (\<forall>n. x n \<in> psfis (u n))"
-proof (safe intro!: exI)
- let "?I n" = "{0<..<n * 2^n}"
- let "?A n i" = "{w \<in> space M. real (i :: nat) / 2^(n::nat) \<le> f w \<and> f w < real (i + 1) / 2^n}"
- let "?u n t" = "\<Sum>i\<in>?I n. real i / 2^n * indicator_fn (?A n i) t"
- let "?x n" = "\<Sum>i\<in>?I n. real i / 2^n * measure M (?A n i)"
-
- let "?w n t" = "if f t < real n then real (natfloor (f t * 2^n)) / 2^n else 0"
-
- { fix t n assume t: "t \<in> space M"
- have "?u n t = ?w n t" (is "_ = (if _ then real ?i / _ else _)")
- proof (cases "f t < real n")
- case True
- with nonneg t
- have i: "?i < n * 2^n"
- by (auto simp: real_of_nat_power[symmetric]
- intro!: less_natfloor mult_nonneg_nonneg)
-
- hence t_in_A: "t \<in> ?A n ?i"
- unfolding divide_le_eq less_divide_eq
- using nonneg t True
- by (auto intro!: real_natfloor_le
- real_natfloor_gt_diff_one[unfolded diff_less_eq]
- simp: real_of_nat_Suc zero_le_mult_iff)
-
- hence *: "real ?i / 2^n \<le> f t"
- "f t < real (?i + 1) / 2^n" by auto
- { fix j assume "t \<in> ?A n j"
- hence "real j / 2^n \<le> f t"
- and "f t < real (j + 1) / 2^n" by auto
- with * have "j \<in> {?i}" unfolding divide_le_eq less_divide_eq
- by auto }
- hence *: "\<And>j. t \<in> ?A n j \<longleftrightarrow> j \<in> {?i}" using t_in_A by auto
-
- have "?u n t = real ?i / 2^n"
- unfolding indicator_fn_def if_distrib *
- setsum_cases[OF finite_greaterThanLessThan]
- using i by (cases "?i = 0") simp_all
- thus ?thesis using True by auto
- next
- case False
- have "?u n t = (\<Sum>i \<in> {0 <..< n*2^n}. 0)"
- proof (rule setsum_cong)
- fix i assume "i \<in> {0 <..< n*2^n}"
- hence "i + 1 \<le> n * 2^n" by simp
- hence "real (i + 1) \<le> real n * 2^n"
- unfolding real_of_nat_le_iff[symmetric]
- by (auto simp: real_of_nat_power[symmetric])
- also have "... \<le> f t * 2^n"
- using False by (auto intro!: mult_nonneg_nonneg)
- finally have "t \<notin> ?A n i"
- by (auto simp: divide_le_eq less_divide_eq)
- thus "real i / 2^n * indicator_fn (?A n i) t = 0"
- unfolding indicator_fn_def by auto
- qed simp
- thus ?thesis using False by auto
- qed }
- note u_at_t = this
-
- show "mono_convergent ?u f (space M)" unfolding mono_convergent_def
- proof safe
- fix t assume t: "t \<in> space M"
- { fix m n :: nat assume "m \<le> n"
- hence *: "(2::real)^n = 2^m * 2^(n - m)" unfolding power_add[symmetric] by auto
- have "real (natfloor (f t * 2^m) * natfloor (2^(n-m))) \<le> real (natfloor (f t * 2 ^ n))"
- apply (subst *)
- apply (subst mult_assoc[symmetric])
- apply (subst real_of_nat_le_iff)
- apply (rule le_mult_natfloor)
- using nonneg[OF t] by (auto intro!: mult_nonneg_nonneg)
- hence "real (natfloor (f t * 2^m)) * 2^n \<le> real (natfloor (f t * 2^n)) * 2^m"
- apply (subst *)
- apply (subst (3) mult_commute)
- apply (subst mult_assoc[symmetric])
- by (auto intro: mult_right_mono simp: natfloor_power real_of_nat_power[symmetric]) }
- thus "incseq (\<lambda>n. ?u n t)" unfolding u_at_t[OF t] unfolding incseq_def
- by (auto simp add: le_divide_eq divide_le_eq less_divide_eq)
-
- show "(\<lambda>i. ?u i t) ----> f t" unfolding u_at_t[OF t]
- proof (rule LIMSEQ_I, safe intro!: exI)
- fix r :: real and n :: nat
- let ?N = "natfloor (1/r) + 1"
- assume "0 < r" and N: "max ?N (natfloor (f t) + 1) \<le> n"
- hence "?N \<le> n" by auto
-
- have "1 / r < real (natfloor (1/r) + 1)" using real_natfloor_add_one_gt
- by (simp add: real_of_nat_Suc)
- also have "... < 2^?N" by (rule two_realpow_gt)
- finally have less_r: "1 / 2^?N < r" using `0 < r`
- by (auto simp: less_divide_eq divide_less_eq algebra_simps)
-
- have "f t < real (natfloor (f t) + 1)" using real_natfloor_add_one_gt[of "f t"] by auto
- also have "... \<le> real n" unfolding real_of_nat_le_iff using N by auto
- finally have "f t < real n" .
-
- have "real (natfloor (f t * 2^n)) \<le> f t * 2^n"
- using nonneg[OF t] by (auto intro!: real_natfloor_le mult_nonneg_nonneg)
- hence less: "real (natfloor (f t * 2^n)) / 2^n \<le> f t" unfolding divide_le_eq by auto
-
- have "f t * 2 ^ n - 1 < real (natfloor (f t * 2^n))" using real_natfloor_gt_diff_one .
- hence "f t - real (natfloor (f t * 2^n)) / 2^n < 1 / 2^n"
- by (auto simp: less_divide_eq divide_less_eq algebra_simps)
- also have "... \<le> 1 / 2^?N" using `?N \<le> n`
- by (auto intro!: divide_left_mono mult_pos_pos simp del: power_Suc)
- also have "... < r" using less_r .
- finally show "norm (?w n t - f t) < r" using `f t < real n` less by auto
- qed
- qed
-
- fix n
- show "?x n \<in> psfis (?u n)"
- proof (rule psfis_intro)
- show "?A n ` ?I n \<subseteq> sets M"
- proof safe
- fix i :: nat
- from Int[OF
- f_borel[unfolded borel_measurable_less_iff, rule_format, of "real (i+1) / 2^n"]
- f_borel[unfolded borel_measurable_ge_iff, rule_format, of "real i / 2^n"]]
- show "?A n i \<in> sets M"
- by (metis Collect_conj_eq Int_commute Int_left_absorb Int_left_commute)
- qed
- show "nonneg (\<lambda>i :: nat. real i / 2 ^ n)"
- unfolding nonneg_def by (auto intro!: divide_nonneg_pos)
- qed auto
-qed
-
-lemma nnfis_dom_conv:
- assumes borel: "f \<in> borel_measurable M"
- and nnfis: "b \<in> nnfis g"
- and ord: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
- and nonneg: "\<And>t. t \<in> space M \<Longrightarrow> 0 \<le> f t"
- shows "\<exists>a. a \<in> nnfis f \<and> a \<le> b"
-proof -
- obtain u x where mc_f: "mono_convergent u f (space M)" and
- psfis: "\<And>n. x n \<in> psfis (u n)"
- using borel_measurable_mon_conv_psfis[OF borel nonneg] by auto
-
- { fix n
- { fix t assume t: "t \<in> space M"
- note mono_convergent_le[OF mc_f this, of n]
- also note ord[OF t]
- finally have "u n t \<le> g t" . }
- from nnfis_mono[OF psfis_nnfis[OF psfis] nnfis this]
- have "x n \<le> b" by simp }
- note x_less_b = this
-
- have "convergent x"
- proof (safe intro!: Bseq_mono_convergent)
- from x_less_b pos_psfis[OF psfis]
- show "Bseq x" by (auto intro!: BseqI'[of _ b])
-
- fix n m :: nat assume *: "n \<le> m"
- show "x n \<le> x m"
- proof (rule psfis_mono[OF `x n \<in> psfis (u n)` `x m \<in> psfis (u m)`])
- fix t assume "t \<in> space M"
- from mc_f[THEN mono_convergentD(1), unfolded incseq_def, OF this]
- show "u n t \<le> u m t" using * by auto
- qed
- qed
- then obtain a where "x ----> a" unfolding convergent_def by auto
- with LIMSEQ_le_const2[OF `x ----> a`] x_less_b mc_f psfis
- show ?thesis unfolding nnfis_def by auto
-qed
-
-lemma the_nnfis[simp]: "a \<in> nnfis f \<Longrightarrow> (THE a. a \<in> nnfis f) = a"
- by (auto intro: the_equality nnfis_unique)
-
-lemma nnfis_cong:
- assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
- shows "nnfis f = nnfis g"
-proof -
- { fix f g :: "'a \<Rightarrow> real" assume cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
- fix x assume "x \<in> nnfis f"
- then guess u y unfolding nnfis_def by safe note x = this
- hence "mono_convergent u g (space M)"
- unfolding mono_convergent_def using cong by auto
- with x(2,3) have "x \<in> nnfis g" unfolding nnfis_def by auto }
- from this[OF cong] this[OF cong[symmetric]]
- show ?thesis by safe
-qed
-
-section "Lebesgue Integral"
-
-definition
- "integrable f \<longleftrightarrow> (\<exists>x. x \<in> nnfis (pos_part f)) \<and> (\<exists>y. y \<in> nnfis (neg_part f))"
-
-definition
- "integral f = (THE i :: real. i \<in> nnfis (pos_part f)) - (THE j. j \<in> nnfis (neg_part f))"
-
-lemma integral_cong:
- assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
- shows "integral f = integral g"
-proof -
- have "nnfis (pos_part f) = nnfis (pos_part g)"
- using cong by (auto simp: pos_part_def intro!: nnfis_cong)
- moreover
- have "nnfis (neg_part f) = nnfis (neg_part g)"
- using cong by (auto simp: neg_part_def intro!: nnfis_cong)
- ultimately show ?thesis
- unfolding integral_def by auto
-qed
-
-lemma nnfis_integral:
- assumes "a \<in> nnfis f"
- shows "integrable f" and "integral f = a"
-proof -
- have a: "a \<in> nnfis (pos_part f)"
- using assms nnfis_pos_on_mspace[OF assms]
- by (auto intro!: nnfis_mon_conv[of "\<lambda>i. f" _ "\<lambda>i. a"]
- LIMSEQ_const simp: mono_convergent_def pos_part_def incseq_def max_def)
-
- have "\<And>t. t \<in> space M \<Longrightarrow> neg_part f t = 0"
- unfolding neg_part_def min_def
- using nnfis_pos_on_mspace[OF assms] by auto
- hence 0: "0 \<in> nnfis (neg_part f)"
- by (auto simp: nnfis_def mono_convergent_def psfis_0 incseq_def
- intro!: LIMSEQ_const exI[of _ "\<lambda> x n. 0"] exI[of _ "\<lambda> n. 0"])
-
- from 0 a show "integrable f" "integral f = a"
- unfolding integrable_def integral_def by auto
-qed
-
-lemma nnfis_minus_nnfis_integral:
- assumes "a \<in> nnfis f" and "b \<in> nnfis g"
- shows "integrable (\<lambda>t. f t - g t)" and "integral (\<lambda>t. f t - g t) = a - b"
-proof -
- have borel: "(\<lambda>t. f t - g t) \<in> borel_measurable M" using assms
- by (blast intro:
- borel_measurable_diff_borel_measurable nnfis_borel_measurable)
-
- have "\<exists>x. x \<in> nnfis (pos_part (\<lambda>t. f t - g t)) \<and> x \<le> a + b"
- (is "\<exists>x. x \<in> nnfis ?pp \<and> _")
- proof (rule nnfis_dom_conv)
- show "?pp \<in> borel_measurable M"
- using borel by (rule pos_part_borel_measurable neg_part_borel_measurable)
- show "a + b \<in> nnfis (\<lambda>t. f t + g t)" using assms by (rule nnfis_add)
- fix t assume "t \<in> space M"
- with assms nnfis_add assms[THEN nnfis_pos_on_mspace[OF _ this]]
- show "?pp t \<le> f t + g t" unfolding pos_part_def by auto
- show "0 \<le> ?pp t" using nonneg_pos_part[of "\<lambda>t. f t - g t"]
- unfolding nonneg_def by auto
- qed
- then obtain x where x: "x \<in> nnfis ?pp" by auto
- moreover
- have "\<exists>x. x \<in> nnfis (neg_part (\<lambda>t. f t - g t)) \<and> x \<le> a + b"
- (is "\<exists>x. x \<in> nnfis ?np \<and> _")
- proof (rule nnfis_dom_conv)
- show "?np \<in> borel_measurable M"
- using borel by (rule pos_part_borel_measurable neg_part_borel_measurable)
- show "a + b \<in> nnfis (\<lambda>t. f t + g t)" using assms by (rule nnfis_add)
- fix t assume "t \<in> space M"
- with assms nnfis_add assms[THEN nnfis_pos_on_mspace[OF _ this]]
- show "?np t \<le> f t + g t" unfolding neg_part_def by auto
- show "0 \<le> ?np t" using nonneg_neg_part[of "\<lambda>t. f t - g t"]
- unfolding nonneg_def by auto
- qed
- then obtain y where y: "y \<in> nnfis ?np" by auto
- ultimately show "integrable (\<lambda>t. f t - g t)"
- unfolding integrable_def by auto
-
- from x and y
- have "a + y \<in> nnfis (\<lambda>t. f t + ?np t)"
- and "b + x \<in> nnfis (\<lambda>t. g t + ?pp t)" using assms by (auto intro: nnfis_add)
- moreover
- have "\<And>t. f t + ?np t = g t + ?pp t"
- unfolding pos_part_def neg_part_def by auto
- ultimately have "a - b = x - y"
- using nnfis_unique by (auto simp: algebra_simps)
- thus "integral (\<lambda>t. f t - g t) = a - b"
- unfolding integral_def
- using the_nnfis[OF x] the_nnfis[OF y] by simp
-qed
-
-lemma integral_borel_measurable:
- "integrable f \<Longrightarrow> f \<in> borel_measurable M"
- unfolding integrable_def
- by (subst pos_part_neg_part_borel_measurable_iff)
- (auto intro: nnfis_borel_measurable)
-
-lemma integral_indicator_fn:
- assumes "a \<in> sets M"
- shows "integral (indicator_fn a) = measure M a"
- and "integrable (indicator_fn a)"
- using psfis_indicator[OF assms, THEN psfis_nnfis]
- by (auto intro!: nnfis_integral)
-
-lemma integral_add:
- assumes "integrable f" and "integrable g"
- shows "integrable (\<lambda>t. f t + g t)"
- and "integral (\<lambda>t. f t + g t) = integral f + integral g"
-proof -
- { fix t
- have "pos_part f t + pos_part g t - (neg_part f t + neg_part g t) =
- f t + g t"
- unfolding pos_part_def neg_part_def by auto }
- note part_sum = this
-
- from assms obtain a b c d where
- a: "a \<in> nnfis (pos_part f)" and b: "b \<in> nnfis (neg_part f)" and
- c: "c \<in> nnfis (pos_part g)" and d: "d \<in> nnfis (neg_part g)"
- unfolding integrable_def by auto
- note sums = nnfis_add[OF a c] nnfis_add[OF b d]
- note int = nnfis_minus_nnfis_integral[OF sums, unfolded part_sum]
-
- show "integrable (\<lambda>t. f t + g t)" using int(1) .
-
- show "integral (\<lambda>t. f t + g t) = integral f + integral g"
- using int(2) sums a b c d by (simp add: the_nnfis integral_def)
-qed
-
-lemma integral_mono:
- assumes "integrable f" and "integrable g"
- and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
- shows "integral f \<le> integral g"
-proof -
- from assms obtain a b c d where
- a: "a \<in> nnfis (pos_part f)" and b: "b \<in> nnfis (neg_part f)" and
- c: "c \<in> nnfis (pos_part g)" and d: "d \<in> nnfis (neg_part g)"
- unfolding integrable_def by auto
-
- have "a \<le> c"
- proof (rule nnfis_mono[OF a c])
- fix t assume "t \<in> space M"
- from mono[OF this] show "pos_part f t \<le> pos_part g t"
- unfolding pos_part_def by auto
- qed
- moreover have "d \<le> b"
- proof (rule nnfis_mono[OF d b])
- fix t assume "t \<in> space M"
- from mono[OF this] show "neg_part g t \<le> neg_part f t"
- unfolding neg_part_def by auto
- qed
- ultimately have "a - b \<le> c - d" by auto
- thus ?thesis unfolding integral_def
- using a b c d by (simp add: the_nnfis)
-qed
-
-lemma integral_uminus:
- assumes "integrable f"
- shows "integrable (\<lambda>t. - f t)"
- and "integral (\<lambda>t. - f t) = - integral f"
-proof -
- have "pos_part f = neg_part (\<lambda>t.-f t)" and "neg_part f = pos_part (\<lambda>t.-f t)"
- unfolding pos_part_def neg_part_def by (auto intro!: ext)
- with assms show "integrable (\<lambda>t.-f t)" and "integral (\<lambda>t.-f t) = -integral f"
- unfolding integrable_def integral_def by simp_all
-qed
-
-lemma integral_times_const:
- assumes "integrable f"
- shows "integrable (\<lambda>t. a * f t)" (is "?P a")
- and "integral (\<lambda>t. a * f t) = a * integral f" (is "?I a")
-proof -
- { fix a :: real assume "0 \<le> a"
- hence "pos_part (\<lambda>t. a * f t) = (\<lambda>t. a * pos_part f t)"
- and "neg_part (\<lambda>t. a * f t) = (\<lambda>t. a * neg_part f t)"
- unfolding pos_part_def neg_part_def max_def min_def
- by (auto intro!: ext simp: zero_le_mult_iff)
- moreover
- obtain x y where x: "x \<in> nnfis (pos_part f)" and y: "y \<in> nnfis (neg_part f)"
- using assms unfolding integrable_def by auto
- ultimately
- have "a * x \<in> nnfis (pos_part (\<lambda>t. a * f t))" and
- "a * y \<in> nnfis (neg_part (\<lambda>t. a * f t))"
- using nnfis_times[OF _ `0 \<le> a`] by auto
- with x y have "?P a \<and> ?I a"
- unfolding integrable_def integral_def by (auto simp: algebra_simps) }
- note int = this
-
- have "?P a \<and> ?I a"
- proof (cases "0 \<le> a")
- case True from int[OF this] show ?thesis .
- next
- case False with int[of "- a"] integral_uminus[of "\<lambda>t. - a * f t"]
- show ?thesis by auto
- qed
- thus "integrable (\<lambda>t. a * f t)"
- and "integral (\<lambda>t. a * f t) = a * integral f" by simp_all
-qed
-
-lemma integral_cmul_indicator:
- assumes "s \<in> sets M"
- shows "integral (\<lambda>x. c * indicator_fn s x) = c * (measure M s)"
- and "integrable (\<lambda>x. c * indicator_fn s x)"
-using assms integral_times_const integral_indicator_fn by auto
-
-lemma integral_zero:
- shows "integral (\<lambda>x. 0) = 0"
- and "integrable (\<lambda>x. 0)"
- using integral_cmul_indicator[OF empty_sets, of 0]
- unfolding indicator_fn_def by auto
-
-lemma integral_setsum:
- assumes "finite S"
- assumes "\<And>n. n \<in> S \<Longrightarrow> integrable (f n)"
- shows "integral (\<lambda>x. \<Sum> i \<in> S. f i x) = (\<Sum> i \<in> S. integral (f i))" (is "?int S")
- and "integrable (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I s")
-proof -
- from assms have "?int S \<and> ?I S"
- proof (induct S)
- case empty thus ?case by (simp add: integral_zero)
- next
- case (insert i S)
- thus ?case
- apply simp
- apply (subst integral_add)
- using assms apply auto
- apply (subst integral_add)
- using assms by auto
- qed
- thus "?int S" and "?I S" by auto
-qed
-
-lemma (in measure_space) integrable_abs:
- assumes "integrable f"
- shows "integrable (\<lambda> x. \<bar>f x\<bar>)"
-using assms
-proof -
- from assms obtain p q where pq: "p \<in> nnfis (pos_part f)" "q \<in> nnfis (neg_part f)"
- unfolding integrable_def by auto
- hence "p + q \<in> nnfis (\<lambda> x. pos_part f x + neg_part f x)"
- using nnfis_add by auto
- hence "p + q \<in> nnfis (\<lambda> x. \<bar>f x\<bar>)" using pos_neg_part_abs[of f] by simp
- thus ?thesis unfolding integrable_def
- using ext[OF pos_part_abs[of f], of "\<lambda> y. y"]
- ext[OF neg_part_abs[of f], of "\<lambda> y. y"]
- using nnfis_0 by auto
-qed
-
-lemma markov_ineq:
- assumes "integrable f" "0 < a" "integrable (\<lambda>x. \<bar>f x\<bar>^n)"
- shows "measure M (f -` {a ..} \<inter> space M) \<le> integral (\<lambda>x. \<bar>f x\<bar>^n) / a^n"
-using assms
-proof -
- from assms have "0 < a ^ n" using real_root_pow_pos by auto
- from assms have "f \<in> borel_measurable M"
- using integral_borel_measurable[OF `integrable f`] by auto
- hence w: "{w . w \<in> space M \<and> a \<le> f w} \<in> sets M"
- using borel_measurable_ge_iff by auto
- have i: "integrable (indicator_fn {w . w \<in> space M \<and> a \<le> f w})"
- using integral_indicator_fn[OF w] by simp
- have v1: "\<And> t. a ^ n * (indicator_fn {w . w \<in> space M \<and> a \<le> f w}) t
- \<le> (f t) ^ n * (indicator_fn {w . w \<in> space M \<and> a \<le> f w}) t"
- unfolding indicator_fn_def
- using `0 < a` power_mono[of a] assms by auto
- have v2: "\<And> t. (f t) ^ n * (indicator_fn {w . w \<in> space M \<and> a \<le> f w}) t \<le> \<bar>f t\<bar> ^ n"
- unfolding indicator_fn_def
- using power_mono[of a _ n] abs_ge_self `a > 0`
- by auto
- have "{w \<in> space M. a \<le> f w} \<inter> space M = {w . a \<le> f w} \<inter> space M"
- using Collect_eq by auto
- from Int_absorb2[OF sets_into_space[OF w]] `0 < a ^ n` sets_into_space[OF w] w this
- have "(a ^ n) * (measure M ((f -` {y . a \<le> y}) \<inter> space M)) =
- (a ^ n) * measure M {w . w \<in> space M \<and> a \<le> f w}"
- unfolding vimage_Collect_eq[of f] by simp
- also have "\<dots> = integral (\<lambda> t. a ^ n * (indicator_fn {w . w \<in> space M \<and> a \<le> f w}) t)"
- using integral_cmul_indicator[OF w] i by auto
- also have "\<dots> \<le> integral (\<lambda> t. \<bar> f t \<bar> ^ n)"
- apply (rule integral_mono)
- using integral_cmul_indicator[OF w]
- `integrable (\<lambda> x. \<bar>f x\<bar> ^ n)` order_trans[OF v1 v2] by auto
- finally show "measure M (f -` {a ..} \<inter> space M) \<le> integral (\<lambda>x. \<bar>f x\<bar>^n) / a^n"
- unfolding atLeast_def
- by (auto intro!: mult_imp_le_div_pos[OF `0 < a ^ n`], simp add: mult_commute)
-qed
-
-lemma (in measure_space) integral_0:
- fixes f :: "'a \<Rightarrow> real"
- assumes "integrable f" "integral f = 0" "nonneg f" and borel: "f \<in> borel_measurable M"
- shows "measure M ({x. f x \<noteq> 0} \<inter> space M) = 0"
-proof -
- have "{x. f x \<noteq> 0} = {x. \<bar>f x\<bar> > 0}" by auto
- moreover
- { fix y assume "y \<in> {x. \<bar> f x \<bar> > 0}"
- hence "\<bar> f y \<bar> > 0" by auto
- hence "\<exists> n. \<bar>f y\<bar> \<ge> inverse (real (Suc n))"
- using ex_inverse_of_nat_Suc_less[of "\<bar>f y\<bar>"] less_imp_le unfolding real_of_nat_def by auto
- hence "y \<in> (\<Union> n. {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))})"
- by auto }
- moreover
- { fix y assume "y \<in> (\<Union> n. {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))})"
- then obtain n where n: "y \<in> {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))}" by auto
- hence "\<bar>f y\<bar> \<ge> inverse (real (Suc n))" by auto
- hence "\<bar>f y\<bar> > 0"
- using real_of_nat_Suc_gt_zero
- positive_imp_inverse_positive[of "real_of_nat (Suc n)"] by fastsimp
- hence "y \<in> {x. \<bar>f x\<bar> > 0}" by auto }
- ultimately have fneq0_UN: "{x. f x \<noteq> 0} = (\<Union> n. {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))})"
- by blast
- { fix n
- have int_one: "integrable (\<lambda> x. \<bar>f x\<bar> ^ 1)" using integrable_abs assms by auto
- have "measure M (f -` {inverse (real (Suc n))..} \<inter> space M)
- \<le> integral (\<lambda> x. \<bar>f x\<bar> ^ 1) / (inverse (real (Suc n)) ^ 1)"
- using markov_ineq[OF `integrable f` _ int_one] real_of_nat_Suc_gt_zero by auto
- hence le0: "measure M (f -` {inverse (real (Suc n))..} \<inter> space M) \<le> 0"
- using assms unfolding nonneg_def by auto
- have "{x. f x \<ge> inverse (real (Suc n))} \<inter> space M \<in> sets M"
- apply (subst Int_commute) unfolding Int_def
- using borel[unfolded borel_measurable_ge_iff] by simp
- hence m0: "measure M ({x. f x \<ge> inverse (real (Suc n))} \<inter> space M) = 0 \<and>
- {x. f x \<ge> inverse (real (Suc n))} \<inter> space M \<in> sets M"
- using positive le0 unfolding atLeast_def by fastsimp }
- moreover hence "range (\<lambda> n. {x. f x \<ge> inverse (real (Suc n))} \<inter> space M) \<subseteq> sets M"
- by auto
- moreover
- { fix n
- have "inverse (real (Suc n)) \<ge> inverse (real (Suc (Suc n)))"
- using less_imp_inverse_less real_of_nat_Suc_gt_zero[of n] by fastsimp
- hence "\<And> x. f x \<ge> inverse (real (Suc n)) \<Longrightarrow> f x \<ge> inverse (real (Suc (Suc n)))" by (rule order_trans)
- hence "{x. f x \<ge> inverse (real (Suc n))} \<inter> space M
- \<subseteq> {x. f x \<ge> inverse (real (Suc (Suc n)))} \<inter> space M" by auto }
- ultimately have "(\<lambda> x. 0) ----> measure M (\<Union> n. {x. f x \<ge> inverse (real (Suc n))} \<inter> space M)"
- using monotone_convergence[of "\<lambda> n. {x. f x \<ge> inverse (real (Suc n))} \<inter> space M"]
- unfolding o_def by (simp del: of_nat_Suc)
- hence "measure M (\<Union> n. {x. f x \<ge> inverse (real (Suc n))} \<inter> space M) = 0"
- using LIMSEQ_const[of 0] LIMSEQ_unique by simp
- hence "measure M ((\<Union> n. {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))}) \<inter> space M) = 0"
- using assms unfolding nonneg_def by auto
- thus "measure M ({x. f x \<noteq> 0} \<inter> space M) = 0" using fneq0_UN by simp
-qed
-
-section "Lebesgue integration on countable spaces"
-
-lemma nnfis_on_countable:
- assumes borel: "f \<in> borel_measurable M"
- and bij: "bij_betw enum S (f ` space M - {0})"
- and enum_zero: "enum ` (-S) \<subseteq> {0}"
- and nn_enum: "\<And>n. 0 \<le> enum n"
- and sums: "(\<lambda>r. enum r * measure M (f -` {enum r} \<inter> space M)) sums x" (is "?sum sums x")
- shows "x \<in> nnfis f"
-proof -
- have inj_enum: "inj_on enum S"
- and range_enum: "enum ` S = f ` space M - {0}"
- using bij by (auto simp: bij_betw_def)
-
- let "?x n z" = "\<Sum>i = 0..<n. enum i * indicator_fn (f -` {enum i} \<inter> space M) z"
-
- show ?thesis
- proof (rule nnfis_mon_conv)
- show "(\<lambda>n. \<Sum>i = 0..<n. ?sum i) ----> x" using sums unfolding sums_def .
- next
- fix n
- show "(\<Sum>i = 0..<n. ?sum i) \<in> nnfis (?x n)"
- proof (induct n)
- case 0 thus ?case by (simp add: nnfis_0)
- next
- case (Suc n) thus ?case using nn_enum
- by (auto intro!: nnfis_add nnfis_times psfis_nnfis[OF psfis_indicator] borel_measurable_vimage[OF borel])
- qed
- next
- show "mono_convergent ?x f (space M)"
- proof (rule mono_convergentI)
- fix x
- show "incseq (\<lambda>n. ?x n x)"
- by (rule incseq_SucI, auto simp: indicator_fn_def nn_enum)
-
- have fin: "\<And>n. finite (enum ` ({0..<n} \<inter> S))" by auto
-
- assume "x \<in> space M"
- hence "f x \<in> enum ` S \<or> f x = 0" using range_enum by auto
- thus "(\<lambda>n. ?x n x) ----> f x"
- proof (rule disjE)
- assume "f x \<in> enum ` S"
- then obtain i where "i \<in> S" and "f x = enum i" by auto
-
- { fix n
- have sum_ranges:
- "i < n \<Longrightarrow> enum`({0..<n} \<inter> S) \<inter> {z. enum i = z \<and> x\<in>space M} = {enum i}"
- "\<not> i < n \<Longrightarrow> enum`({0..<n} \<inter> S) \<inter> {z. enum i = z \<and> x\<in>space M} = {}"
- using `x \<in> space M` `i \<in> S` inj_enum[THEN inj_on_iff] by auto
- have "?x n x =
- (\<Sum>i \<in> {0..<n} \<inter> S. enum i * indicator_fn (f -` {enum i} \<inter> space M) x)"
- using enum_zero by (auto intro!: setsum_mono_zero_cong_right)
- also have "... =
- (\<Sum>z \<in> enum`({0..<n} \<inter> S). z * indicator_fn (f -` {z} \<inter> space M) x)"
- using inj_enum[THEN subset_inj_on] by (auto simp: setsum_reindex)
- also have "... = (if i < n then f x else 0)"
- unfolding indicator_fn_def if_distrib[where x=1 and y=0]
- setsum_cases[OF fin]
- using sum_ranges `f x = enum i`
- by auto
- finally have "?x n x = (if i < n then f x else 0)" . }
- note sum_equals_if = this
-
- show ?thesis unfolding sum_equals_if
- by (rule LIMSEQ_offset[where k="i + 1"]) (auto intro!: LIMSEQ_const)
- next
- assume "f x = 0"
- { fix n have "?x n x = 0"
- unfolding indicator_fn_def if_distrib[where x=1 and y=0]
- setsum_cases[OF finite_atLeastLessThan]
- using `f x = 0` `x \<in> space M`
- by (auto split: split_if) }
- thus ?thesis using `f x = 0` by (auto intro!: LIMSEQ_const)
- qed
- qed
- qed
-qed
-
-lemma integral_on_countable:
- fixes enum :: "nat \<Rightarrow> real"
- assumes borel: "f \<in> borel_measurable M"
- and bij: "bij_betw enum S (f ` space M)"
- and enum_zero: "enum ` (-S) \<subseteq> {0}"
- and abs_summable: "summable (\<lambda>r. \<bar>enum r * measure M (f -` {enum r} \<inter> space M)\<bar>)"
- shows "integrable f"
- and "integral f = (\<Sum>r. enum r * measure M (f -` {enum r} \<inter> space M))" (is "_ = suminf (?sum f enum)")
-proof -
- { fix f enum
- assume borel: "f \<in> borel_measurable M"
- and bij: "bij_betw enum S (f ` space M)"
- and enum_zero: "enum ` (-S) \<subseteq> {0}"
- and abs_summable: "summable (\<lambda>r. \<bar>enum r * measure M (f -` {enum r} \<inter> space M)\<bar>)"
- have inj_enum: "inj_on enum S" and range_enum: "f ` space M = enum ` S"
- using bij unfolding bij_betw_def by auto
-
- have [simp, intro]: "\<And>X. 0 \<le> measure M (f -` {X} \<inter> space M)"
- by (rule positive, rule borel_measurable_vimage[OF borel])
-
- have "(\<Sum>r. ?sum (pos_part f) (pos_part enum) r) \<in> nnfis (pos_part f) \<and>
- summable (\<lambda>r. ?sum (pos_part f) (pos_part enum) r)"
- proof (rule conjI, rule nnfis_on_countable)
- have pos_f_image: "pos_part f ` space M - {0} = f ` space M \<inter> {0<..}"
- unfolding pos_part_def max_def by auto
-
- show "bij_betw (pos_part enum) {x \<in> S. 0 < enum x} (pos_part f ` space M - {0})"
- unfolding bij_betw_def pos_f_image
- unfolding pos_part_def max_def range_enum
- by (auto intro!: inj_onI simp: inj_enum[THEN inj_on_eq_iff])
-
- show "\<And>n. 0 \<le> pos_part enum n" unfolding pos_part_def max_def by auto
-
- show "pos_part f \<in> borel_measurable M"
- by (rule pos_part_borel_measurable[OF borel])
-
- show "pos_part enum ` (- {x \<in> S. 0 < enum x}) \<subseteq> {0}"
- unfolding pos_part_def max_def using enum_zero by auto
-
- show "summable (\<lambda>r. ?sum (pos_part f) (pos_part enum) r)"
- proof (rule summable_comparison_test[OF _ abs_summable], safe intro!: exI[of _ 0])
- fix n :: nat
- have "pos_part enum n \<noteq> 0 \<Longrightarrow> (pos_part f -` {enum n} \<inter> space M) =
- (if 0 < enum n then (f -` {enum n} \<inter> space M) else {})"
- unfolding pos_part_def max_def by (auto split: split_if_asm)
- thus "norm (?sum (pos_part f) (pos_part enum) n) \<le> \<bar>?sum f enum n \<bar>"
- by (cases "pos_part enum n = 0",
- auto simp: pos_part_def max_def abs_mult not_le split: split_if_asm intro!: mult_nonpos_nonneg)
- qed
- thus "(\<lambda>r. ?sum (pos_part f) (pos_part enum) r) sums (\<Sum>r. ?sum (pos_part f) (pos_part enum) r)"
- by (rule summable_sums)
- qed }
- note pos = this
-
- note pos_part = pos[OF assms(1-4)]
- moreover
- have neg_part_to_pos_part:
- "\<And>f :: _ \<Rightarrow> real. neg_part f = pos_part (uminus \<circ> f)"
- by (auto simp: pos_part_def neg_part_def min_def max_def expand_fun_eq)
-
- have neg_part: "(\<Sum>r. ?sum (neg_part f) (neg_part enum) r) \<in> nnfis (neg_part f) \<and>
- summable (\<lambda>r. ?sum (neg_part f) (neg_part enum) r)"
- unfolding neg_part_to_pos_part
- proof (rule pos)
- show "uminus \<circ> f \<in> borel_measurable M" unfolding comp_def
- by (rule borel_measurable_uminus_borel_measurable[OF borel])
-
- show "bij_betw (uminus \<circ> enum) S ((uminus \<circ> f) ` space M)"
- using bij unfolding bij_betw_def
- by (auto intro!: comp_inj_on simp: image_compose)
-
- show "(uminus \<circ> enum) ` (- S) \<subseteq> {0}"
- using enum_zero by auto
-
- have minus_image: "\<And>r. (uminus \<circ> f) -` {(uminus \<circ> enum) r} \<inter> space M = f -` {enum r} \<inter> space M"
- by auto
- show "summable (\<lambda>r. \<bar>(uminus \<circ> enum) r * measure_space.measure M ((uminus \<circ> f) -` {(uminus \<circ> enum) r} \<inter> space M)\<bar>)"
- unfolding minus_image using abs_summable by simp
- qed
- ultimately show "integrable f" unfolding integrable_def by auto
-
- { fix r
- have "?sum (pos_part f) (pos_part enum) r - ?sum (neg_part f) (neg_part enum) r = ?sum f enum r"
- proof (cases rule: linorder_cases)
- assume "0 < enum r"
- hence "pos_part f -` {enum r} \<inter> space M = f -` {enum r} \<inter> space M"
- unfolding pos_part_def max_def by (auto split: split_if_asm)
- with `0 < enum r` show ?thesis unfolding pos_part_def neg_part_def min_def max_def expand_fun_eq
- by auto
- next
- assume "enum r < 0"
- hence "neg_part f -` {- enum r} \<inter> space M = f -` {enum r} \<inter> space M"
- unfolding neg_part_def min_def by (auto split: split_if_asm)
- with `enum r < 0` show ?thesis unfolding pos_part_def neg_part_def min_def max_def expand_fun_eq
- by auto
- qed (simp add: neg_part_def pos_part_def) }
- note sum_diff_eq_sum = this
-
- have "(\<Sum>r. ?sum (pos_part f) (pos_part enum) r) - (\<Sum>r. ?sum (neg_part f) (neg_part enum) r)
- = (\<Sum>r. ?sum (pos_part f) (pos_part enum) r - ?sum (neg_part f) (neg_part enum) r)"
- using neg_part pos_part by (auto intro: suminf_diff)
- also have "... = (\<Sum>r. ?sum f enum r)" unfolding sum_diff_eq_sum ..
- finally show "integral f = suminf (?sum f enum)"
- unfolding integral_def using pos_part neg_part
- by (auto dest: the_nnfis)
-qed
-
-section "Lebesgue integration on finite space"
-
-lemma integral_finite_on_sets:
- assumes "f \<in> borel_measurable M" and "finite (space M)" and "a \<in> sets M"
- shows "integral (\<lambda>x. f x * indicator_fn a x) =
- (\<Sum> r \<in> f`a. r * measure M (f -` {r} \<inter> a))" (is "integral ?f = _")
-proof -
- { fix x assume "x \<in> a"
- with assms have "f -` {f x} \<inter> space M \<in> sets M"
- by (subst Int_commute)
- (auto simp: vimage_def Int_def
- intro!: borel_measurable_const
- borel_measurable_eq_borel_measurable)
- from Int[OF this assms(3)]
- sets_into_space[OF assms(3), THEN Int_absorb1]
- have "f -` {f x} \<inter> a \<in> sets M" by (simp add: Int_assoc) }
- note vimage_f = this
-
- have "finite a"
- using assms(2,3) sets_into_space
- by (auto intro: finite_subset)
-
- have "integral (\<lambda>x. f x * indicator_fn a x) =
- integral (\<lambda>x. \<Sum>i\<in>f ` a. i * indicator_fn (f -` {i} \<inter> a) x)"
- (is "_ = integral (\<lambda>x. setsum (?f x) _)")
- unfolding indicator_fn_def if_distrib
- using `finite a` by (auto simp: setsum_cases intro!: integral_cong)
- also have "\<dots> = (\<Sum>i\<in>f`a. integral (\<lambda>x. ?f x i))"
- proof (rule integral_setsum, safe)
- fix n x assume "x \<in> a"
- thus "integrable (\<lambda>y. ?f y (f x))"
- using integral_indicator_fn(2)[OF vimage_f]
- by (auto intro!: integral_times_const)
- qed (simp add: `finite a`)
- also have "\<dots> = (\<Sum>i\<in>f`a. i * measure M (f -` {i} \<inter> a))"
- using integral_cmul_indicator[OF vimage_f]
- by (auto intro!: setsum_cong)
- finally show ?thesis .
-qed
-
-lemma integral_finite:
- assumes "f \<in> borel_measurable M" and "finite (space M)"
- shows "integral f = (\<Sum> r \<in> f ` space M. r * measure M (f -` {r} \<inter> space M))"
- using integral_finite_on_sets[OF assms top]
- integral_cong[of "\<lambda>x. f x * indicator_fn (space M) x" f]
- by (auto simp add: indicator_fn_def)
-
-section "Radon–Nikodym derivative"
-
-definition
- "RN_deriv v \<equiv> SOME f. measure_space (M\<lparr>measure := v\<rparr>) \<and>
- f \<in> borel_measurable M \<and>
- (\<forall>a \<in> sets M. (integral (\<lambda>x. f x * indicator_fn a x) = v a))"
-
-end
-
-lemma sigma_algebra_cong:
- fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
- assumes *: "sigma_algebra M"
- and cong: "space M = space M'" "sets M = sets M'"
- shows "sigma_algebra M'"
-using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
-
-lemma finite_Pow_additivity_sufficient:
- assumes "finite (space M)" and "sets M = Pow (space M)"
- and "positive M (measure M)" and "additive M (measure M)"
- shows "finite_measure_space M"
-proof -
- have "sigma_algebra M"
- using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow])
-
- have "measure_space M"
- by (rule Measure.finite_additivity_sufficient) (fact+)
- thus ?thesis
- unfolding finite_measure_space_def finite_measure_space_axioms_def
- using assms by simp
-qed
-
-lemma finite_measure_spaceI:
- assumes "measure_space M" and "finite (space M)" and "sets M = Pow (space M)"
- shows "finite_measure_space M"
- unfolding finite_measure_space_def finite_measure_space_axioms_def
- using assms by simp
-
-lemma (in finite_measure_space) integral_finite_singleton:
- "integral f = (\<Sum>x \<in> space M. f x * measure M {x})"
-proof -
- have "f \<in> borel_measurable M"
- unfolding borel_measurable_le_iff
- using sets_eq_Pow by auto
- { fix r let ?x = "f -` {r} \<inter> space M"
- have "?x \<subseteq> space M" by auto
- with finite_space sets_eq_Pow have "measure M ?x = (\<Sum>i \<in> ?x. measure M {i})"
- by (auto intro!: measure_real_sum_image) }
- note measure_eq_setsum = this
- show ?thesis
- unfolding integral_finite[OF `f \<in> borel_measurable M` finite_space]
- measure_eq_setsum setsum_right_distrib
- apply (subst setsum_Sigma)
- apply (simp add: finite_space)
- apply (simp add: finite_space)
- proof (rule setsum_reindex_cong[symmetric])
- fix a assume "a \<in> Sigma (f ` space M) (\<lambda>x. f -` {x} \<inter> space M)"
- thus "(\<lambda>(x, y). x * measure M {y}) a = f (snd a) * measure_space.measure M {snd a}"
- by auto
- qed (auto intro!: image_eqI inj_onI)
-qed
-
-lemma (in finite_measure_space) RN_deriv_finite_singleton:
- fixes v :: "'a set \<Rightarrow> real"
- assumes ms_v: "measure_space (M\<lparr>measure := v\<rparr>)"
- and eq_0: "\<And>x. \<lbrakk> x \<in> space M ; measure M {x} = 0 \<rbrakk> \<Longrightarrow> v {x} = 0"
- and "x \<in> space M" and "measure M {x} \<noteq> 0"
- shows "RN_deriv v x = v {x} / (measure M {x})" (is "_ = ?v x")
- unfolding RN_deriv_def
-proof (rule someI2_ex[where Q = "\<lambda>f. f x = ?v x"], rule exI[where x = ?v], safe)
- show "(\<lambda>a. v {a} / measure_space.measure M {a}) \<in> borel_measurable M"
- unfolding borel_measurable_le_iff using sets_eq_Pow by auto
-next
- fix a assume "a \<in> sets M"
- hence "a \<subseteq> space M" and "finite a"
- using sets_into_space finite_space by (auto intro: finite_subset)
- have *: "\<And>x a. x \<in> space M \<Longrightarrow> (if measure M {x} = 0 then 0 else v {x} * indicator_fn a x) =
- v {x} * indicator_fn a x" using eq_0 by auto
-
- from measure_space.measure_real_sum_image[OF ms_v, of a]
- sets_eq_Pow `a \<in> sets M` sets_into_space `finite a`
- have "v a = (\<Sum>x\<in>a. v {x})" by auto
- thus "integral (\<lambda>x. v {x} / measure_space.measure M {x} * indicator_fn a x) = v a"
- apply (simp add: eq_0 integral_finite_singleton)
- apply (unfold divide_1)
- by (simp add: * indicator_fn_def if_distrib setsum_cases finite_space `a \<subseteq> space M` Int_absorb1)
-next
- fix w assume "w \<in> borel_measurable M"
- assume int_eq_v: "\<forall>a\<in>sets M. integral (\<lambda>x. w x * indicator_fn a x) = v a"
- have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
-
- have "w x * measure M {x} =
- (\<Sum>y\<in>space M. w y * indicator_fn {x} y * measure M {y})"
- apply (subst (3) mult_commute)
- unfolding indicator_fn_def if_distrib setsum_cases[OF finite_space]
- using `x \<in> space M` by simp
- also have "... = v {x}"
- using int_eq_v[rule_format, OF `{x} \<in> sets M`]
- by (simp add: integral_finite_singleton)
- finally show "w x = v {x} / measure M {x}"
- using `measure M {x} \<noteq> 0` by (simp add: eq_divide_eq)
-qed fact
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Lebesgue_Integration.thy Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,2012 @@
+(* Author: Armin Heller, Johannes Hoelzl, TU Muenchen *)
+
+header {*Lebesgue Integration*}
+
+theory Lebesgue_Integration
+imports Measure Borel
+begin
+
+section "@{text \<mu>}-null sets"
+
+abbreviation (in measure_space) "null_sets \<equiv> {N\<in>sets M. \<mu> N = 0}"
+
+lemma sums_If_finite:
+ assumes finite: "finite {r. P r}"
+ shows "(\<lambda>r. if P r then f r else 0) sums (\<Sum>r\<in>{r. P r}. f r)" (is "?F sums _")
+proof cases
+ assume "{r. P r} = {}" hence "\<And>r. \<not> P r" by auto
+ thus ?thesis by (simp add: sums_zero)
+next
+ assume not_empty: "{r. P r} \<noteq> {}"
+ have "?F sums (\<Sum>r = 0..< Suc (Max {r. P r}). ?F r)"
+ by (rule series_zero)
+ (auto simp add: Max_less_iff[OF finite not_empty] less_eq_Suc_le[symmetric])
+ also have "(\<Sum>r = 0..< Suc (Max {r. P r}). ?F r) = (\<Sum>r\<in>{r. P r}. f r)"
+ by (subst setsum_cases)
+ (auto intro!: setsum_cong simp: Max_ge_iff[OF finite not_empty] less_Suc_eq_le)
+ finally show ?thesis .
+qed
+
+lemma sums_single:
+ "(\<lambda>r. if r = i then f r else 0) sums f i"
+ using sums_If_finite[of "\<lambda>r. r = i" f] by simp
+
+section "Simple function"
+
+text {*
+
+Our simple functions are not restricted to positive real numbers. Instead
+they are just functions with a finite range and are measurable when singleton
+sets are measurable.
+
+*}
+
+definition (in sigma_algebra) "simple_function g \<longleftrightarrow>
+ finite (g ` space M) \<and>
+ (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
+
+lemma (in sigma_algebra) simple_functionD:
+ assumes "simple_function g"
+ shows "finite (g ` space M)"
+ "x \<in> g ` space M \<Longrightarrow> g -` {x} \<inter> space M \<in> sets M"
+ using assms unfolding simple_function_def by auto
+
+lemma (in sigma_algebra) simple_function_indicator_representation:
+ fixes f ::"'a \<Rightarrow> pinfreal"
+ assumes f: "simple_function f" and x: "x \<in> space M"
+ shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
+ (is "?l = ?r")
+proof -
+ have "?r = (\<Sum>y \<in> f ` space M.
+ (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
+ by (auto intro!: setsum_cong2)
+ also have "... = f x * indicator (f -` {f x} \<inter> space M) x"
+ using assms by (auto dest: simple_functionD simp: setsum_delta)
+ also have "... = f x" using x by (auto simp: indicator_def)
+ finally show ?thesis by auto
+qed
+
+lemma (in measure_space) simple_function_notspace:
+ "simple_function (\<lambda>x. h x * indicator (- space M) x::pinfreal)" (is "simple_function ?h")
+proof -
+ have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
+ hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
+ have "?h -` {0} \<inter> space M = space M" by auto
+ thus ?thesis unfolding simple_function_def by auto
+qed
+
+lemma (in sigma_algebra) simple_function_cong:
+ assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
+ shows "simple_function f \<longleftrightarrow> simple_function g"
+proof -
+ have "f ` space M = g ` space M"
+ "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
+ using assms by (auto intro!: image_eqI)
+ thus ?thesis unfolding simple_function_def using assms by simp
+qed
+
+lemma (in sigma_algebra) borel_measurable_simple_function:
+ assumes "simple_function f"
+ shows "f \<in> borel_measurable M"
+proof (rule borel_measurableI)
+ fix S
+ let ?I = "f ` (f -` S \<inter> space M)"
+ have *: "(\<Union>x\<in>?I. f -` {x} \<inter> space M) = f -` S \<inter> space M" (is "?U = _") by auto
+ have "finite ?I"
+ using assms unfolding simple_function_def by (auto intro: finite_subset)
+ hence "?U \<in> sets M"
+ apply (rule finite_UN)
+ using assms unfolding simple_function_def by auto
+ thus "f -` S \<inter> space M \<in> sets M" unfolding * .
+qed
+
+lemma (in sigma_algebra) simple_function_borel_measurable:
+ fixes f :: "'a \<Rightarrow> 'x::t2_space"
+ assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
+ shows "simple_function f"
+ using assms unfolding simple_function_def
+ by (auto intro: borel_measurable_vimage)
+
+lemma (in sigma_algebra) simple_function_const[intro, simp]:
+ "simple_function (\<lambda>x. c)"
+ by (auto intro: finite_subset simp: simple_function_def)
+
+lemma (in sigma_algebra) simple_function_compose[intro, simp]:
+ assumes "simple_function f"
+ shows "simple_function (g \<circ> f)"
+ unfolding simple_function_def
+proof safe
+ show "finite ((g \<circ> f) ` space M)"
+ using assms unfolding simple_function_def by (auto simp: image_compose)
+next
+ fix x assume "x \<in> space M"
+ let ?G = "g -` {g (f x)} \<inter> (f`space M)"
+ have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
+ (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
+ show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
+ using assms unfolding simple_function_def *
+ by (rule_tac finite_UN) (auto intro!: finite_UN)
+qed
+
+lemma (in sigma_algebra) simple_function_indicator[intro, simp]:
+ assumes "A \<in> sets M"
+ shows "simple_function (indicator A)"
+proof -
+ have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
+ by (auto simp: indicator_def)
+ hence "finite ?S" by (rule finite_subset) simp
+ moreover have "- A \<inter> space M = space M - A" by auto
+ ultimately show ?thesis unfolding simple_function_def
+ using assms by (auto simp: indicator_def_raw)
+qed
+
+lemma (in sigma_algebra) simple_function_Pair[intro, simp]:
+ assumes "simple_function f"
+ assumes "simple_function g"
+ shows "simple_function (\<lambda>x. (f x, g x))" (is "simple_function ?p")
+ unfolding simple_function_def
+proof safe
+ show "finite (?p ` space M)"
+ using assms unfolding simple_function_def
+ by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto
+next
+ fix x assume "x \<in> space M"
+ have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
+ (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
+ by auto
+ with `x \<in> space M` show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
+ using assms unfolding simple_function_def by auto
+qed
+
+lemma (in sigma_algebra) simple_function_compose1:
+ assumes "simple_function f"
+ shows "simple_function (\<lambda>x. g (f x))"
+ using simple_function_compose[OF assms, of g]
+ by (simp add: comp_def)
+
+lemma (in sigma_algebra) simple_function_compose2:
+ assumes "simple_function f" and "simple_function g"
+ shows "simple_function (\<lambda>x. h (f x) (g x))"
+proof -
+ have "simple_function ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))"
+ using assms by auto
+ thus ?thesis by (simp_all add: comp_def)
+qed
+
+lemmas (in sigma_algebra) simple_function_add[intro, simp] = simple_function_compose2[where h="op +"]
+ and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"]
+ and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
+ and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
+ and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"]
+ and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
+
+lemma (in sigma_algebra) simple_function_setsum[intro, simp]:
+ assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function (f i)"
+ shows "simple_function (\<lambda>x. \<Sum>i\<in>P. f i x)"
+proof cases
+ assume "finite P" from this assms show ?thesis by induct auto
+qed auto
+
+lemma (in sigma_algebra) simple_function_le_measurable:
+ assumes "simple_function f" "simple_function g"
+ shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
+proof -
+ have *: "{x \<in> space M. f x \<le> g x} =
+ (\<Union>(F, G)\<in>f`space M \<times> g`space M.
+ if F \<le> G then (f -` {F} \<inter> space M) \<inter> (g -` {G} \<inter> space M) else {})"
+ apply (auto split: split_if_asm)
+ apply (rule_tac x=x in bexI)
+ apply (rule_tac x=x in bexI)
+ by simp_all
+ have **: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow>
+ (f -` {f x} \<inter> space M) \<inter> (g -` {g y} \<inter> space M) \<in> sets M"
+ using assms unfolding simple_function_def by auto
+ have "finite (f`space M \<times> g`space M)"
+ using assms unfolding simple_function_def by auto
+ thus ?thesis unfolding *
+ apply (rule finite_UN)
+ using assms unfolding simple_function_def
+ by (auto intro!: **)
+qed
+
+lemma setsum_indicator_disjoint_family:
+ fixes f :: "'d \<Rightarrow> 'e::semiring_1"
+ assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
+ shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
+proof -
+ have "P \<inter> {i. x \<in> A i} = {j}"
+ using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
+ by auto
+ thus ?thesis
+ unfolding indicator_def
+ by (simp add: if_distrib setsum_cases[OF `finite P`])
+qed
+
+lemma LeastI2_wellorder:
+ fixes a :: "_ :: wellorder"
+ assumes "P a"
+ and "\<And>a. \<lbrakk> P a; \<forall>b. P b \<longrightarrow> a \<le> b \<rbrakk> \<Longrightarrow> Q a"
+ shows "Q (Least P)"
+proof (rule LeastI2_order)
+ show "P (Least P)" using `P a` by (rule LeastI)
+next
+ fix y assume "P y" thus "Least P \<le> y" by (rule Least_le)
+next
+ fix x assume "P x" "\<forall>y. P y \<longrightarrow> x \<le> y" thus "Q x" by (rule assms(2))
+qed
+
+lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
+ fixes u :: "'a \<Rightarrow> pinfreal"
+ assumes u: "u \<in> borel_measurable M"
+ shows "\<exists>f. (\<forall>i. simple_function (f i) \<and> (\<forall>x\<in>space M. f i x \<noteq> \<omega>)) \<and> f \<up> u"
+proof -
+ have "\<exists>f. \<forall>x j. (of_nat j \<le> u x \<longrightarrow> f x j = j*2^j) \<and>
+ (u x < of_nat j \<longrightarrow> of_nat (f x j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f x j)))"
+ (is "\<exists>f. \<forall>x j. ?P x j (f x j)")
+ proof(rule choice, rule, rule choice, rule)
+ fix x j show "\<exists>n. ?P x j n"
+ proof cases
+ assume *: "u x < of_nat j"
+ then obtain r where r: "u x = Real r" "0 \<le> r" by (cases "u x") auto
+ from reals_Archimedean6a[of "r * 2^j"]
+ obtain n :: nat where "real n \<le> r * 2 ^ j" "r * 2 ^ j < real (Suc n)"
+ using `0 \<le> r` by (auto simp: zero_le_power zero_le_mult_iff)
+ thus ?thesis using r * by (auto intro!: exI[of _ n])
+ qed auto
+ qed
+ then obtain f where top: "\<And>j x. of_nat j \<le> u x \<Longrightarrow> f x j = j*2^j" and
+ upper: "\<And>j x. u x < of_nat j \<Longrightarrow> of_nat (f x j) \<le> u x * 2^j" and
+ lower: "\<And>j x. u x < of_nat j \<Longrightarrow> u x * 2^j < of_nat (Suc (f x j))" by blast
+
+ { fix j x P
+ assume 1: "of_nat j \<le> u x \<Longrightarrow> P (j * 2^j)"
+ assume 2: "\<And>k. \<lbrakk> u x < of_nat j ; of_nat k \<le> u x * 2^j ; u x * 2^j < of_nat (Suc k) \<rbrakk> \<Longrightarrow> P k"
+ have "P (f x j)"
+ proof cases
+ assume "of_nat j \<le> u x" thus "P (f x j)"
+ using top[of j x] 1 by auto
+ next
+ assume "\<not> of_nat j \<le> u x"
+ hence "u x < of_nat j" "of_nat (f x j) \<le> u x * 2^j" "u x * 2^j < of_nat (Suc (f x j))"
+ using upper lower by auto
+ from 2[OF this] show "P (f x j)" .
+ qed }
+ note fI = this
+
+ { fix j x
+ have "f x j = j * 2 ^ j \<longleftrightarrow> of_nat j \<le> u x"
+ by (rule fI, simp, cases "u x") (auto split: split_if_asm) }
+ note f_eq = this
+
+ { fix j x
+ have "f x j \<le> j * 2 ^ j"
+ proof (rule fI)
+ fix k assume *: "u x < of_nat j"
+ assume "of_nat k \<le> u x * 2 ^ j"
+ also have "\<dots> \<le> of_nat (j * 2^j)"
+ using * by (cases "u x") (auto simp: zero_le_mult_iff)
+ finally show "k \<le> j*2^j" by (auto simp del: real_of_nat_mult)
+ qed simp }
+ note f_upper = this
+
+ let "?g j x" = "of_nat (f x j) / 2^j :: pinfreal"
+ show ?thesis unfolding simple_function_def isoton_fun_expand unfolding isoton_def
+ proof (safe intro!: exI[of _ ?g])
+ fix j
+ have *: "?g j ` space M \<subseteq> (\<lambda>x. of_nat x / 2^j) ` {..j * 2^j}"
+ using f_upper by auto
+ thus "finite (?g j ` space M)" by (rule finite_subset) auto
+ next
+ fix j t assume "t \<in> space M"
+ have **: "?g j -` {?g j t} \<inter> space M = {x \<in> space M. f t j = f x j}"
+ by (auto simp: power_le_zero_eq Real_eq_Real mult_le_0_iff)
+
+ show "?g j -` {?g j t} \<inter> space M \<in> sets M"
+ proof cases
+ assume "of_nat j \<le> u t"
+ hence "?g j -` {?g j t} \<inter> space M = {x\<in>space M. of_nat j \<le> u x}"
+ unfolding ** f_eq[symmetric] by auto
+ thus "?g j -` {?g j t} \<inter> space M \<in> sets M"
+ using u by auto
+ next
+ assume not_t: "\<not> of_nat j \<le> u t"
+ hence less: "f t j < j*2^j" using f_eq[symmetric] `f t j \<le> j*2^j` by auto
+ have split_vimage: "?g j -` {?g j t} \<inter> space M =
+ {x\<in>space M. of_nat (f t j)/2^j \<le> u x} \<inter> {x\<in>space M. u x < of_nat (Suc (f t j))/2^j}"
+ unfolding **
+ proof safe
+ fix x assume [simp]: "f t j = f x j"
+ have *: "\<not> of_nat j \<le> u x" using not_t f_eq[symmetric] by simp
+ hence "of_nat (f t j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f t j))"
+ using upper lower by auto
+ hence "of_nat (f t j) / 2^j \<le> u x \<and> u x < of_nat (Suc (f t j))/2^j" using *
+ by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps)
+ thus "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j" by auto
+ next
+ fix x
+ assume "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j"
+ hence "of_nat (f t j) \<le> u x * 2 ^ j \<and> u x * 2 ^ j < of_nat (Suc (f t j))"
+ by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps)
+ hence 1: "of_nat (f t j) \<le> u x * 2 ^ j" and 2: "u x * 2 ^ j < of_nat (Suc (f t j))" by auto
+ note 2
+ also have "\<dots> \<le> of_nat (j*2^j)"
+ using less by (auto simp: zero_le_mult_iff simp del: real_of_nat_mult)
+ finally have bound_ux: "u x < of_nat j"
+ by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq)
+ show "f t j = f x j"
+ proof (rule antisym)
+ from 1 lower[OF bound_ux]
+ show "f t j \<le> f x j" by (cases "u x") (auto split: split_if_asm)
+ from upper[OF bound_ux] 2
+ show "f x j \<le> f t j" by (cases "u x") (auto split: split_if_asm)
+ qed
+ qed
+ show ?thesis unfolding split_vimage using u by auto
+ qed
+ next
+ fix j t assume "?g j t = \<omega>" thus False by (auto simp: power_le_zero_eq)
+ next
+ fix t
+ { fix i
+ have "f t i * 2 \<le> f t (Suc i)"
+ proof (rule fI)
+ assume "of_nat (Suc i) \<le> u t"
+ hence "of_nat i \<le> u t" by (cases "u t") auto
+ thus "f t i * 2 \<le> Suc i * 2 ^ Suc i" unfolding f_eq[symmetric] by simp
+ next
+ fix k
+ assume *: "u t * 2 ^ Suc i < of_nat (Suc k)"
+ show "f t i * 2 \<le> k"
+ proof (rule fI)
+ assume "of_nat i \<le> u t"
+ hence "of_nat (i * 2^Suc i) \<le> u t * 2^Suc i"
+ by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
+ also have "\<dots> < of_nat (Suc k)" using * by auto
+ finally show "i * 2 ^ i * 2 \<le> k"
+ by (auto simp del: real_of_nat_mult)
+ next
+ fix j assume "of_nat j \<le> u t * 2 ^ i"
+ with * show "j * 2 \<le> k" by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq)
+ qed
+ qed
+ thus "?g i t \<le> ?g (Suc i) t"
+ by (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) }
+ hence mono: "mono (\<lambda>i. ?g i t)" unfolding mono_iff_le_Suc by auto
+
+ show "(SUP j. of_nat (f t j) / 2 ^ j) = u t"
+ proof (rule pinfreal_SUPI)
+ fix j show "of_nat (f t j) / 2 ^ j \<le> u t"
+ proof (rule fI)
+ assume "of_nat j \<le> u t" thus "of_nat (j * 2 ^ j) / 2 ^ j \<le> u t"
+ by (cases "u t") (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps)
+ next
+ fix k assume "of_nat k \<le> u t * 2 ^ j"
+ thus "of_nat k / 2 ^ j \<le> u t"
+ by (cases "u t")
+ (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps zero_le_mult_iff)
+ qed
+ next
+ fix y :: pinfreal assume *: "\<And>j. j \<in> UNIV \<Longrightarrow> of_nat (f t j) / 2 ^ j \<le> y"
+ show "u t \<le> y"
+ proof (cases "u t")
+ case (preal r)
+ show ?thesis
+ proof (rule ccontr)
+ assume "\<not> u t \<le> y"
+ then obtain p where p: "y = Real p" "0 \<le> p" "p < r" using preal by (cases y) auto
+ with LIMSEQ_inverse_realpow_zero[of 2, unfolded LIMSEQ_iff, rule_format, of "r - p"]
+ obtain n where n: "\<And>N. n \<le> N \<Longrightarrow> inverse (2^N) < r - p" by auto
+ let ?N = "max n (natfloor r + 1)"
+ have "u t < of_nat ?N" "n \<le> ?N"
+ using ge_natfloor_plus_one_imp_gt[of r n] preal
+ by (auto simp: max_def real_Suc_natfloor)
+ from lower[OF this(1)]
+ have "r < (real (f t ?N) + 1) / 2 ^ ?N" unfolding less_divide_eq
+ using preal by (auto simp add: power_le_zero_eq zero_le_mult_iff real_of_nat_Suc split: split_if_asm)
+ hence "u t < of_nat (f t ?N) / 2 ^ ?N + 1 / 2 ^ ?N"
+ using preal by (auto simp: field_simps divide_real_def[symmetric])
+ with n[OF `n \<le> ?N`] p preal *[of ?N]
+ show False
+ by (cases "f t ?N") (auto simp: power_le_zero_eq split: split_if_asm)
+ qed
+ next
+ case infinite
+ { fix j have "f t j = j*2^j" using top[of j t] infinite by simp
+ hence "of_nat j \<le> y" using *[of j]
+ by (cases y) (auto simp: power_le_zero_eq zero_le_mult_iff field_simps) }
+ note all_less_y = this
+ show ?thesis unfolding infinite
+ proof (rule ccontr)
+ assume "\<not> \<omega> \<le> y" then obtain r where r: "y = Real r" "0 \<le> r" by (cases y) auto
+ moreover obtain n ::nat where "r < real n" using ex_less_of_nat by (auto simp: real_eq_of_nat)
+ with all_less_y[of n] r show False by auto
+ qed
+ qed
+ qed
+ qed
+qed
+
+lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
+ fixes u :: "'a \<Rightarrow> pinfreal"
+ assumes "u \<in> borel_measurable M"
+ obtains (x) f where "f \<up> u" "\<And>i. simple_function (f i)" "\<And>i. \<omega>\<notin>f i`space M"
+proof -
+ from borel_measurable_implies_simple_function_sequence[OF assms]
+ obtain f where x: "\<And>i. simple_function (f i)" "f \<up> u"
+ and fin: "\<And>i. \<And>x. x\<in>space M \<Longrightarrow> f i x \<noteq> \<omega>" by auto
+ { fix i from fin[of _ i] have "\<omega> \<notin> f i`space M" by fastsimp }
+ with x show thesis by (auto intro!: that[of f])
+qed
+
+section "Simple integral"
+
+definition (in measure_space)
+ "simple_integral f = (\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M))"
+
+lemma (in measure_space) simple_integral_cong:
+ assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
+ shows "simple_integral f = simple_integral g"
+proof -
+ have "f ` space M = g ` space M"
+ "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
+ using assms by (auto intro!: image_eqI)
+ thus ?thesis unfolding simple_integral_def by simp
+qed
+
+lemma (in measure_space) simple_integral_const[simp]:
+ "simple_integral (\<lambda>x. c) = c * \<mu> (space M)"
+proof (cases "space M = {}")
+ case True thus ?thesis unfolding simple_integral_def by simp
+next
+ case False hence "(\<lambda>x. c) ` space M = {c}" by auto
+ thus ?thesis unfolding simple_integral_def by simp
+qed
+
+lemma (in measure_space) simple_function_partition:
+ assumes "simple_function f" and "simple_function g"
+ shows "simple_integral f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. contents (f`A) * \<mu> A)"
+ (is "_ = setsum _ (?p ` space M)")
+proof-
+ let "?sub x" = "?p ` (f -` {x} \<inter> space M)"
+ let ?SIGMA = "Sigma (f`space M) ?sub"
+
+ have [intro]:
+ "finite (f ` space M)"
+ "finite (g ` space M)"
+ using assms unfolding simple_function_def by simp_all
+
+ { fix A
+ have "?p ` (A \<inter> space M) \<subseteq>
+ (\<lambda>(x,y). f -` {x} \<inter> g -` {y} \<inter> space M) ` (f`space M \<times> g`space M)"
+ by auto
+ hence "finite (?p ` (A \<inter> space M))"
+ by (rule finite_subset) (auto intro: finite_SigmaI finite_imageI) }
+ note this[intro, simp]
+
+ { fix x assume "x \<in> space M"
+ have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto
+ moreover {
+ fix x y
+ have *: "f -` {f x} \<inter> g -` {g x} \<inter> space M
+ = (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)" by auto
+ assume "x \<in> space M" "y \<in> space M"
+ hence "f -` {f x} \<inter> g -` {g x} \<inter> space M \<in> sets M"
+ using assms unfolding simple_function_def * by auto }
+ ultimately
+ have "\<mu> (f -` {f x} \<inter> space M) = setsum (\<mu>) (?sub (f x))"
+ by (subst measure_finitely_additive) auto }
+ hence "simple_integral f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
+ unfolding simple_integral_def
+ by (subst setsum_Sigma[symmetric],
+ auto intro!: setsum_cong simp: setsum_right_distrib[symmetric])
+ also have "\<dots> = (\<Sum>A\<in>?p ` space M. contents (f`A) * \<mu> A)"
+ proof -
+ have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI)
+ have "(\<lambda>A. (contents (f ` A), A)) ` ?p ` space M
+ = (\<lambda>x. (f x, ?p x)) ` space M"
+ proof safe
+ fix x assume "x \<in> space M"
+ thus "(f x, ?p x) \<in> (\<lambda>A. (contents (f`A), A)) ` ?p ` space M"
+ by (auto intro!: image_eqI[of _ _ "?p x"])
+ qed auto
+ thus ?thesis
+ apply (auto intro!: setsum_reindex_cong[of "\<lambda>A. (contents (f`A), A)"] inj_onI)
+ apply (rule_tac x="xa" in image_eqI)
+ by simp_all
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) simple_integral_add[simp]:
+ assumes "simple_function f" and "simple_function g"
+ shows "simple_integral (\<lambda>x. f x + g x) = simple_integral f + simple_integral g"
+proof -
+ { fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
+ assume "x \<in> space M"
+ hence "(\<lambda>a. f a + g a) ` ?S = {f x + g x}" "f ` ?S = {f x}" "g ` ?S = {g x}"
+ "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M = ?S"
+ by auto }
+ thus ?thesis
+ unfolding
+ simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]]
+ simple_function_partition[OF `simple_function f` `simple_function g`]
+ simple_function_partition[OF `simple_function g` `simple_function f`]
+ apply (subst (3) Int_commute)
+ by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong)
+qed
+
+lemma (in measure_space) simple_integral_setsum[simp]:
+ assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function (f i)"
+ shows "simple_integral (\<lambda>x. \<Sum>i\<in>P. f i x) = (\<Sum>i\<in>P. simple_integral (f i))"
+proof cases
+ assume "finite P"
+ from this assms show ?thesis
+ by induct (auto simp: simple_function_setsum simple_integral_add)
+qed auto
+
+lemma (in measure_space) simple_integral_mult[simp]:
+ assumes "simple_function f"
+ shows "simple_integral (\<lambda>x. c * f x) = c * simple_integral f"
+proof -
+ note mult = simple_function_mult[OF simple_function_const[of c] assms]
+ { fix x let ?S = "f -` {f x} \<inter> (\<lambda>x. c * f x) -` {c * f x} \<inter> space M"
+ assume "x \<in> space M"
+ hence "(\<lambda>x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}"
+ by auto }
+ thus ?thesis
+ unfolding simple_function_partition[OF mult assms]
+ simple_function_partition[OF assms mult]
+ by (auto simp: setsum_right_distrib field_simps intro!: setsum_cong)
+qed
+
+lemma (in measure_space) simple_integral_mono:
+ assumes "simple_function f" and "simple_function g"
+ and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
+ shows "simple_integral f \<le> simple_integral g"
+ unfolding
+ simple_function_partition[OF `simple_function f` `simple_function g`]
+ simple_function_partition[OF `simple_function g` `simple_function f`]
+ apply (subst Int_commute)
+proof (safe intro!: setsum_mono)
+ fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
+ assume "x \<in> space M"
+ hence "f ` ?S = {f x}" "g ` ?S = {g x}" by auto
+ thus "contents (f`?S) * \<mu> ?S \<le> contents (g`?S) * \<mu> ?S"
+ using mono[OF `x \<in> space M`] by (auto intro!: mult_right_mono)
+qed
+
+lemma (in measure_space) simple_integral_indicator:
+ assumes "A \<in> sets M"
+ assumes "simple_function f"
+ shows "simple_integral (\<lambda>x. f x * indicator A x) =
+ (\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M \<inter> A))"
+proof cases
+ assume "A = space M"
+ moreover hence "simple_integral (\<lambda>x. f x * indicator A x) = simple_integral f"
+ by (auto intro!: simple_integral_cong)
+ moreover have "\<And>X. X \<inter> space M \<inter> space M = X \<inter> space M" by auto
+ ultimately show ?thesis by (simp add: simple_integral_def)
+next
+ assume "A \<noteq> space M"
+ then obtain x where x: "x \<in> space M" "x \<notin> A" using sets_into_space[OF assms(1)] by auto
+ have I: "(\<lambda>x. f x * indicator A x) ` space M = f ` A \<union> {0}" (is "?I ` _ = _")
+ proof safe
+ fix y assume "?I y \<notin> f ` A" hence "y \<notin> A" by auto thus "?I y = 0" by auto
+ next
+ fix y assume "y \<in> A" thus "f y \<in> ?I ` space M"
+ using sets_into_space[OF assms(1)] by (auto intro!: image_eqI[of _ _ y])
+ next
+ show "0 \<in> ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x])
+ qed
+ have *: "simple_integral (\<lambda>x. f x * indicator A x) =
+ (\<Sum>x \<in> f ` space M \<union> {0}. x * \<mu> (f -` {x} \<inter> space M \<inter> A))"
+ unfolding simple_integral_def I
+ proof (rule setsum_mono_zero_cong_left)
+ show "finite (f ` space M \<union> {0})"
+ using assms(2) unfolding simple_function_def by auto
+ show "f ` A \<union> {0} \<subseteq> f`space M \<union> {0}"
+ using sets_into_space[OF assms(1)] by auto
+ have "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}" by (auto simp: image_iff)
+ thus "\<forall>i\<in>f ` space M \<union> {0} - (f ` A \<union> {0}).
+ i * \<mu> (f -` {i} \<inter> space M \<inter> A) = 0" by auto
+ next
+ fix x assume "x \<in> f`A \<union> {0}"
+ hence "x \<noteq> 0 \<Longrightarrow> ?I -` {x} \<inter> space M = f -` {x} \<inter> space M \<inter> A"
+ by (auto simp: indicator_def split: split_if_asm)
+ thus "x * \<mu> (?I -` {x} \<inter> space M) =
+ x * \<mu> (f -` {x} \<inter> space M \<inter> A)" by (cases "x = 0") simp_all
+ qed
+ show ?thesis unfolding *
+ using assms(2) unfolding simple_function_def
+ by (auto intro!: setsum_mono_zero_cong_right)
+qed
+
+lemma (in measure_space) simple_integral_indicator_only[simp]:
+ assumes "A \<in> sets M"
+ shows "simple_integral (indicator A) = \<mu> A"
+proof cases
+ assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
+ thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
+next
+ assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::pinfreal}" by auto
+ thus ?thesis
+ using simple_integral_indicator[OF assms simple_function_const[of 1]]
+ using sets_into_space[OF assms]
+ by (auto intro!: arg_cong[where f="\<mu>"])
+qed
+
+lemma (in measure_space) simple_integral_null_set:
+ assumes "simple_function u" "N \<in> null_sets"
+ shows "simple_integral (\<lambda>x. u x * indicator N x) = 0"
+proof -
+ have "simple_integral (\<lambda>x. u x * indicator N x) \<le>
+ simple_integral (\<lambda>x. \<omega> * indicator N x)"
+ using assms
+ by (safe intro!: simple_integral_mono simple_function_mult simple_function_indicator simple_function_const) simp
+ also have "... = 0" apply(subst simple_integral_mult)
+ using assms(2) by auto
+ finally show ?thesis by auto
+qed
+
+lemma (in measure_space) simple_integral_cong':
+ assumes f: "simple_function f" and g: "simple_function g"
+ and mea: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
+ shows "simple_integral f = simple_integral g"
+proof -
+ let ?h = "\<lambda>h. \<lambda>x. (h x * indicator {x\<in>space M. f x = g x} x
+ + h x * indicator {x\<in>space M. f x \<noteq> g x} x
+ + h x * indicator (-space M) x::pinfreal)"
+ have *:"\<And>h. h = ?h h" unfolding indicator_def apply rule by auto
+ have mea_neq:"{x \<in> space M. f x \<noteq> g x} \<in> sets M" using f g by (auto simp: borel_measurable_simple_function)
+ then have mea_nullset: "{x \<in> space M. f x \<noteq> g x} \<in> null_sets" using mea by auto
+ have h1:"\<And>h::_=>pinfreal. simple_function h \<Longrightarrow>
+ simple_function (\<lambda>x. h x * indicator {x\<in>space M. f x = g x} x)"
+ apply(safe intro!: simple_function_add simple_function_mult simple_function_indicator)
+ using f g by (auto simp: borel_measurable_simple_function)
+ have h2:"\<And>h::_\<Rightarrow>pinfreal. simple_function h \<Longrightarrow>
+ simple_function (\<lambda>x. h x * indicator {x\<in>space M. f x \<noteq> g x} x)"
+ apply(safe intro!: simple_function_add simple_function_mult simple_function_indicator)
+ by(rule mea_neq)
+ have **:"\<And>a b c d e f. a = b \<Longrightarrow> c = d \<Longrightarrow> e = f \<Longrightarrow> a+c+e = b+d+f" by auto
+ note *** = simple_integral_add[OF simple_function_add[OF h1 h2] simple_function_notspace]
+ simple_integral_add[OF h1 h2]
+ show ?thesis apply(subst *[of g]) apply(subst *[of f])
+ unfolding ***[OF f f] ***[OF g g]
+ proof(rule **) case goal1 show ?case apply(rule arg_cong[where f=simple_integral]) apply rule
+ unfolding indicator_def by auto
+ next note * = simple_integral_null_set[OF _ mea_nullset]
+ case goal2 show ?case unfolding *[OF f] *[OF g] ..
+ next case goal3 show ?case apply(rule simple_integral_cong) by auto
+ qed
+qed
+
+section "Continuous posititve integration"
+
+definition (in measure_space)
+ "positive_integral f =
+ (SUP g : {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}. simple_integral g)"
+
+lemma (in measure_space) positive_integral_alt1:
+ "positive_integral f =
+ (SUP g : {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}. simple_integral g)"
+ unfolding positive_integral_def SUPR_def
+proof (safe intro!: arg_cong[where f=Sup])
+ fix g let ?g = "\<lambda>x. if x \<in> space M then g x else f x"
+ assume "simple_function g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
+ hence "?g \<le> f" "simple_function ?g" "simple_integral g = simple_integral ?g"
+ "\<omega> \<notin> g`space M"
+ unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong)
+ thus "simple_integral g \<in> simple_integral ` {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}"
+ by auto
+next
+ fix g assume "simple_function g" "g \<le> f" "\<omega> \<notin> g`space M"
+ hence "simple_function g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>"
+ by (auto simp add: le_fun_def image_iff)
+ thus "simple_integral g \<in> simple_integral ` {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}"
+ by auto
+qed
+
+lemma (in measure_space) positive_integral_alt:
+ "positive_integral f =
+ (SUP g : {g. simple_function g \<and> g \<le> f}. simple_integral g)"
+ apply(rule order_class.antisym) unfolding positive_integral_def
+ apply(rule SUPR_subset) apply blast apply(rule SUPR_mono_lim)
+proof safe fix u assume u:"simple_function u" and uf:"u \<le> f"
+ let ?u = "\<lambda>n x. if u x = \<omega> then Real (real (n::nat)) else u x"
+ have su:"\<And>n. simple_function (?u n)" using simple_function_compose1[OF u] .
+ show "\<exists>b. \<forall>n. b n \<in> {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g ` space M} \<and>
+ (\<lambda>n. simple_integral (b n)) ----> simple_integral u"
+ apply(rule_tac x="?u" in exI, safe) apply(rule su)
+ proof- fix n::nat have "?u n \<le> u" unfolding le_fun_def by auto
+ also note uf finally show "?u n \<le> f" .
+ let ?s = "{x \<in> space M. u x = \<omega>}"
+ show "(\<lambda>n. simple_integral (?u n)) ----> simple_integral u"
+ proof(cases "\<mu> ?s = 0")
+ case True have *:"\<And>n. {x\<in>space M. ?u n x \<noteq> u x} = {x\<in>space M. u x = \<omega>}" by auto
+ have *:"\<And>n. simple_integral (?u n) = simple_integral u"
+ apply(rule simple_integral_cong'[OF su u]) unfolding * True ..
+ show ?thesis unfolding * by auto
+ next case False note m0=this
+ have s:"{x \<in> space M. u x = \<omega>} \<in> sets M" using u by (auto simp: borel_measurable_simple_function)
+ have "\<omega> = simple_integral (\<lambda>x. \<omega> * indicator {x\<in>space M. u x = \<omega>} x)"
+ apply(subst simple_integral_mult) using s
+ unfolding simple_integral_indicator_only[OF s] using False by auto
+ also have "... \<le> simple_integral u"
+ apply (rule simple_integral_mono)
+ apply (rule simple_function_mult)
+ apply (rule simple_function_const)
+ apply(rule ) prefer 3 apply(subst indicator_def)
+ using s u by auto
+ finally have *:"simple_integral u = \<omega>" by auto
+ show ?thesis unfolding * Lim_omega_pos
+ proof safe case goal1
+ from real_arch_simple[of "B / real (\<mu> ?s)"] guess N0 .. note N=this
+ def N \<equiv> "Suc N0" have N:"real N \<ge> B / real (\<mu> ?s)" "N > 0"
+ unfolding N_def using N by auto
+ show ?case apply-apply(rule_tac x=N in exI,safe)
+ proof- case goal1
+ have "Real B \<le> Real (real N) * \<mu> ?s"
+ proof(cases "\<mu> ?s = \<omega>")
+ case True thus ?thesis using `B>0` N by auto
+ next case False
+ have *:"B \<le> real N * real (\<mu> ?s)"
+ using N(1) apply-apply(subst (asm) pos_divide_le_eq)
+ apply rule using m0 False by auto
+ show ?thesis apply(subst Real_real'[THEN sym,OF False])
+ apply(subst pinfreal_times,subst if_P) defer
+ apply(subst pinfreal_less_eq,subst if_P) defer
+ using * N `B>0` by(auto intro: mult_nonneg_nonneg)
+ qed
+ also have "... \<le> Real (real n) * \<mu> ?s"
+ apply(rule mult_right_mono) using goal1 by auto
+ also have "... = simple_integral (\<lambda>x. Real (real n) * indicator ?s x)"
+ apply(subst simple_integral_mult) apply(rule simple_function_indicator[OF s])
+ unfolding simple_integral_indicator_only[OF s] ..
+ also have "... \<le> simple_integral (\<lambda>x. if u x = \<omega> then Real (real n) else u x)"
+ apply(rule simple_integral_mono) apply(rule simple_function_mult)
+ apply(rule simple_function_const)
+ apply(rule simple_function_indicator) apply(rule s su)+ by auto
+ finally show ?case .
+ qed qed qed
+ fix x assume x:"\<omega> = (if u x = \<omega> then Real (real n) else u x)" "x \<in> space M"
+ hence "u x = \<omega>" apply-apply(rule ccontr) by auto
+ hence "\<omega> = Real (real n)" using x by auto
+ thus False by auto
+ qed
+qed
+
+lemma (in measure_space) positive_integral_cong:
+ assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
+ shows "positive_integral f = positive_integral g"
+proof -
+ have "\<And>h. (\<forall>x\<in>space M. h x \<le> f x \<and> h x \<noteq> \<omega>) = (\<forall>x\<in>space M. h x \<le> g x \<and> h x \<noteq> \<omega>)"
+ using assms by auto
+ thus ?thesis unfolding positive_integral_alt1 by auto
+qed
+
+lemma (in measure_space) positive_integral_eq_simple_integral:
+ assumes "simple_function f"
+ shows "positive_integral f = simple_integral f"
+ unfolding positive_integral_alt
+proof (safe intro!: pinfreal_SUPI)
+ fix g assume "simple_function g" "g \<le> f"
+ with assms show "simple_integral g \<le> simple_integral f"
+ by (auto intro!: simple_integral_mono simp: le_fun_def)
+next
+ fix y assume "\<forall>x. x\<in>{g. simple_function g \<and> g \<le> f} \<longrightarrow> simple_integral x \<le> y"
+ with assms show "simple_integral f \<le> y" by auto
+qed
+
+lemma (in measure_space) positive_integral_mono:
+ assumes mono: "\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x"
+ shows "positive_integral u \<le> positive_integral v"
+ unfolding positive_integral_alt1
+proof (safe intro!: SUPR_mono)
+ fix a assume a: "simple_function a" and "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>"
+ with mono have "\<forall>x\<in>space M. a x \<le> v x \<and> a x \<noteq> \<omega>" by fastsimp
+ with a show "\<exists>b\<in>{g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}. simple_integral a \<le> simple_integral b"
+ by (auto intro!: bexI[of _ a])
+qed
+
+lemma (in measure_space) positive_integral_SUP_approx:
+ assumes "f \<up> s"
+ and f: "\<And>i. f i \<in> borel_measurable M"
+ and "simple_function u"
+ and le: "u \<le> s" and real: "\<omega> \<notin> u`space M"
+ shows "simple_integral u \<le> (SUP i. positive_integral (f i))" (is "_ \<le> ?S")
+proof (rule pinfreal_le_mult_one_interval)
+ fix a :: pinfreal assume "0 < a" "a < 1"
+ hence "a \<noteq> 0" by auto
+ let "?B i" = "{x \<in> space M. a * u x \<le> f i x}"
+ have B: "\<And>i. ?B i \<in> sets M"
+ using f `simple_function u` by (auto simp: borel_measurable_simple_function)
+
+ let "?uB i x" = "u x * indicator (?B i) x"
+
+ { fix i have "?B i \<subseteq> ?B (Suc i)"
+ proof safe
+ fix i x assume "a * u x \<le> f i x"
+ also have "\<dots> \<le> f (Suc i) x"
+ using `f \<up> s` unfolding isoton_def le_fun_def by auto
+ finally show "a * u x \<le> f (Suc i) x" .
+ qed }
+ note B_mono = this
+
+ have u: "\<And>x. x \<in> space M \<Longrightarrow> u -` {u x} \<inter> space M \<in> sets M"
+ using `simple_function u` by (auto simp add: simple_function_def)
+
+ { fix i
+ have "(\<lambda>n. (u -` {i} \<inter> space M) \<inter> ?B n) \<up> (u -` {i} \<inter> space M)" using B_mono unfolding isoton_def
+ proof safe
+ fix x assume "x \<in> space M"
+ show "x \<in> (\<Union>i. (u -` {u x} \<inter> space M) \<inter> ?B i)"
+ proof cases
+ assume "u x = 0" thus ?thesis using `x \<in> space M` by simp
+ next
+ assume "u x \<noteq> 0"
+ with `a < 1` real `x \<in> space M`
+ have "a * u x < 1 * u x" by (rule_tac pinfreal_mult_strict_right_mono) (auto simp: image_iff)
+ also have "\<dots> \<le> (SUP i. f i x)" using le `f \<up> s`
+ unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def)
+ finally obtain i where "a * u x < f i x" unfolding SUPR_def
+ by (auto simp add: less_Sup_iff)
+ hence "a * u x \<le> f i x" by auto
+ thus ?thesis using `x \<in> space M` by auto
+ qed
+ qed auto }
+ note measure_conv = measure_up[OF u Int[OF u B] this]
+
+ have "simple_integral u = (SUP i. simple_integral (?uB i))"
+ unfolding simple_integral_indicator[OF B `simple_function u`]
+ proof (subst SUPR_pinfreal_setsum, safe)
+ fix x n assume "x \<in> space M"
+ have "\<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f n x})
+ \<le> \<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f (Suc n) x})"
+ using B_mono Int[OF u[OF `x \<in> space M`] B] by (auto intro!: measure_mono)
+ thus "u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B n)
+ \<le> u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B (Suc n))"
+ by (auto intro: mult_left_mono)
+ next
+ show "simple_integral u =
+ (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (u -` {i} \<inter> space M \<inter> ?B n))"
+ using measure_conv unfolding simple_integral_def isoton_def
+ by (auto intro!: setsum_cong simp: pinfreal_SUP_cmult)
+ qed
+ moreover
+ have "a * (SUP i. simple_integral (?uB i)) \<le> ?S"
+ unfolding pinfreal_SUP_cmult[symmetric]
+ proof (safe intro!: SUP_mono)
+ fix i
+ have "a * simple_integral (?uB i) = simple_integral (\<lambda>x. a * ?uB i x)"
+ using B `simple_function u`
+ by (subst simple_integral_mult[symmetric]) (auto simp: field_simps)
+ also have "\<dots> \<le> positive_integral (f i)"
+ proof -
+ have "\<And>x. a * ?uB i x \<le> f i x" unfolding indicator_def by auto
+ hence *: "simple_function (\<lambda>x. a * ?uB i x)" using B assms(3)
+ by (auto intro!: simple_integral_mono)
+ show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric]
+ by (auto intro!: positive_integral_mono simp: indicator_def)
+ qed
+ finally show "a * simple_integral (?uB i) \<le> positive_integral (f i)"
+ by auto
+ qed
+ ultimately show "a * simple_integral u \<le> ?S" by simp
+qed
+
+text {* Beppo-Levi monotone convergence theorem *}
+lemma (in measure_space) positive_integral_isoton:
+ assumes "f \<up> u" "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>i. positive_integral (f i)) \<up> positive_integral u"
+ unfolding isoton_def
+proof safe
+ fix i show "positive_integral (f i) \<le> positive_integral (f (Suc i))"
+ apply (rule positive_integral_mono)
+ using `f \<up> u` unfolding isoton_def le_fun_def by auto
+next
+ have "u \<in> borel_measurable M"
+ using borel_measurable_SUP[of UNIV f] assms by (auto simp: isoton_def)
+ have u: "u = (SUP i. f i)" using `f \<up> u` unfolding isoton_def by auto
+
+ show "(SUP i. positive_integral (f i)) = positive_integral u"
+ proof (rule antisym)
+ from `f \<up> u`[THEN isoton_Sup, unfolded le_fun_def]
+ show "(SUP j. positive_integral (f j)) \<le> positive_integral u"
+ by (auto intro!: SUP_leI positive_integral_mono)
+ next
+ show "positive_integral u \<le> (SUP i. positive_integral (f i))"
+ unfolding positive_integral_def[of u]
+ by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms])
+ qed
+qed
+
+lemma (in measure_space) SUP_simple_integral_sequences:
+ assumes f: "f \<up> u" "\<And>i. simple_function (f i)"
+ and g: "g \<up> u" "\<And>i. simple_function (g i)"
+ shows "(SUP i. simple_integral (f i)) = (SUP i. simple_integral (g i))"
+ (is "SUPR _ ?F = SUPR _ ?G")
+proof -
+ have "(SUP i. ?F i) = (SUP i. positive_integral (f i))"
+ using assms by (simp add: positive_integral_eq_simple_integral)
+ also have "\<dots> = positive_integral u"
+ using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]]
+ unfolding isoton_def by simp
+ also have "\<dots> = (SUP i. positive_integral (g i))"
+ using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]]
+ unfolding isoton_def by simp
+ also have "\<dots> = (SUP i. ?G i)"
+ using assms by (simp add: positive_integral_eq_simple_integral)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) positive_integral_const[simp]:
+ "positive_integral (\<lambda>x. c) = c * \<mu> (space M)"
+ by (subst positive_integral_eq_simple_integral) auto
+
+lemma (in measure_space) positive_integral_isoton_simple:
+ assumes "f \<up> u" and e: "\<And>i. simple_function (f i)"
+ shows "(\<lambda>i. simple_integral (f i)) \<up> positive_integral u"
+ using positive_integral_isoton[OF `f \<up> u` e(1)[THEN borel_measurable_simple_function]]
+ unfolding positive_integral_eq_simple_integral[OF e] .
+
+lemma (in measure_space) positive_integral_linear:
+ assumes f: "f \<in> borel_measurable M"
+ and g: "g \<in> borel_measurable M"
+ shows "positive_integral (\<lambda>x. a * f x + g x) =
+ a * positive_integral f + positive_integral g"
+ (is "positive_integral ?L = _")
+proof -
+ from borel_measurable_implies_simple_function_sequence'[OF f] guess u .
+ note u = this positive_integral_isoton_simple[OF this(1-2)]
+ from borel_measurable_implies_simple_function_sequence'[OF g] guess v .
+ note v = this positive_integral_isoton_simple[OF this(1-2)]
+ let "?L' i x" = "a * u i x + v i x"
+
+ have "?L \<in> borel_measurable M"
+ using assms by simp
+ from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
+ note positive_integral_isoton_simple[OF this(1-2)] and l = this
+ moreover have
+ "(SUP i. simple_integral (l i)) = (SUP i. simple_integral (?L' i))"
+ proof (rule SUP_simple_integral_sequences[OF l(1-2)])
+ show "?L' \<up> ?L" "\<And>i. simple_function (?L' i)"
+ using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right)
+ qed
+ moreover from u v have L'_isoton:
+ "(\<lambda>i. simple_integral (?L' i)) \<up> a * positive_integral f + positive_integral g"
+ by (simp add: isoton_add isoton_cmult_right)
+ ultimately show ?thesis by (simp add: isoton_def)
+qed
+
+lemma (in measure_space) positive_integral_cmult:
+ assumes "f \<in> borel_measurable M"
+ shows "positive_integral (\<lambda>x. c * f x) = c * positive_integral f"
+ using positive_integral_linear[OF assms, of "\<lambda>x. 0" c] by auto
+
+lemma (in measure_space) positive_integral_indicator[simp]:
+ "A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. indicator A x) = \<mu> A"
+by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator)
+
+lemma (in measure_space) positive_integral_cmult_indicator:
+ "A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. c * indicator A x) = c * \<mu> A"
+by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator)
+
+lemma (in measure_space) positive_integral_add:
+ assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ shows "positive_integral (\<lambda>x. f x + g x) = positive_integral f + positive_integral g"
+ using positive_integral_linear[OF assms, of 1] by simp
+
+lemma (in measure_space) positive_integral_setsum:
+ assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "positive_integral (\<lambda>x. \<Sum>i\<in>P. f i x) = (\<Sum>i\<in>P. positive_integral (f i))"
+proof cases
+ assume "finite P"
+ from this assms show ?thesis
+ proof induct
+ case (insert i P)
+ have "f i \<in> borel_measurable M"
+ "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M"
+ using insert by (auto intro!: borel_measurable_pinfreal_setsum)
+ from positive_integral_add[OF this]
+ show ?case using insert by auto
+ qed simp
+qed simp
+
+lemma (in measure_space) positive_integral_diff:
+ assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
+ and fin: "positive_integral g \<noteq> \<omega>"
+ and mono: "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x"
+ shows "positive_integral (\<lambda>x. f x - g x) = positive_integral f - positive_integral g"
+proof -
+ have borel: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+ using f g by (rule borel_measurable_pinfreal_diff)
+ have "positive_integral (\<lambda>x. f x - g x) + positive_integral g =
+ positive_integral f"
+ unfolding positive_integral_add[OF borel g, symmetric]
+ proof (rule positive_integral_cong)
+ fix x assume "x \<in> space M"
+ from mono[OF this] show "f x - g x + g x = f x"
+ by (cases "f x", cases "g x", simp, simp, cases "g x", auto)
+ qed
+ with mono show ?thesis
+ by (subst minus_pinfreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono)
+qed
+
+lemma (in measure_space) positive_integral_psuminf:
+ assumes "\<And>i. f i \<in> borel_measurable M"
+ shows "positive_integral (\<lambda>x. \<Sum>\<^isub>\<infinity> i. f i x) = (\<Sum>\<^isub>\<infinity> i. positive_integral (f i))"
+proof -
+ have "(\<lambda>i. positive_integral (\<lambda>x. \<Sum>i<i. f i x)) \<up> positive_integral (\<lambda>x. \<Sum>\<^isub>\<infinity>i. f i x)"
+ by (rule positive_integral_isoton)
+ (auto intro!: borel_measurable_pinfreal_setsum assms positive_integral_mono
+ arg_cong[where f=Sup]
+ simp: isoton_def le_fun_def psuminf_def expand_fun_eq SUPR_def Sup_fun_def)
+ thus ?thesis
+ by (auto simp: isoton_def psuminf_def positive_integral_setsum[OF assms])
+qed
+
+text {* Fatou's lemma: convergence theorem on limes inferior *}
+lemma (in measure_space) positive_integral_lim_INF:
+ fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> pinfreal"
+ assumes "\<And>i. u i \<in> borel_measurable M"
+ shows "positive_integral (SUP n. INF m. u (m + n)) \<le>
+ (SUP n. INF m. positive_integral (u (m + n)))"
+proof -
+ have "(SUP n. INF m. u (m + n)) \<in> borel_measurable M"
+ by (auto intro!: borel_measurable_SUP borel_measurable_INF assms)
+
+ have "(\<lambda>n. INF m. u (m + n)) \<up> (SUP n. INF m. u (m + n))"
+ proof (unfold isoton_def, safe)
+ fix i show "(INF m. u (m + i)) \<le> (INF m. u (m + Suc i))"
+ by (rule INF_mono[where N=Suc]) simp
+ qed
+ from positive_integral_isoton[OF this] assms
+ have "positive_integral (SUP n. INF m. u (m + n)) =
+ (SUP n. positive_integral (INF m. u (m + n)))"
+ unfolding isoton_def by (simp add: borel_measurable_INF)
+ also have "\<dots> \<le> (SUP n. INF m. positive_integral (u (m + n)))"
+ by (auto intro!: SUP_mono[where N="\<lambda>x. x"] INFI_bound positive_integral_mono INF_leI simp: INFI_fun_expand)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) measure_space_density:
+ assumes borel: "u \<in> borel_measurable M"
+ shows "measure_space M (\<lambda>A. positive_integral (\<lambda>x. u x * indicator A x))" (is "measure_space M ?v")
+proof
+ show "?v {} = 0" by simp
+ show "countably_additive M ?v"
+ unfolding countably_additive_def
+ proof safe
+ fix A :: "nat \<Rightarrow> 'a set"
+ assume "range A \<subseteq> sets M"
+ hence "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
+ using borel by (auto intro: borel_measurable_indicator)
+ moreover assume "disjoint_family A"
+ note psuminf_indicator[OF this]
+ ultimately show "(\<Sum>\<^isub>\<infinity>n. ?v (A n)) = ?v (\<Union>x. A x)"
+ by (simp add: positive_integral_psuminf[symmetric])
+ qed
+qed
+
+lemma (in measure_space) positive_integral_null_set:
+ assumes borel: "u \<in> borel_measurable M" and "N \<in> null_sets"
+ shows "positive_integral (\<lambda>x. u x * indicator N x) = 0" (is "?I = 0")
+proof -
+ have "N \<in> sets M" using `N \<in> null_sets` by auto
+ have "(\<lambda>i x. min (of_nat i) (u x) * indicator N x) \<up> (\<lambda>x. u x * indicator N x)"
+ unfolding isoton_fun_expand
+ proof (safe intro!: isoton_cmult_left, unfold isoton_def, safe)
+ fix j i show "min (of_nat j) (u i) \<le> min (of_nat (Suc j)) (u i)"
+ by (rule min_max.inf_mono) auto
+ next
+ fix i show "(SUP j. min (of_nat j) (u i)) = u i"
+ proof (cases "u i")
+ case infinite
+ moreover hence "\<And>j. min (of_nat j) (u i) = of_nat j"
+ by (auto simp: min_def)
+ ultimately show ?thesis by (simp add: Sup_\<omega>)
+ next
+ case (preal r)
+ obtain j where "r \<le> of_nat j" using ex_le_of_nat ..
+ hence "u i \<le> of_nat j" using preal by (auto simp: real_of_nat_def)
+ show ?thesis
+ proof (rule pinfreal_SUPI)
+ fix y assume "\<And>j. j \<in> UNIV \<Longrightarrow> min (of_nat j) (u i) \<le> y"
+ note this[of j]
+ moreover have "min (of_nat j) (u i) = u i"
+ using `u i \<le> of_nat j` by (auto simp: min_def)
+ ultimately show "u i \<le> y" by simp
+ qed simp
+ qed
+ qed
+ from positive_integral_isoton[OF this]
+ have "?I = (SUP i. positive_integral (\<lambda>x. min (of_nat i) (u x) * indicator N x))"
+ unfolding isoton_def using borel `N \<in> sets M` by (simp add: borel_measurable_indicator)
+ also have "\<dots> \<le> (SUP i. positive_integral (\<lambda>x. of_nat i * indicator N x))"
+ proof (rule SUP_mono, rule positive_integral_mono)
+ fix x i show "min (of_nat i) (u x) * indicator N x \<le> of_nat i * indicator N x"
+ by (cases "x \<in> N") auto
+ qed
+ also have "\<dots> = 0"
+ using `N \<in> null_sets` by (simp add: positive_integral_cmult_indicator)
+ finally show ?thesis by simp
+qed
+
+lemma (in measure_space) positive_integral_Markov_inequality:
+ assumes borel: "u \<in> borel_measurable M" and "A \<in> sets M" and c: "c \<noteq> \<omega>"
+ shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * positive_integral (\<lambda>x. u x * indicator A x)"
+ (is "\<mu> ?A \<le> _ * ?PI")
+proof -
+ have "?A \<in> sets M"
+ using `A \<in> sets M` borel by auto
+ hence "\<mu> ?A = positive_integral (\<lambda>x. indicator ?A x)"
+ using positive_integral_indicator by simp
+ also have "\<dots> \<le> positive_integral (\<lambda>x. c * (u x * indicator A x))"
+ proof (rule positive_integral_mono)
+ fix x assume "x \<in> space M"
+ show "indicator ?A x \<le> c * (u x * indicator A x)"
+ by (cases "x \<in> ?A") auto
+ qed
+ also have "\<dots> = c * positive_integral (\<lambda>x. u x * indicator A x)"
+ using assms
+ by (auto intro!: positive_integral_cmult borel_measurable_indicator)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) positive_integral_0_iff:
+ assumes borel: "u \<in> borel_measurable M"
+ shows "positive_integral u = 0 \<longleftrightarrow> \<mu> {x\<in>space M. u x \<noteq> 0} = 0"
+ (is "_ \<longleftrightarrow> \<mu> ?A = 0")
+proof -
+ have A: "?A \<in> sets M" using borel by auto
+ have u: "positive_integral (\<lambda>x. u x * indicator ?A x) = positive_integral u"
+ by (auto intro!: positive_integral_cong simp: indicator_def)
+
+ show ?thesis
+ proof
+ assume "\<mu> ?A = 0"
+ hence "?A \<in> null_sets" using `?A \<in> sets M` by auto
+ from positive_integral_null_set[OF borel this]
+ have "0 = positive_integral (\<lambda>x. u x * indicator ?A x)" by simp
+ thus "positive_integral u = 0" unfolding u by simp
+ next
+ assume *: "positive_integral u = 0"
+ let "?M n" = "{x \<in> space M. 1 \<le> of_nat n * u x}"
+ have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
+ proof -
+ { fix n
+ from positive_integral_Markov_inequality[OF borel `?A \<in> sets M`, of "of_nat n"]
+ have "\<mu> (?M n \<inter> ?A) = 0" unfolding * u by simp }
+ thus ?thesis by simp
+ qed
+ also have "\<dots> = \<mu> (\<Union>n. ?M n \<inter> ?A)"
+ proof (safe intro!: continuity_from_below)
+ fix n show "?M n \<inter> ?A \<in> sets M"
+ using borel by (auto intro!: Int)
+ next
+ fix n x assume "1 \<le> of_nat n * u x"
+ also have "\<dots> \<le> of_nat (Suc n) * u x"
+ by (subst (1 2) mult_commute) (auto intro!: pinfreal_mult_cancel)
+ finally show "1 \<le> of_nat (Suc n) * u x" .
+ qed
+ also have "\<dots> = \<mu> ?A"
+ proof (safe intro!: arg_cong[where f="\<mu>"])
+ fix x assume "u x \<noteq> 0" and [simp, intro]: "x \<in> space M"
+ show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
+ proof (cases "u x")
+ case (preal r)
+ obtain j where "1 / r \<le> of_nat j" using ex_le_of_nat ..
+ hence "1 / r * r \<le> of_nat j * r" using preal unfolding mult_le_cancel_right by auto
+ hence "1 \<le> of_nat j * r" using preal `u x \<noteq> 0` by auto
+ thus ?thesis using `u x \<noteq> 0` preal by (auto simp: real_of_nat_def[symmetric])
+ qed auto
+ qed
+ finally show "\<mu> ?A = 0" by simp
+ qed
+qed
+
+lemma (in measure_space) positive_integral_cong_on_null_sets:
+ assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
+ and measure: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
+ shows "positive_integral f = positive_integral g"
+proof -
+ let ?N = "{x\<in>space M. f x \<noteq> g x}" and ?E = "{x\<in>space M. f x = g x}"
+ let "?A h x" = "h x * indicator ?E x :: pinfreal"
+ let "?B h x" = "h x * indicator ?N x :: pinfreal"
+
+ have A: "positive_integral (?A f) = positive_integral (?A g)"
+ by (auto intro!: positive_integral_cong simp: indicator_def)
+
+ have [intro]: "?N \<in> sets M" "?E \<in> sets M" using f g by auto
+ hence "?N \<in> null_sets" using measure by auto
+ hence B: "positive_integral (?B f) = positive_integral (?B g)"
+ using f g by (simp add: positive_integral_null_set)
+
+ have "positive_integral f = positive_integral (\<lambda>x. ?A f x + ?B f x)"
+ by (auto intro!: positive_integral_cong simp: indicator_def)
+ also have "\<dots> = positive_integral (?A f) + positive_integral (?B f)"
+ using f g by (auto intro!: positive_integral_add borel_measurable_indicator)
+ also have "\<dots> = positive_integral (\<lambda>x. ?A g x + ?B g x)"
+ unfolding A B using f g by (auto intro!: positive_integral_add[symmetric] borel_measurable_indicator)
+ also have "\<dots> = positive_integral g"
+ by (auto intro!: positive_integral_cong simp: indicator_def)
+ finally show ?thesis by simp
+qed
+
+section "Lebesgue Integral"
+
+definition (in measure_space) integrable where
+ "integrable f \<longleftrightarrow> f \<in> borel_measurable M \<and>
+ positive_integral (\<lambda>x. Real (f x)) \<noteq> \<omega> \<and>
+ positive_integral (\<lambda>x. Real (- f x)) \<noteq> \<omega>"
+
+lemma (in measure_space) integrableD[dest]:
+ assumes "integrable f"
+ shows "f \<in> borel_measurable M"
+ "positive_integral (\<lambda>x. Real (f x)) \<noteq> \<omega>"
+ "positive_integral (\<lambda>x. Real (- f x)) \<noteq> \<omega>"
+ using assms unfolding integrable_def by auto
+
+definition (in measure_space) integral where
+ "integral f =
+ real (positive_integral (\<lambda>x. Real (f x))) -
+ real (positive_integral (\<lambda>x. Real (- f x)))"
+
+lemma (in measure_space) integral_cong:
+ assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
+ shows "integral f = integral g"
+ using assms by (simp cong: positive_integral_cong add: integral_def)
+
+lemma (in measure_space) integrable_cong:
+ "(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable f \<longleftrightarrow> integrable g"
+ by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
+
+lemma (in measure_space) integral_eq_positive_integral:
+ assumes "\<And>x. 0 \<le> f x"
+ shows "integral f = real (positive_integral (\<lambda>x. Real (f x)))"
+proof -
+ have "\<And>x. Real (- f x) = 0" using assms by simp
+ thus ?thesis by (simp del: Real_eq_0 add: integral_def)
+qed
+
+lemma (in measure_space) integral_minus[intro, simp]:
+ assumes "integrable f"
+ shows "integrable (\<lambda>x. - f x)" "integral (\<lambda>x. - f x) = - integral f"
+ using assms by (auto simp: integrable_def integral_def)
+
+lemma (in measure_space) integral_of_positive_diff:
+ assumes integrable: "integrable u" "integrable v"
+ and f_def: "\<And>x. f x = u x - v x" and pos: "\<And>x. 0 \<le> u x" "\<And>x. 0 \<le> v x"
+ shows "integrable f" and "integral f = integral u - integral v"
+proof -
+ let ?PI = positive_integral
+ let "?f x" = "Real (f x)"
+ let "?mf x" = "Real (- f x)"
+ let "?u x" = "Real (u x)"
+ let "?v x" = "Real (v x)"
+
+ from borel_measurable_diff[of u v] integrable
+ have f_borel: "?f \<in> borel_measurable M" and
+ mf_borel: "?mf \<in> borel_measurable M" and
+ v_borel: "?v \<in> borel_measurable M" and
+ u_borel: "?u \<in> borel_measurable M" and
+ "f \<in> borel_measurable M"
+ by (auto simp: f_def[symmetric] integrable_def)
+
+ have "?PI (\<lambda>x. Real (u x - v x)) \<le> ?PI ?u"
+ using pos by (auto intro!: positive_integral_mono)
+ moreover have "?PI (\<lambda>x. Real (v x - u x)) \<le> ?PI ?v"
+ using pos by (auto intro!: positive_integral_mono)
+ ultimately show f: "integrable f"
+ using `integrable u` `integrable v` `f \<in> borel_measurable M`
+ by (auto simp: integrable_def f_def)
+ hence mf: "integrable (\<lambda>x. - f x)" ..
+
+ have *: "\<And>x. Real (- v x) = 0" "\<And>x. Real (- u x) = 0"
+ using pos by auto
+
+ have "\<And>x. ?u x + ?mf x = ?v x + ?f x"
+ unfolding f_def using pos by simp
+ hence "?PI (\<lambda>x. ?u x + ?mf x) = ?PI (\<lambda>x. ?v x + ?f x)" by simp
+ hence "real (?PI ?u + ?PI ?mf) = real (?PI ?v + ?PI ?f)"
+ using positive_integral_add[OF u_borel mf_borel]
+ using positive_integral_add[OF v_borel f_borel]
+ by auto
+ then show "integral f = integral u - integral v"
+ using f mf `integrable u` `integrable v`
+ unfolding integral_def integrable_def *
+ by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?v", cases "?PI ?u")
+ (auto simp add: field_simps)
+qed
+
+lemma (in measure_space) integral_linear:
+ assumes "integrable f" "integrable g" and "0 \<le> a"
+ shows "integrable (\<lambda>t. a * f t + g t)"
+ and "integral (\<lambda>t. a * f t + g t) = a * integral f + integral g"
+proof -
+ let ?PI = positive_integral
+ let "?f x" = "Real (f x)"
+ let "?g x" = "Real (g x)"
+ let "?mf x" = "Real (- f x)"
+ let "?mg x" = "Real (- g x)"
+ let "?p t" = "max 0 (a * f t) + max 0 (g t)"
+ let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)"
+
+ have pos: "?f \<in> borel_measurable M" "?g \<in> borel_measurable M"
+ and neg: "?mf \<in> borel_measurable M" "?mg \<in> borel_measurable M"
+ and p: "?p \<in> borel_measurable M"
+ and n: "?n \<in> borel_measurable M"
+ using assms by (simp_all add: integrable_def)
+
+ have *: "\<And>x. Real (?p x) = Real a * ?f x + ?g x"
+ "\<And>x. Real (?n x) = Real a * ?mf x + ?mg x"
+ "\<And>x. Real (- ?p x) = 0"
+ "\<And>x. Real (- ?n x) = 0"
+ using `0 \<le> a` by (auto simp: max_def min_def zero_le_mult_iff mult_le_0_iff add_nonpos_nonpos)
+
+ note linear =
+ positive_integral_linear[OF pos]
+ positive_integral_linear[OF neg]
+
+ have "integrable ?p" "integrable ?n"
+ "\<And>t. a * f t + g t = ?p t - ?n t" "\<And>t. 0 \<le> ?p t" "\<And>t. 0 \<le> ?n t"
+ using assms p n unfolding integrable_def * linear by auto
+ note diff = integral_of_positive_diff[OF this]
+
+ show "integrable (\<lambda>t. a * f t + g t)" by (rule diff)
+
+ from assms show "integral (\<lambda>t. a * f t + g t) = a * integral f + integral g"
+ unfolding diff(2) unfolding integral_def * linear integrable_def
+ by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg")
+ (auto simp add: field_simps zero_le_mult_iff)
+qed
+
+lemma (in measure_space) integral_add[simp, intro]:
+ assumes "integrable f" "integrable g"
+ shows "integrable (\<lambda>t. f t + g t)"
+ and "integral (\<lambda>t. f t + g t) = integral f + integral g"
+ using assms integral_linear[where a=1] by auto
+
+lemma (in measure_space) integral_zero[simp, intro]:
+ shows "integrable (\<lambda>x. 0)"
+ and "integral (\<lambda>x. 0) = 0"
+ unfolding integrable_def integral_def
+ by (auto simp add: borel_measurable_const)
+
+lemma (in measure_space) integral_cmult[simp, intro]:
+ assumes "integrable f"
+ shows "integrable (\<lambda>t. a * f t)" (is ?P)
+ and "integral (\<lambda>t. a * f t) = a * integral f" (is ?I)
+proof -
+ have "integrable (\<lambda>t. a * f t) \<and> integral (\<lambda>t. a * f t) = a * integral f"
+ proof (cases rule: le_cases)
+ assume "0 \<le> a" show ?thesis
+ using integral_linear[OF assms integral_zero(1) `0 \<le> a`]
+ by (simp add: integral_zero)
+ next
+ assume "a \<le> 0" hence "0 \<le> - a" by auto
+ have *: "\<And>t. - a * t + 0 = (-a) * t" by simp
+ show ?thesis using integral_linear[OF assms integral_zero(1) `0 \<le> - a`]
+ integral_minus(1)[of "\<lambda>t. - a * f t"]
+ unfolding * integral_zero by simp
+ qed
+ thus ?P ?I by auto
+qed
+
+lemma (in measure_space) integral_mono:
+ assumes fg: "integrable f" "integrable g"
+ and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
+ shows "integral f \<le> integral g"
+ using fg unfolding integral_def integrable_def diff_minus
+proof (safe intro!: add_mono real_of_pinfreal_mono le_imp_neg_le positive_integral_mono)
+ fix x assume "x \<in> space M" from mono[OF this]
+ show "Real (f x) \<le> Real (g x)" "Real (- g x) \<le> Real (- f x)" by auto
+qed
+
+lemma (in measure_space) integral_diff[simp, intro]:
+ assumes f: "integrable f" and g: "integrable g"
+ shows "integrable (\<lambda>t. f t - g t)"
+ and "integral (\<lambda>t. f t - g t) = integral f - integral g"
+ using integral_add[OF f integral_minus(1)[OF g]]
+ unfolding diff_minus integral_minus(2)[OF g]
+ by auto
+
+lemma (in measure_space) integral_indicator[simp, intro]:
+ assumes "a \<in> sets M" and "\<mu> a \<noteq> \<omega>"
+ shows "integral (indicator a) = real (\<mu> a)" (is ?int)
+ and "integrable (indicator a)" (is ?able)
+proof -
+ have *:
+ "\<And>A x. Real (indicator A x) = indicator A x"
+ "\<And>A x. Real (- indicator A x) = 0" unfolding indicator_def by auto
+ show ?int ?able
+ using assms unfolding integral_def integrable_def
+ by (auto simp: * positive_integral_indicator borel_measurable_indicator)
+qed
+
+lemma (in measure_space) integral_cmul_indicator:
+ assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<omega>"
+ shows "integrable (\<lambda>x. c * indicator A x)" (is ?P)
+ and "integral (\<lambda>x. c * indicator A x) = c * real (\<mu> A)" (is ?I)
+proof -
+ show ?P
+ proof (cases "c = 0")
+ case False with assms show ?thesis by simp
+ qed simp
+
+ show ?I
+ proof (cases "c = 0")
+ case False with assms show ?thesis by simp
+ qed simp
+qed
+
+lemma (in measure_space) integral_setsum[simp, intro]:
+ assumes "\<And>n. n \<in> S \<Longrightarrow> integrable (f n)"
+ shows "integral (\<lambda>x. \<Sum> i \<in> S. f i x) = (\<Sum> i \<in> S. integral (f i))" (is "?int S")
+ and "integrable (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I S")
+proof -
+ have "?int S \<and> ?I S"
+ proof (cases "finite S")
+ assume "finite S"
+ from this assms show ?thesis by (induct S) simp_all
+ qed simp
+ thus "?int S" and "?I S" by auto
+qed
+
+lemma (in measure_space) integrable_abs:
+ assumes "integrable f"
+ shows "integrable (\<lambda> x. \<bar>f x\<bar>)"
+proof -
+ have *:
+ "\<And>x. Real \<bar>f x\<bar> = Real (f x) + Real (- f x)"
+ "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
+ have abs: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" and
+ f: "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
+ "(\<lambda>x. Real (- f x)) \<in> borel_measurable M"
+ using assms unfolding integrable_def by auto
+ from abs assms show ?thesis unfolding integrable_def *
+ using positive_integral_linear[OF f, of 1] by simp
+qed
+
+lemma (in measure_space) integrable_bound:
+ assumes "integrable f"
+ and f: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+ "\<And>x. x \<in> space M \<Longrightarrow> \<bar>g x\<bar> \<le> f x"
+ assumes borel: "g \<in> borel_measurable M"
+ shows "integrable g"
+proof -
+ have "positive_integral (\<lambda>x. Real (g x)) \<le> positive_integral (\<lambda>x. Real \<bar>g x\<bar>)"
+ by (auto intro!: positive_integral_mono)
+ also have "\<dots> \<le> positive_integral (\<lambda>x. Real (f x))"
+ using f by (auto intro!: positive_integral_mono)
+ also have "\<dots> < \<omega>"
+ using `integrable f` unfolding integrable_def by (auto simp: pinfreal_less_\<omega>)
+ finally have pos: "positive_integral (\<lambda>x. Real (g x)) < \<omega>" .
+
+ have "positive_integral (\<lambda>x. Real (- g x)) \<le> positive_integral (\<lambda>x. Real (\<bar>g x\<bar>))"
+ by (auto intro!: positive_integral_mono)
+ also have "\<dots> \<le> positive_integral (\<lambda>x. Real (f x))"
+ using f by (auto intro!: positive_integral_mono)
+ also have "\<dots> < \<omega>"
+ using `integrable f` unfolding integrable_def by (auto simp: pinfreal_less_\<omega>)
+ finally have neg: "positive_integral (\<lambda>x. Real (- g x)) < \<omega>" .
+
+ from neg pos borel show ?thesis
+ unfolding integrable_def by auto
+qed
+
+lemma (in measure_space) integrable_abs_iff:
+ "f \<in> borel_measurable M \<Longrightarrow> integrable (\<lambda> x. \<bar>f x\<bar>) \<longleftrightarrow> integrable f"
+ by (auto intro!: integrable_bound[where g=f] integrable_abs)
+
+lemma (in measure_space) integrable_max:
+ assumes int: "integrable f" "integrable g"
+ shows "integrable (\<lambda> x. max (f x) (g x))"
+proof (rule integrable_bound)
+ show "integrable (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
+ using int by (simp add: integrable_abs)
+ show "(\<lambda>x. max (f x) (g x)) \<in> borel_measurable M"
+ using int unfolding integrable_def by auto
+next
+ fix x assume "x \<in> space M"
+ show "0 \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" "\<bar>max (f x) (g x)\<bar> \<le> \<bar>f x\<bar> + \<bar>g x\<bar>"
+ by auto
+qed
+
+lemma (in measure_space) integrable_min:
+ assumes int: "integrable f" "integrable g"
+ shows "integrable (\<lambda> x. min (f x) (g x))"
+proof (rule integrable_bound)
+ show "integrable (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)"
+ using int by (simp add: integrable_abs)
+ show "(\<lambda>x. min (f x) (g x)) \<in> borel_measurable M"
+ using int unfolding integrable_def by auto
+next
+ fix x assume "x \<in> space M"
+ show "0 \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" "\<bar>min (f x) (g x)\<bar> \<le> \<bar>f x\<bar> + \<bar>g x\<bar>"
+ by auto
+qed
+
+lemma (in measure_space) integral_triangle_inequality:
+ assumes "integrable f"
+ shows "\<bar>integral f\<bar> \<le> integral (\<lambda>x. \<bar>f x\<bar>)"
+proof -
+ have "\<bar>integral f\<bar> = max (integral f) (- integral f)" by auto
+ also have "\<dots> \<le> integral (\<lambda>x. \<bar>f x\<bar>)"
+ using assms integral_minus(2)[of f, symmetric]
+ by (auto intro!: integral_mono integrable_abs simp del: integral_minus)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) integral_positive:
+ assumes "integrable f" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+ shows "0 \<le> integral f"
+proof -
+ have "0 = integral (\<lambda>x. 0)" by (auto simp: integral_zero)
+ also have "\<dots> \<le> integral f"
+ using assms by (rule integral_mono[OF integral_zero(1)])
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) integral_monotone_convergence_pos:
+ assumes i: "\<And>i. integrable (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
+ and pos: "\<And>x i. 0 \<le> f i x"
+ and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
+ and ilim: "(\<lambda>i. integral (f i)) ----> x"
+ shows "integrable u"
+ and "integral u = x"
+proof -
+ { fix x have "0 \<le> u x"
+ using mono pos[of 0 x] incseq_le[OF _ lim, of x 0]
+ by (simp add: mono_def incseq_def) }
+ note pos_u = this
+
+ hence [simp]: "\<And>i x. Real (- f i x) = 0" "\<And>x. Real (- u x) = 0"
+ using pos by auto
+
+ have SUP_F: "\<And>x. (SUP n. Real (f n x)) = Real (u x)"
+ using mono pos pos_u lim by (rule SUP_eq_LIMSEQ[THEN iffD2])
+
+ have borel_f: "\<And>i. (\<lambda>x. Real (f i x)) \<in> borel_measurable M"
+ using i unfolding integrable_def by auto
+ hence "(SUP i. (\<lambda>x. Real (f i x))) \<in> borel_measurable M"
+ by auto
+ hence borel_u: "u \<in> borel_measurable M"
+ using pos_u by (auto simp: borel_measurable_Real_eq SUPR_fun_expand SUP_F)
+
+ have integral_eq: "\<And>n. positive_integral (\<lambda>x. Real (f n x)) = Real (integral (f n))"
+ using i unfolding integral_def integrable_def by (auto simp: Real_real)
+
+ have pos_integral: "\<And>n. 0 \<le> integral (f n)"
+ using pos i by (auto simp: integral_positive)
+ hence "0 \<le> x"
+ using LIMSEQ_le_const[OF ilim, of 0] by auto
+
+ have "(\<lambda>i. positive_integral (\<lambda>x. Real (f i x))) \<up> positive_integral (\<lambda>x. Real (u x))"
+ proof (rule positive_integral_isoton)
+ from SUP_F mono pos
+ show "(\<lambda>i x. Real (f i x)) \<up> (\<lambda>x. Real (u x))"
+ unfolding isoton_fun_expand by (auto simp: isoton_def mono_def)
+ qed (rule borel_f)
+ hence pI: "positive_integral (\<lambda>x. Real (u x)) =
+ (SUP n. positive_integral (\<lambda>x. Real (f n x)))"
+ unfolding isoton_def by simp
+ also have "\<dots> = Real x" unfolding integral_eq
+ proof (rule SUP_eq_LIMSEQ[THEN iffD2])
+ show "mono (\<lambda>n. integral (f n))"
+ using mono i by (auto simp: mono_def intro!: integral_mono)
+ show "\<And>n. 0 \<le> integral (f n)" using pos_integral .
+ show "0 \<le> x" using `0 \<le> x` .
+ show "(\<lambda>n. integral (f n)) ----> x" using ilim .
+ qed
+ finally show "integrable u" "integral u = x" using borel_u `0 \<le> x`
+ unfolding integrable_def integral_def by auto
+qed
+
+lemma (in measure_space) integral_monotone_convergence:
+ assumes f: "\<And>i. integrable (f i)" and "mono f"
+ and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
+ and ilim: "(\<lambda>i. integral (f i)) ----> x"
+ shows "integrable u"
+ and "integral u = x"
+proof -
+ have 1: "\<And>i. integrable (\<lambda>x. f i x - f 0 x)"
+ using f by (auto intro!: integral_diff)
+ have 2: "\<And>x. mono (\<lambda>n. f n x - f 0 x)" using `mono f`
+ unfolding mono_def le_fun_def by auto
+ have 3: "\<And>x n. 0 \<le> f n x - f 0 x" using `mono f`
+ unfolding mono_def le_fun_def by (auto simp: field_simps)
+ have 4: "\<And>x. (\<lambda>i. f i x - f 0 x) ----> u x - f 0 x"
+ using lim by (auto intro!: LIMSEQ_diff)
+ have 5: "(\<lambda>i. integral (\<lambda>x. f i x - f 0 x)) ----> x - integral (f 0)"
+ using f ilim by (auto intro!: LIMSEQ_diff simp: integral_diff)
+ note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5]
+ have "integrable (\<lambda>x. (u x - f 0 x) + f 0 x)"
+ using diff(1) f by (rule integral_add(1))
+ with diff(2) f show "integrable u" "integral u = x"
+ by (auto simp: integral_diff)
+qed
+
+lemma (in measure_space) integral_0_iff:
+ assumes "integrable f"
+ shows "integral (\<lambda>x. \<bar>f x\<bar>) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
+proof -
+ have *: "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto
+ have "integrable (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs)
+ hence "(\<lambda>x. Real (\<bar>f x\<bar>)) \<in> borel_measurable M"
+ "positive_integral (\<lambda>x. Real \<bar>f x\<bar>) \<noteq> \<omega>" unfolding integrable_def by auto
+ from positive_integral_0_iff[OF this(1)] this(2)
+ show ?thesis unfolding integral_def *
+ by (simp add: real_of_pinfreal_eq_0)
+qed
+
+lemma LIMSEQ_max:
+ "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
+ by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
+
+lemma (in sigma_algebra) borel_measurable_LIMSEQ:
+ fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
+ assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
+ and u: "\<And>i. u i \<in> borel_measurable M"
+ shows "u' \<in> borel_measurable M"
+proof -
+ let "?pu x i" = "max (u i x) 0"
+ let "?nu x i" = "max (- u i x) 0"
+
+ { fix x assume x: "x \<in> space M"
+ have "(?pu x) ----> max (u' x) 0"
+ "(?nu x) ----> max (- u' x) 0"
+ using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
+ from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
+ have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
+ "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
+ by (simp_all add: Real_max'[symmetric]) }
+ note eq = this
+
+ have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
+ by auto
+
+ have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M"
+ "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M"
+ using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real)
+ with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space]
+ have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
+ "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M"
+ unfolding SUPR_fun_expand INFI_fun_expand by auto
+ note this[THEN borel_measurable_real]
+ from borel_measurable_diff[OF this]
+ show ?thesis unfolding * .
+qed
+
+lemma (in measure_space) integral_dominated_convergence:
+ assumes u: "\<And>i. integrable (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
+ and w: "integrable w" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> w x"
+ and u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
+ shows "integrable u'"
+ and "(\<lambda>i. integral (\<lambda>x. \<bar>u i x - u' x\<bar>)) ----> 0" (is "?lim_diff")
+ and "(\<lambda>i. integral (u i)) ----> integral u'" (is ?lim)
+proof -
+ { fix x j assume x: "x \<in> space M"
+ from u'[OF x] have "(\<lambda>i. \<bar>u i x\<bar>) ----> \<bar>u' x\<bar>" by (rule LIMSEQ_imp_rabs)
+ from LIMSEQ_le_const2[OF this]
+ have "\<bar>u' x\<bar> \<le> w x" using bound[OF x] by auto }
+ note u'_bound = this
+
+ from u[unfolded integrable_def]
+ have u'_borel: "u' \<in> borel_measurable M"
+ using u' by (blast intro: borel_measurable_LIMSEQ[of u])
+
+ show "integrable u'"
+ proof (rule integrable_bound)
+ show "integrable w" by fact
+ show "u' \<in> borel_measurable M" by fact
+ next
+ fix x assume x: "x \<in> space M"
+ thus "0 \<le> w x" by fact
+ show "\<bar>u' x\<bar> \<le> w x" using u'_bound[OF x] .
+ qed
+
+ let "?diff n x" = "2 * w x - \<bar>u n x - u' x\<bar>"
+ have diff: "\<And>n. integrable (\<lambda>x. \<bar>u n x - u' x\<bar>)"
+ using w u `integrable u'`
+ by (auto intro!: integral_add integral_diff integral_cmult integrable_abs)
+
+ { fix j x assume x: "x \<in> space M"
+ have "\<bar>u j x - u' x\<bar> \<le> \<bar>u j x\<bar> + \<bar>u' x\<bar>" by auto
+ also have "\<dots> \<le> w x + w x"
+ by (rule add_mono[OF bound[OF x] u'_bound[OF x]])
+ finally have "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp }
+ note diff_less_2w = this
+
+ have PI_diff: "\<And>m n. positive_integral (\<lambda>x. Real (?diff (m + n) x)) =
+ positive_integral (\<lambda>x. Real (2 * w x)) - positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)"
+ using diff w diff_less_2w
+ by (subst positive_integral_diff[symmetric])
+ (auto simp: integrable_def intro!: positive_integral_cong)
+
+ have "integrable (\<lambda>x. 2 * w x)"
+ using w by (auto intro: integral_cmult)
+ hence I2w_fin: "positive_integral (\<lambda>x. Real (2 * w x)) \<noteq> \<omega>" and
+ borel_2w: "(\<lambda>x. Real (2 * w x)) \<in> borel_measurable M"
+ unfolding integrable_def by auto
+
+ have "(INF n. SUP m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)) = 0" (is "?lim_SUP = 0")
+ proof cases
+ assume eq_0: "positive_integral (\<lambda>x. Real (2 * w x)) = 0"
+ have "\<And>i. positive_integral (\<lambda>x. Real \<bar>u i x - u' x\<bar>) \<le> positive_integral (\<lambda>x. Real (2 * w x))"
+ proof (rule positive_integral_mono)
+ fix i x assume "x \<in> space M" from diff_less_2w[OF this, of i]
+ show "Real \<bar>u i x - u' x\<bar> \<le> Real (2 * w x)" by auto
+ qed
+ hence "\<And>i. positive_integral (\<lambda>x. Real \<bar>u i x - u' x\<bar>) = 0" using eq_0 by auto
+ thus ?thesis by simp
+ next
+ assume neq_0: "positive_integral (\<lambda>x. Real (2 * w x)) \<noteq> 0"
+ have "positive_integral (\<lambda>x. Real (2 * w x)) = positive_integral (SUP n. INF m. (\<lambda>x. Real (?diff (m + n) x)))"
+ proof (rule positive_integral_cong, unfold SUPR_fun_expand INFI_fun_expand, subst add_commute)
+ fix x assume x: "x \<in> space M"
+ show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))"
+ proof (rule LIMSEQ_imp_lim_INF[symmetric])
+ fix j show "0 \<le> ?diff j x" using diff_less_2w[OF x, of j] by simp
+ next
+ have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>"
+ using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs)
+ thus "(\<lambda>i. ?diff i x) ----> 2 * w x" by simp
+ qed
+ qed
+ also have "\<dots> \<le> (SUP n. INF m. positive_integral (\<lambda>x. Real (?diff (m + n) x)))"
+ using u'_borel w u unfolding integrable_def
+ by (auto intro!: positive_integral_lim_INF)
+ also have "\<dots> = positive_integral (\<lambda>x. Real (2 * w x)) -
+ (INF n. SUP m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>))"
+ unfolding PI_diff pinfreal_INF_minus[OF I2w_fin] pinfreal_SUP_minus ..
+ finally show ?thesis using neq_0 I2w_fin by (rule pinfreal_le_minus_imp_0)
+ qed
+
+ have [simp]: "\<And>n m x. Real (- \<bar>u (m + n) x - u' x\<bar>) = 0" by auto
+
+ have [simp]: "\<And>n m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>) =
+ Real (integral (\<lambda>x. \<bar>u (n + m) x - u' x\<bar>))"
+ using diff by (subst add_commute) (simp add: integral_def integrable_def Real_real)
+
+ have "(SUP n. INF m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)) \<le> ?lim_SUP"
+ (is "?lim_INF \<le> _") by (subst (1 2) add_commute) (rule lim_INF_le_lim_SUP)
+ hence "?lim_INF = Real 0" "?lim_SUP = Real 0" using `?lim_SUP = 0` by auto
+ thus ?lim_diff using diff by (auto intro!: integral_positive lim_INF_eq_lim_SUP)
+
+ show ?lim
+ proof (rule LIMSEQ_I)
+ fix r :: real assume "0 < r"
+ from LIMSEQ_D[OF `?lim_diff` this]
+ obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> integral (\<lambda>x. \<bar>u n x - u' x\<bar>) < r"
+ using diff by (auto simp: integral_positive)
+
+ show "\<exists>N. \<forall>n\<ge>N. norm (integral (u n) - integral u') < r"
+ proof (safe intro!: exI[of _ N])
+ fix n assume "N \<le> n"
+ have "\<bar>integral (u n) - integral u'\<bar> = \<bar>integral (\<lambda>x. u n x - u' x)\<bar>"
+ using u `integrable u'` by (auto simp: integral_diff)
+ also have "\<dots> \<le> integral (\<lambda>x. \<bar>u n x - u' x\<bar>)" using u `integrable u'`
+ by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff)
+ also note N[OF `N \<le> n`]
+ finally show "norm (integral (u n) - integral u') < r" by simp
+ qed
+ qed
+qed
+
+lemma (in measure_space) integral_sums:
+ assumes borel: "\<And>i. integrable (f i)"
+ and summable: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>f i x\<bar>)"
+ and sums: "summable (\<lambda>i. integral (\<lambda>x. \<bar>f i x\<bar>))"
+ shows "integrable (\<lambda>x. (\<Sum>i. f i x))" (is "integrable ?S")
+ and "(\<lambda>i. integral (f i)) sums integral (\<lambda>x. (\<Sum>i. f i x))" (is ?integral)
+proof -
+ have "\<forall>x\<in>space M. \<exists>w. (\<lambda>i. \<bar>f i x\<bar>) sums w"
+ using summable unfolding summable_def by auto
+ from bchoice[OF this]
+ obtain w where w: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. \<bar>f i x\<bar>) sums w x" by auto
+
+ let "?w y" = "if y \<in> space M then w y else 0"
+
+ obtain x where abs_sum: "(\<lambda>i. integral (\<lambda>x. \<bar>f i x\<bar>)) sums x"
+ using sums unfolding summable_def ..
+
+ have 1: "\<And>n. integrable (\<lambda>x. \<Sum>i = 0..<n. f i x)"
+ using borel by (auto intro!: integral_setsum)
+
+ { fix j x assume [simp]: "x \<in> space M"
+ have "\<bar>\<Sum>i = 0..< j. f i x\<bar> \<le> (\<Sum>i = 0..< j. \<bar>f i x\<bar>)" by (rule setsum_abs)
+ also have "\<dots> \<le> w x" using w[of x] series_pos_le[of "\<lambda>i. \<bar>f i x\<bar>"] unfolding sums_iff by auto
+ finally have "\<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x" by simp }
+ note 2 = this
+
+ have 3: "integrable ?w"
+ proof (rule integral_monotone_convergence(1))
+ let "?F n y" = "(\<Sum>i = 0..<n. \<bar>f i y\<bar>)"
+ let "?w' n y" = "if y \<in> space M then ?F n y else 0"
+ have "\<And>n. integrable (?F n)"
+ using borel by (auto intro!: integral_setsum integrable_abs)
+ thus "\<And>n. integrable (?w' n)" by (simp cong: integrable_cong)
+ show "mono ?w'"
+ by (auto simp: mono_def le_fun_def intro!: setsum_mono2)
+ { fix x show "(\<lambda>n. ?w' n x) ----> ?w x"
+ using w by (cases "x \<in> space M") (simp_all add: LIMSEQ_const sums_def) }
+ have *: "\<And>n. integral (?w' n) = (\<Sum>i = 0..< n. integral (\<lambda>x. \<bar>f i x\<bar>))"
+ using borel by (simp add: integral_setsum integrable_abs cong: integral_cong)
+ from abs_sum
+ show "(\<lambda>i. integral (?w' i)) ----> x" unfolding * sums_def .
+ qed
+
+ have 4: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> ?w x" using 2[of _ 0] by simp
+
+ from summable[THEN summable_rabs_cancel]
+ have 5: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>n. \<Sum>i = 0..<n. f i x) ----> (\<Sum>i. f i x)"
+ by (auto intro: summable_sumr_LIMSEQ_suminf)
+
+ note int = integral_dominated_convergence(1,3)[OF 1 2 3 4 5]
+
+ from int show "integrable ?S" by simp
+
+ show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF borel]
+ using int(2) by simp
+qed
+
+section "Lebesgue integration on countable spaces"
+
+lemma (in measure_space) integral_on_countable:
+ assumes f: "f \<in> borel_measurable M"
+ and bij: "bij_betw enum S (f ` space M)"
+ and enum_zero: "enum ` (-S) \<subseteq> {0}"
+ and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
+ and abs_summable: "summable (\<lambda>r. \<bar>enum r * real (\<mu> (f -` {enum r} \<inter> space M))\<bar>)"
+ shows "integrable f"
+ and "(\<lambda>r. enum r * real (\<mu> (f -` {enum r} \<inter> space M))) sums integral f" (is ?sums)
+proof -
+ let "?A r" = "f -` {enum r} \<inter> space M"
+ let "?F r x" = "enum r * indicator (?A r) x"
+ have enum_eq: "\<And>r. enum r * real (\<mu> (?A r)) = integral (?F r)"
+ using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
+
+ { fix x assume "x \<in> space M"
+ hence "f x \<in> enum ` S" using bij unfolding bij_betw_def by auto
+ then obtain i where "i\<in>S" "enum i = f x" by auto
+ have F: "\<And>j. ?F j x = (if j = i then f x else 0)"
+ proof cases
+ fix j assume "j = i"
+ thus "?thesis j" using `x \<in> space M` `enum i = f x` by (simp add: indicator_def)
+ next
+ fix j assume "j \<noteq> i"
+ show "?thesis j" using bij `i \<in> S` `j \<noteq> i` `enum i = f x` enum_zero
+ by (cases "j \<in> S") (auto simp add: indicator_def bij_betw_def inj_on_def)
+ qed
+ hence F_abs: "\<And>j. \<bar>if j = i then f x else 0\<bar> = (if j = i then \<bar>f x\<bar> else 0)" by auto
+ have "(\<lambda>i. ?F i x) sums f x"
+ "(\<lambda>i. \<bar>?F i x\<bar>) sums \<bar>f x\<bar>"
+ by (auto intro!: sums_single simp: F F_abs) }
+ note F_sums_f = this(1) and F_abs_sums_f = this(2)
+
+ have int_f: "integral f = integral (\<lambda>x. \<Sum>r. ?F r x)" "integrable f = integrable (\<lambda>x. \<Sum>r. ?F r x)"
+ using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff)
+
+ { fix r
+ have "integral (\<lambda>x. \<bar>?F r x\<bar>) = integral (\<lambda>x. \<bar>enum r\<bar> * indicator (?A r) x)"
+ by (auto simp: indicator_def intro!: integral_cong)
+ also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
+ using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
+ finally have "integral (\<lambda>x. \<bar>?F r x\<bar>) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
+ by (simp add: abs_mult_pos real_pinfreal_pos) }
+ note int_abs_F = this
+
+ have 1: "\<And>i. integrable (\<lambda>x. ?F i x)"
+ using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
+
+ have 2: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>?F i x\<bar>)"
+ using F_abs_sums_f unfolding sums_iff by auto
+
+ from integral_sums(2)[OF 1 2, unfolded int_abs_F, OF _ abs_summable]
+ show ?sums unfolding enum_eq int_f by simp
+
+ from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable]
+ show "integrable f" unfolding int_f by simp
+qed
+
+section "Lebesgue integration on finite space"
+
+lemma (in measure_space) integral_on_finite:
+ assumes f: "f \<in> borel_measurable M" and finite: "finite (f`space M)"
+ and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
+ shows "integrable f"
+ and "integral (\<lambda>x. f x) =
+ (\<Sum> r \<in> f`space M. r * real (\<mu> (f -` {r} \<inter> space M)))" (is "?integral")
+proof -
+ let "?A r" = "f -` {r} \<inter> space M"
+ let "?S x" = "\<Sum>r\<in>f`space M. r * indicator (?A r) x"
+
+ { fix x assume "x \<in> space M"
+ have "f x = (\<Sum>r\<in>f`space M. if x \<in> ?A r then r else 0)"
+ using finite `x \<in> space M` by (simp add: setsum_cases)
+ also have "\<dots> = ?S x"
+ by (auto intro!: setsum_cong)
+ finally have "f x = ?S x" . }
+ note f_eq = this
+
+ have f_eq_S: "integrable f \<longleftrightarrow> integrable ?S" "integral f = integral ?S"
+ by (auto intro!: integrable_cong integral_cong simp only: f_eq)
+
+ show "integrable f" ?integral using fin f f_eq_S
+ by (simp_all add: integral_cmul_indicator borel_measurable_vimage)
+qed
+
+lemma sigma_algebra_cong:
+ fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
+ assumes *: "sigma_algebra M"
+ and cong: "space M = space M'" "sets M = sets M'"
+ shows "sigma_algebra M'"
+using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
+
+lemma finite_Pow_additivity_sufficient:
+ assumes "finite (space M)" and "sets M = Pow (space M)"
+ and "positive \<mu>" and "additive M \<mu>"
+ and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
+ shows "finite_measure_space M \<mu>"
+proof -
+ have "sigma_algebra M"
+ using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow])
+
+ have "measure_space M \<mu>"
+ by (rule sigma_algebra.finite_additivity_sufficient) (fact+)
+ thus ?thesis
+ unfolding finite_measure_space_def finite_measure_space_axioms_def
+ using assms by simp
+qed
+
+lemma finite_measure_spaceI:
+ assumes "measure_space M \<mu>" and "finite (space M)" and "sets M = Pow (space M)"
+ and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
+ shows "finite_measure_space M \<mu>"
+ unfolding finite_measure_space_def finite_measure_space_axioms_def
+ using assms by simp
+
+lemma (in finite_measure_space) borel_measurable_finite[intro, simp]:
+ fixes f :: "'a \<Rightarrow> real"
+ shows "f \<in> borel_measurable M"
+ unfolding measurable_def sets_eq_Pow by auto
+
+lemma (in finite_measure_space) integral_finite_singleton:
+ shows "integrable f"
+ and "integral f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
+proof -
+ have 1: "f \<in> borel_measurable M"
+ unfolding measurable_def sets_eq_Pow by auto
+
+ have 2: "finite (f`space M)" using finite_space by simp
+
+ have 3: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>"
+ using finite_measure[unfolded sets_eq_Pow] by simp
+
+ show "integrable f"
+ by (rule integral_on_finite(1)[OF 1 2 3]) simp
+
+ { fix r let ?x = "f -` {r} \<inter> space M"
+ have "?x \<subseteq> space M" by auto
+ with finite_space sets_eq_Pow finite_single_measure
+ have "real (\<mu> ?x) = (\<Sum>i \<in> ?x. real (\<mu> {i}))"
+ using real_measure_setsum_singleton[of ?x] by auto }
+ note measure_eq_setsum = this
+
+ have "integral f = (\<Sum>r\<in>f`space M. r * real (\<mu> (f -` {r} \<inter> space M)))"
+ by (rule integral_on_finite(2)[OF 1 2 3]) simp
+ also have "\<dots> = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))"
+ unfolding measure_eq_setsum setsum_right_distrib
+ apply (subst setsum_Sigma)
+ apply (simp add: finite_space)
+ apply (simp add: finite_space)
+ proof (rule setsum_reindex_cong[symmetric])
+ fix a assume "a \<in> Sigma (f ` space M) (\<lambda>x. f -` {x} \<inter> space M)"
+ thus "(\<lambda>(x, y). x * real (\<mu> {y})) a = f (snd a) * real (\<mu> {snd a})"
+ by auto
+ qed (auto intro!: image_eqI inj_onI)
+ finally show ?I .
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Lebesgue_Measure.thy Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,576 @@
+
+header {* Lebsegue measure *}
+(* Author: Robert Himmelmann, TU Muenchen *)
+
+theory Lebesgue_Measure
+ imports Gauge_Measure Measure Lebesgue_Integration
+begin
+
+subsection {* Various *}
+
+lemma seq_offset_iff:"f ----> l \<longleftrightarrow> (\<lambda>i. f (i + k)) ----> l"
+ using seq_offset_rev seq_offset[of f l k] by auto
+
+lemma has_integral_disjoint_union: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
+ assumes "(f has_integral i) s" "(f has_integral j) t" "s \<inter> t = {}"
+ shows "(f has_integral (i + j)) (s \<union> t)"
+ apply(rule has_integral_union[OF assms(1-2)]) unfolding assms by auto
+
+lemma lim_eq: assumes "\<forall>n>N. f n = g n" shows "(f ----> l) \<longleftrightarrow> (g ----> l)" using assms
+proof(induct N arbitrary: f g) case 0
+ hence *:"\<And>n. f (Suc n) = g (Suc n)" by auto
+ show ?case apply(subst LIMSEQ_Suc_iff[THEN sym]) apply(subst(2) LIMSEQ_Suc_iff[THEN sym])
+ unfolding * ..
+next case (Suc n)
+ show ?case apply(subst LIMSEQ_Suc_iff[THEN sym]) apply(subst(2) LIMSEQ_Suc_iff[THEN sym])
+ apply(rule Suc(1)) using Suc(2) by auto
+qed
+
+subsection {* Standard Cubes *}
+
+definition cube where
+ "cube (n::nat) \<equiv> {\<chi>\<chi> i. - real n .. (\<chi>\<chi> i. real n)::_::ordered_euclidean_space}"
+
+lemma cube_subset[intro]:"n\<le>N \<Longrightarrow> cube n \<subseteq> (cube N::'a::ordered_euclidean_space set)"
+ apply(auto simp: eucl_le[where 'a='a] cube_def) apply(erule_tac[!] x=i in allE)+ by auto
+
+lemma ball_subset_cube:"ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n"
+ unfolding cube_def subset_eq mem_interval apply safe unfolding euclidean_lambda_beta'
+proof- fix x::'a and i assume as:"x\<in>ball 0 (real n)" "i<DIM('a)"
+ thus "- real n \<le> x $$ i" "real n \<ge> x $$ i"
+ using component_le_norm[of x i] by(auto simp: dist_norm)
+qed
+
+lemma mem_big_cube: obtains n where "x \<in> cube n"
+proof- from real_arch_lt[of "norm x"] guess n ..
+ thus ?thesis apply-apply(rule that[where n=n])
+ apply(rule ball_subset_cube[unfolded subset_eq,rule_format])
+ by (auto simp add:dist_norm)
+qed
+
+lemma Union_inter_cube:"\<Union>{s \<inter> cube n |n. n \<in> UNIV} = s"
+proof safe case goal1
+ from mem_big_cube[of x] guess n . note n=this
+ show ?case unfolding Union_iff apply(rule_tac x="s \<inter> cube n" in bexI)
+ using n goal1 by auto
+qed
+
+lemma gmeasurable_cube[intro]:"gmeasurable (cube n)"
+ unfolding cube_def by auto
+
+lemma gmeasure_le_inter_cube[intro]: fixes s::"'a::ordered_euclidean_space set"
+ assumes "gmeasurable (s \<inter> cube n)" shows "gmeasure (s \<inter> cube n) \<le> gmeasure (cube n::'a set)"
+ apply(rule has_gmeasure_subset[of "s\<inter>cube n" _ "cube n"])
+ unfolding has_gmeasure_measure[THEN sym] using assms by auto
+
+
+subsection {* Measurability *}
+
+definition lmeasurable :: "('a::ordered_euclidean_space) set => bool" where
+ "lmeasurable s \<equiv> (\<forall>n::nat. gmeasurable (s \<inter> cube n))"
+
+lemma lmeasurableD[dest]:assumes "lmeasurable s"
+ shows "\<And>n. gmeasurable (s \<inter> cube n)"
+ using assms unfolding lmeasurable_def by auto
+
+lemma measurable_imp_lmeasurable: assumes"gmeasurable s"
+ shows "lmeasurable s" unfolding lmeasurable_def cube_def
+ using assms gmeasurable_interval by auto
+
+lemma lmeasurable_empty[intro]: "lmeasurable {}"
+ using gmeasurable_empty apply- apply(drule_tac measurable_imp_lmeasurable) .
+
+lemma lmeasurable_union[intro]: assumes "lmeasurable s" "lmeasurable t"
+ shows "lmeasurable (s \<union> t)"
+ using assms unfolding lmeasurable_def Int_Un_distrib2
+ by(auto intro:gmeasurable_union)
+
+lemma lmeasurable_countable_unions_strong:
+ fixes s::"nat => 'a::ordered_euclidean_space set"
+ assumes "\<And>n::nat. lmeasurable(s n)"
+ shows "lmeasurable(\<Union>{ s(n) |n. n \<in> UNIV })"
+ unfolding lmeasurable_def
+proof fix n::nat
+ have *:"\<Union>{s n |n. n \<in> UNIV} \<inter> cube n = \<Union>{s k \<inter> cube n |k. k \<in> UNIV}" by auto
+ show "gmeasurable (\<Union>{s n |n. n \<in> UNIV} \<inter> cube n)" unfolding *
+ apply(rule gmeasurable_countable_unions_strong)
+ apply(rule assms[unfolded lmeasurable_def,rule_format])
+ proof- fix k::nat
+ show "gmeasure (\<Union>{s ka \<inter> cube n |ka. ka \<le> k}) \<le> gmeasure (cube n::'a set)"
+ apply(rule measure_subset) apply(rule gmeasurable_finite_unions)
+ using assms(1)[unfolded lmeasurable_def] by auto
+ qed
+qed
+
+lemma lmeasurable_inter[intro]: fixes s::"'a :: ordered_euclidean_space set"
+ assumes "lmeasurable s" "lmeasurable t" shows "lmeasurable (s \<inter> t)"
+ unfolding lmeasurable_def
+proof fix n::nat
+ have *:"s \<inter> t \<inter> cube n = (s \<inter> cube n) \<inter> (t \<inter> cube n)" by auto
+ show "gmeasurable (s \<inter> t \<inter> cube n)"
+ using assms unfolding lmeasurable_def *
+ using gmeasurable_inter[of "s \<inter> cube n" "t \<inter> cube n"] by auto
+qed
+
+lemma lmeasurable_complement[intro]: assumes "lmeasurable s"
+ shows "lmeasurable (UNIV - s)"
+ unfolding lmeasurable_def
+proof fix n::nat
+ have *:"(UNIV - s) \<inter> cube n = cube n - (s \<inter> cube n)" by auto
+ show "gmeasurable ((UNIV - s) \<inter> cube n)" unfolding *
+ apply(rule gmeasurable_diff) using assms unfolding lmeasurable_def by auto
+qed
+
+lemma lmeasurable_finite_unions:
+ assumes "finite f" "\<And>s. s \<in> f \<Longrightarrow> lmeasurable s"
+ shows "lmeasurable (\<Union> f)" unfolding lmeasurable_def
+proof fix n::nat have *:"(\<Union>f \<inter> cube n) = \<Union>{x \<inter> cube n | x . x\<in>f}" by auto
+ show "gmeasurable (\<Union>f \<inter> cube n)" unfolding *
+ apply(rule gmeasurable_finite_unions) unfolding simple_image
+ using assms unfolding lmeasurable_def by auto
+qed
+
+lemma negligible_imp_lmeasurable[dest]: fixes s::"'a::ordered_euclidean_space set"
+ assumes "negligible s" shows "lmeasurable s"
+ unfolding lmeasurable_def
+proof case goal1
+ have *:"\<And>x. (if x \<in> cube n then indicator s x else 0) = (if x \<in> s \<inter> cube n then 1 else 0)"
+ unfolding indicator_def_raw by auto
+ note assms[unfolded negligible_def,rule_format,of "(\<chi>\<chi> i. - real n)::'a" "\<chi>\<chi> i. real n"]
+ thus ?case apply-apply(rule gmeasurableI[of _ 0]) unfolding has_gmeasure_def
+ apply(subst(asm) has_integral_restrict_univ[THEN sym]) unfolding cube_def[symmetric]
+ apply(subst has_integral_restrict_univ[THEN sym]) unfolding * .
+qed
+
+
+section {* The Lebesgue Measure *}
+
+definition has_lmeasure (infixr "has'_lmeasure" 80) where
+ "s has_lmeasure m \<equiv> lmeasurable s \<and> ((\<lambda>n. Real (gmeasure (s \<inter> cube n))) ---> m) sequentially"
+
+lemma has_lmeasureD: assumes "s has_lmeasure m"
+ shows "lmeasurable s" "gmeasurable (s \<inter> cube n)"
+ "((\<lambda>n. Real (gmeasure (s \<inter> cube n))) ---> m) sequentially"
+ using assms unfolding has_lmeasure_def lmeasurable_def by auto
+
+lemma has_lmeasureI: assumes "lmeasurable s" "((\<lambda>n. Real (gmeasure (s \<inter> cube n))) ---> m) sequentially"
+ shows "s has_lmeasure m" using assms unfolding has_lmeasure_def by auto
+
+definition lmeasure where
+ "lmeasure s \<equiv> SOME m. s has_lmeasure m"
+
+lemma has_lmeasure_has_gmeasure: assumes "s has_lmeasure (Real m)" "m\<ge>0"
+ shows "s has_gmeasure m"
+proof- note s = has_lmeasureD[OF assms(1)]
+ have *:"(\<lambda>n. (gmeasure (s \<inter> cube n))) ----> m"
+ using s(3) apply(subst (asm) lim_Real) using s(2) assms(2) by auto
+
+ have "(\<lambda>x. if x \<in> s then 1 else (0::real)) integrable_on UNIV \<and>
+ (\<lambda>k. integral UNIV (\<lambda>x. if x \<in> s \<inter> cube k then 1 else (0::real)))
+ ----> integral UNIV (\<lambda>x. if x \<in> s then 1 else 0)"
+ proof(rule monotone_convergence_increasing)
+ have "\<forall>n. gmeasure (s \<inter> cube n) \<le> m" apply(rule ccontr) unfolding not_all not_le
+ proof(erule exE) fix k assume k:"m < gmeasure (s \<inter> cube k)"
+ hence "gmeasure (s \<inter> cube k) - m > 0" by auto
+ from *[unfolded Lim_sequentially,rule_format,OF this] guess N ..
+ note this[unfolded dist_real_def,rule_format,of "N + k"]
+ moreover have "gmeasure (s \<inter> cube (N + k)) \<ge> gmeasure (s \<inter> cube k)" apply-
+ apply(rule measure_subset) prefer 3 using s(2)
+ using cube_subset[of k "N + k"] by auto
+ ultimately show False by auto
+ qed
+ thus *:"bounded {integral UNIV (\<lambda>x. if x \<in> s \<inter> cube k then 1 else (0::real)) |k. True}"
+ unfolding integral_measure_univ[OF s(2)] bounded_def apply-
+ apply(rule_tac x=0 in exI,rule_tac x=m in exI) unfolding dist_real_def
+ by (auto simp: measure_pos_le)
+
+ show "\<forall>k. (\<lambda>x. if x \<in> s \<inter> cube k then (1::real) else 0) integrable_on UNIV"
+ unfolding integrable_restrict_univ
+ using s(2) unfolding gmeasurable_def has_gmeasure_def by auto
+ have *:"\<And>n. n \<le> Suc n" by auto
+ show "\<forall>k. \<forall>x\<in>UNIV. (if x \<in> s \<inter> cube k then 1 else 0) \<le> (if x \<in> s \<inter> cube (Suc k) then 1 else (0::real))"
+ using cube_subset[OF *] by fastsimp
+ show "\<forall>x\<in>UNIV. (\<lambda>k. if x \<in> s \<inter> cube k then 1 else 0) ----> (if x \<in> s then 1 else (0::real))"
+ unfolding Lim_sequentially
+ proof safe case goal1 from real_arch_lt[of "norm x"] guess N .. note N = this
+ show ?case apply(rule_tac x=N in exI)
+ proof safe case goal1
+ have "x \<in> cube n" using cube_subset[OF goal1] N
+ using ball_subset_cube[of N] by(auto simp: dist_norm)
+ thus ?case using `e>0` by auto
+ qed
+ qed
+ qed note ** = conjunctD2[OF this]
+ hence *:"m = integral UNIV (\<lambda>x. if x \<in> s then 1 else 0)" apply-
+ apply(rule LIMSEQ_unique[OF _ **(2)]) unfolding measure_integral_univ[THEN sym,OF s(2)] using * .
+ show ?thesis unfolding has_gmeasure * apply(rule integrable_integral) using ** by auto
+qed
+
+lemma has_lmeasure_unique: "s has_lmeasure m1 \<Longrightarrow> s has_lmeasure m2 \<Longrightarrow> m1 = m2"
+ unfolding has_lmeasure_def apply(rule Lim_unique) using trivial_limit_sequentially by auto
+
+lemma lmeasure_unique[intro]: assumes "A has_lmeasure m" shows "lmeasure A = m"
+ using assms unfolding lmeasure_def lmeasurable_def apply-
+ apply(rule some_equality) defer apply(rule has_lmeasure_unique) by auto
+
+lemma glmeasurable_finite: assumes "lmeasurable s" "lmeasure s \<noteq> \<omega>"
+ shows "gmeasurable s"
+proof- have "\<exists>B. \<forall>n. gmeasure (s \<inter> cube n) \<le> B"
+ proof(rule ccontr) case goal1
+ note as = this[unfolded not_ex not_all not_le]
+ have "s has_lmeasure \<omega>" apply- apply(rule has_lmeasureI[OF assms(1)])
+ unfolding Lim_omega
+ proof fix B::real
+ from as[rule_format,of B] guess N .. note N = this
+ have "\<And>n. N \<le> n \<Longrightarrow> B \<le> gmeasure (s \<inter> cube n)"
+ apply(rule order_trans[where y="gmeasure (s \<inter> cube N)"]) defer
+ apply(rule measure_subset) prefer 3
+ using cube_subset N assms(1)[unfolded lmeasurable_def] by auto
+ thus "\<exists>N. \<forall>n\<ge>N. Real B \<le> Real (gmeasure (s \<inter> cube n))" apply-
+ apply(subst Real_max') apply(rule_tac x=N in exI,safe)
+ unfolding pinfreal_less_eq apply(subst if_P) by auto
+ qed note lmeasure_unique[OF this]
+ thus False using assms(2) by auto
+ qed then guess B .. note B=this
+
+ show ?thesis apply(rule gmeasurable_nested_unions[of "\<lambda>n. s \<inter> cube n",
+ unfolded Union_inter_cube,THEN conjunct1, where B1=B])
+ proof- fix n::nat
+ show " gmeasurable (s \<inter> cube n)" using assms by auto
+ show "gmeasure (s \<inter> cube n) \<le> B" using B by auto
+ show "s \<inter> cube n \<subseteq> s \<inter> cube (Suc n)"
+ by (rule Int_mono) (simp_all add: cube_subset)
+ qed
+qed
+
+lemma lmeasure_empty[intro]:"lmeasure {} = 0"
+ apply(rule lmeasure_unique)
+ unfolding has_lmeasure_def by auto
+
+lemma lmeasurableI[dest]:"s has_lmeasure m \<Longrightarrow> lmeasurable s"
+ unfolding has_lmeasure_def by auto
+
+lemma has_gmeasure_has_lmeasure: assumes "s has_gmeasure m"
+ shows "s has_lmeasure (Real m)"
+proof- have gmea:"gmeasurable s" using assms by auto
+ have m:"m \<ge> 0" using assms by auto
+ have *:"m = gmeasure (\<Union>{s \<inter> cube n |n. n \<in> UNIV})" unfolding Union_inter_cube
+ using assms by(rule measure_unique[THEN sym])
+ show ?thesis unfolding has_lmeasure_def
+ apply(rule,rule measurable_imp_lmeasurable[OF gmea])
+ apply(subst lim_Real) apply(rule,rule,rule m) unfolding *
+ apply(rule gmeasurable_nested_unions[THEN conjunct2, where B1="gmeasure s"])
+ proof- fix n::nat show *:"gmeasurable (s \<inter> cube n)"
+ using gmeasurable_inter[OF gmea gmeasurable_cube] .
+ show "gmeasure (s \<inter> cube n) \<le> gmeasure s" apply(rule measure_subset)
+ apply(rule * gmea)+ by auto
+ show "s \<inter> cube n \<subseteq> s \<inter> cube (Suc n)" using cube_subset[of n "Suc n"] by auto
+ qed
+qed
+
+lemma gmeasure_lmeasure: assumes "gmeasurable s" shows "lmeasure s = Real (gmeasure s)"
+proof- note has_gmeasure_measureI[OF assms]
+ note has_gmeasure_has_lmeasure[OF this]
+ thus ?thesis by(rule lmeasure_unique)
+qed
+
+lemma has_lmeasure_lmeasure: "lmeasurable s \<longleftrightarrow> s has_lmeasure (lmeasure s)" (is "?l = ?r")
+proof assume ?l let ?f = "\<lambda>n. Real (gmeasure (s \<inter> cube n))"
+ have "\<forall>n m. n\<ge>m \<longrightarrow> ?f n \<ge> ?f m" unfolding pinfreal_less_eq apply safe
+ apply(subst if_P) defer apply(rule measure_subset) prefer 3
+ apply(drule cube_subset) using `?l` by auto
+ from lim_pinfreal_increasing[OF this] guess l . note l=this
+ hence "s has_lmeasure l" using `?l` apply-apply(rule has_lmeasureI) by auto
+ thus ?r using lmeasure_unique by auto
+next assume ?r thus ?l unfolding has_lmeasure_def by auto
+qed
+
+lemma lmeasure_subset[dest]: assumes "lmeasurable s" "lmeasurable t" "s \<subseteq> t"
+ shows "lmeasure s \<le> lmeasure t"
+proof(cases "lmeasure t = \<omega>")
+ case False have som:"lmeasure s \<noteq> \<omega>"
+ proof(rule ccontr,unfold not_not) assume as:"lmeasure s = \<omega>"
+ have "t has_lmeasure \<omega>" using assms(2) apply(rule has_lmeasureI)
+ unfolding Lim_omega
+ proof case goal1
+ note assms(1)[unfolded has_lmeasure_lmeasure]
+ note has_lmeasureD(3)[OF this,unfolded as Lim_omega,rule_format,of B]
+ then guess N .. note N = this
+ show ?case apply(rule_tac x=N in exI) apply safe
+ apply(rule order_trans) apply(rule N[rule_format],assumption)
+ unfolding pinfreal_less_eq apply(subst if_P)defer
+ apply(rule measure_subset) using assms by auto
+ qed
+ thus False using lmeasure_unique False by auto
+ qed
+
+ note assms(1)[unfolded has_lmeasure_lmeasure] note has_lmeasureD(3)[OF this]
+ hence "(\<lambda>n. Real (gmeasure (s \<inter> cube n))) ----> Real (real (lmeasure s))"
+ unfolding Real_real'[OF som] .
+ hence l1:"(\<lambda>n. gmeasure (s \<inter> cube n)) ----> real (lmeasure s)"
+ apply-apply(subst(asm) lim_Real) by auto
+
+ note assms(2)[unfolded has_lmeasure_lmeasure] note has_lmeasureD(3)[OF this]
+ hence "(\<lambda>n. Real (gmeasure (t \<inter> cube n))) ----> Real (real (lmeasure t))"
+ unfolding Real_real'[OF False] .
+ hence l2:"(\<lambda>n. gmeasure (t \<inter> cube n)) ----> real (lmeasure t)"
+ apply-apply(subst(asm) lim_Real) by auto
+
+ have "real (lmeasure s) \<le> real (lmeasure t)" apply(rule LIMSEQ_le[OF l1 l2])
+ apply(rule_tac x=0 in exI,safe) apply(rule measure_subset) using assms by auto
+ hence "Real (real (lmeasure s)) \<le> Real (real (lmeasure t))"
+ unfolding pinfreal_less_eq by auto
+ thus ?thesis unfolding Real_real'[OF som] Real_real'[OF False] .
+qed auto
+
+lemma has_lmeasure_negligible_unions_image:
+ assumes "finite s" "\<And>x. x \<in> s ==> lmeasurable(f x)"
+ "\<And>x y. x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x \<noteq> y \<Longrightarrow> negligible((f x) \<inter> (f y))"
+ shows "(\<Union> (f ` s)) has_lmeasure (setsum (\<lambda>x. lmeasure(f x)) s)"
+ unfolding has_lmeasure_def
+proof show lmeaf:"lmeasurable (\<Union>f ` s)" apply(rule lmeasurable_finite_unions)
+ using assms(1-2) by auto
+ show "(\<lambda>n. Real (gmeasure (\<Union>f ` s \<inter> cube n))) ----> (\<Sum>x\<in>s. lmeasure (f x))" (is ?l)
+ proof(cases "\<exists>x\<in>s. lmeasure (f x) = \<omega>")
+ case False hence *:"(\<Sum>x\<in>s. lmeasure (f x)) \<noteq> \<omega>" apply-
+ apply(rule setsum_neq_omega) using assms(1) by auto
+ have gmea:"\<And>x. x\<in>s \<Longrightarrow> gmeasurable (f x)" apply(rule glmeasurable_finite) using False assms(2) by auto
+ have "(\<Sum>x\<in>s. lmeasure (f x)) = (\<Sum>x\<in>s. Real (gmeasure (f x)))" apply(rule setsum_cong2)
+ apply(rule gmeasure_lmeasure) using False assms(2) gmea by auto
+ also have "... = Real (\<Sum>x\<in>s. (gmeasure (f x)))" apply(rule setsum_Real) by auto
+ finally have sum:"(\<Sum>x\<in>s. lmeasure (f x)) = Real (\<Sum>x\<in>s. gmeasure (f x))" .
+ have sum_0:"(\<Sum>x\<in>s. gmeasure (f x)) \<ge> 0" apply(rule setsum_nonneg) by auto
+ have int_un:"\<Union>f ` s has_gmeasure (\<Sum>x\<in>s. gmeasure (f x))"
+ apply(rule has_gmeasure_negligible_unions_image) using assms gmea by auto
+
+ have unun:"\<Union>{\<Union>f ` s \<inter> cube n |n. n \<in> UNIV} = \<Union>f ` s" unfolding simple_image
+ proof safe fix x y assume as:"x \<in> f y" "y \<in> s"
+ from mem_big_cube[of x] guess n . note n=this
+ thus "x \<in> \<Union>range (\<lambda>n. \<Union>f ` s \<inter> cube n)" unfolding Union_iff
+ apply-apply(rule_tac x="\<Union>f ` s \<inter> cube n" in bexI) using as by auto
+ qed
+ show ?l apply(subst Real_real'[OF *,THEN sym])apply(subst lim_Real)
+ apply rule apply rule unfolding sum real_Real if_P[OF sum_0] apply(rule sum_0)
+ unfolding measure_unique[OF int_un,THEN sym] apply(subst(2) unun[THEN sym])
+ apply(rule has_gmeasure_nested_unions[THEN conjunct2])
+ proof- fix n::nat
+ show *:"gmeasurable (\<Union>f ` s \<inter> cube n)" using lmeaf unfolding lmeasurable_def by auto
+ thus "gmeasure (\<Union>f ` s \<inter> cube n) \<le> gmeasure (\<Union>f ` s)"
+ apply(rule measure_subset) using int_un by auto
+ show "\<Union>f ` s \<inter> cube n \<subseteq> \<Union>f ` s \<inter> cube (Suc n)"
+ using cube_subset[of n "Suc n"] by auto
+ qed
+
+ next case True then guess X .. note X=this
+ hence sum:"(\<Sum>x\<in>s. lmeasure (f x)) = \<omega>" using setsum_\<omega>[THEN iffD2, of s] assms by fastsimp
+ show ?l unfolding sum Lim_omega
+ proof fix B::real
+ have Xm:"(f X) has_lmeasure \<omega>" using X by (metis assms(2) has_lmeasure_lmeasure)
+ note this[unfolded has_lmeasure_def,THEN conjunct2, unfolded Lim_omega]
+ from this[rule_format,of B] guess N .. note N=this[rule_format]
+ show "\<exists>N. \<forall>n\<ge>N. Real B \<le> Real (gmeasure (\<Union>f ` s \<inter> cube n))"
+ apply(rule_tac x=N in exI)
+ proof safe case goal1 show ?case apply(rule order_trans[OF N[OF goal1]])
+ unfolding pinfreal_less_eq apply(subst if_P) defer
+ apply(rule measure_subset) using has_lmeasureD(2)[OF Xm]
+ using lmeaf unfolding lmeasurable_def using X(1) by auto
+ qed qed qed qed
+
+lemma has_lmeasure_negligible_unions:
+ assumes"finite f" "\<And>s. s \<in> f ==> s has_lmeasure (m s)"
+ "\<And>s t. s \<in> f \<Longrightarrow> t \<in> f \<Longrightarrow> s \<noteq> t ==> negligible (s\<inter>t)"
+ shows "(\<Union> f) has_lmeasure (setsum m f)"
+proof- have *:"setsum m f = setsum lmeasure f" apply(rule setsum_cong2)
+ apply(subst lmeasure_unique[OF assms(2)]) by auto
+ show ?thesis unfolding *
+ apply(rule has_lmeasure_negligible_unions_image[where s=f and f=id,unfolded image_id id_apply])
+ using assms by auto
+qed
+
+lemma has_lmeasure_disjoint_unions:
+ assumes"finite f" "\<And>s. s \<in> f ==> s has_lmeasure (m s)"
+ "\<And>s t. s \<in> f \<Longrightarrow> t \<in> f \<Longrightarrow> s \<noteq> t ==> s \<inter> t = {}"
+ shows "(\<Union> f) has_lmeasure (setsum m f)"
+proof- have *:"setsum m f = setsum lmeasure f" apply(rule setsum_cong2)
+ apply(subst lmeasure_unique[OF assms(2)]) by auto
+ show ?thesis unfolding *
+ apply(rule has_lmeasure_negligible_unions_image[where s=f and f=id,unfolded image_id id_apply])
+ using assms by auto
+qed
+
+lemma has_lmeasure_nested_unions:
+ assumes "\<And>n. lmeasurable(s n)" "\<And>n. s(n) \<subseteq> s(Suc n)"
+ shows "lmeasurable(\<Union> { s n | n. n \<in> UNIV }) \<and>
+ (\<lambda>n. lmeasure(s n)) ----> lmeasure(\<Union> { s(n) | n. n \<in> UNIV })" (is "?mea \<and> ?lim")
+proof- have cube:"\<And>m. \<Union> { s(n) | n. n \<in> UNIV } \<inter> cube m = \<Union> { s(n) \<inter> cube m | n. n \<in> UNIV }" by blast
+ have 3:"\<And>n. \<forall>m\<ge>n. s n \<subseteq> s m" apply(rule transitive_stepwise_le) using assms(2) by auto
+ have mea:"?mea" unfolding lmeasurable_def cube apply rule
+ apply(rule_tac B1="gmeasure (cube n)" in has_gmeasure_nested_unions[THEN conjunct1])
+ prefer 3 apply rule using assms(1) unfolding lmeasurable_def
+ by(auto intro!:assms(2)[unfolded subset_eq,rule_format])
+ show ?thesis apply(rule,rule mea)
+ proof(cases "lmeasure(\<Union> { s(n) | n. n \<in> UNIV }) = \<omega>")
+ case True show ?lim unfolding True Lim_omega
+ proof(rule ccontr) case goal1 note this[unfolded not_all not_ex]
+ hence "\<exists>B. \<forall>n. \<exists>m\<ge>n. Real B > lmeasure (s m)" by(auto simp add:not_le)
+ from this guess B .. note B=this[rule_format]
+
+ have "\<forall>n. gmeasurable (s n) \<and> gmeasure (s n) \<le> max B 0"
+ proof safe fix n::nat from B[of n] guess m .. note m=this
+ hence *:"lmeasure (s n) < Real B" apply-apply(rule le_less_trans)
+ apply(rule lmeasure_subset[OF assms(1,1)]) apply(rule 3[rule_format]) by auto
+ thus **:"gmeasurable (s n)" apply-apply(rule glmeasurable_finite[OF assms(1)]) by auto
+ thus "gmeasure (s n) \<le> max B 0" using * unfolding gmeasure_lmeasure[OF **] Real_max'[of B]
+ unfolding pinfreal_less apply- apply(subst(asm) if_P) by auto
+ qed
+ hence "\<And>n. gmeasurable (s n)" "\<And>n. gmeasure (s n) \<le> max B 0" by auto
+ note g = conjunctD2[OF has_gmeasure_nested_unions[of s, OF this assms(2)]]
+ show False using True unfolding gmeasure_lmeasure[OF g(1)] by auto
+ qed
+ next let ?B = "lmeasure (\<Union>{s n |n. n \<in> UNIV})"
+ case False note gmea_lim = glmeasurable_finite[OF mea this]
+ have ls:"\<And>n. lmeasure (s n) \<le> lmeasure (\<Union>{s n |n. n \<in> UNIV})"
+ apply(rule lmeasure_subset) using assms(1) mea by auto
+ have "\<And>n. lmeasure (s n) \<noteq> \<omega>"
+ proof(rule ccontr,safe) case goal1
+ show False using False ls[of n] unfolding goal1 by auto
+ qed
+ note gmea = glmeasurable_finite[OF assms(1) this]
+
+ have "\<And>n. gmeasure (s n) \<le> real ?B" unfolding gmeasure_lmeasure[OF gmea_lim]
+ unfolding real_Real apply(subst if_P,rule) apply(rule measure_subset)
+ using gmea gmea_lim by auto
+ note has_gmeasure_nested_unions[of s, OF gmea this assms(2)]
+ thus ?lim unfolding gmeasure_lmeasure[OF gmea] gmeasure_lmeasure[OF gmea_lim]
+ apply-apply(subst lim_Real) by auto
+ qed
+qed
+
+lemma has_lmeasure_countable_negligible_unions:
+ assumes "\<And>n. lmeasurable(s n)" "\<And>m n. m \<noteq> n \<Longrightarrow> negligible(s m \<inter> s n)"
+ shows "(\<lambda>m. setsum (\<lambda>n. lmeasure(s n)) {..m}) ----> (lmeasure(\<Union> { s(n) |n. n \<in> UNIV }))"
+proof- have *:"\<And>n. (\<Union> (s ` {0..n})) has_lmeasure (setsum (\<lambda>k. lmeasure(s k)) {0..n})"
+ apply(rule has_lmeasure_negligible_unions_image) using assms by auto
+ have **:"(\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV}) = (\<Union>{s n |n. n \<in> UNIV})" unfolding simple_image by fastsimp
+ have "lmeasurable (\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV}) \<and>
+ (\<lambda>n. lmeasure (\<Union>(s ` {0..n}))) ----> lmeasure (\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV})"
+ apply(rule has_lmeasure_nested_unions) apply(rule has_lmeasureD(1)[OF *])
+ apply(rule Union_mono,rule image_mono) by auto
+ note lem = conjunctD2[OF this,unfolded **]
+ show ?thesis using lem(2) unfolding lmeasure_unique[OF *] unfolding atLeast0AtMost .
+qed
+
+lemma lmeasure_eq_0: assumes "negligible s" shows "lmeasure s = 0"
+proof- note mea=negligible_imp_lmeasurable[OF assms]
+ have *:"\<And>n. (gmeasure (s \<inter> cube n)) = 0"
+ unfolding gmeasurable_measure_eq_0[OF mea[unfolded lmeasurable_def,rule_format]]
+ using assms by auto
+ show ?thesis
+ apply(rule lmeasure_unique) unfolding has_lmeasure_def
+ apply(rule,rule mea) unfolding * by auto
+qed
+
+lemma negligible_img_gmeasurable: fixes s::"'a::ordered_euclidean_space set"
+ assumes "negligible s" shows "gmeasurable s"
+ apply(rule glmeasurable_finite)
+ using lmeasure_eq_0[OF assms] negligible_imp_lmeasurable[OF assms] by auto
+
+
+
+
+section {* Instantiation of _::euclidean_space as measure_space *}
+
+definition lebesgue_space :: "'a::ordered_euclidean_space algebra" where
+ "lebesgue_space = \<lparr> space = UNIV, sets = lmeasurable \<rparr>"
+
+lemma lebesgue_measurable[simp]:"A \<in> sets lebesgue_space \<longleftrightarrow> lmeasurable A"
+ unfolding lebesgue_space_def by(auto simp: mem_def)
+
+lemma mem_gmeasurable[simp]: "A \<in> gmeasurable \<longleftrightarrow> gmeasurable A"
+ unfolding mem_def ..
+
+interpretation lebesgue: measure_space lebesgue_space lmeasure
+ apply(intro_locales) unfolding measure_space_axioms_def countably_additive_def
+ unfolding sigma_algebra_axioms_def algebra_def
+ unfolding lebesgue_measurable
+proof safe
+ fix A::"nat => _" assume as:"range A \<subseteq> sets lebesgue_space" "disjoint_family A"
+ "lmeasurable (UNION UNIV A)"
+ have *:"UNION UNIV A = \<Union>range A" by auto
+ show "(\<Sum>\<^isub>\<infinity>n. lmeasure (A n)) = lmeasure (UNION UNIV A)"
+ unfolding psuminf_def apply(rule SUP_Lim_pinfreal)
+ proof- fix n m::nat assume mn:"m\<le>n"
+ have *:"\<And>m. (\<Sum>n<m. lmeasure (A n)) = lmeasure (\<Union>A ` {..<m})"
+ apply(subst lmeasure_unique[OF has_lmeasure_negligible_unions[where m=lmeasure]])
+ apply(rule finite_imageI) apply rule apply(subst has_lmeasure_lmeasure[THEN sym])
+ proof- fix m::nat
+ show "(\<Sum>n<m. lmeasure (A n)) = setsum lmeasure (A ` {..<m})"
+ apply(subst setsum_reindex_nonzero) unfolding o_def apply rule
+ apply(rule lmeasure_eq_0) using as(2) unfolding disjoint_family_on_def
+ apply(erule_tac x=x in ballE,safe,erule_tac x=y in ballE) by auto
+ next fix m s assume "s \<in> A ` {..<m}"
+ hence "s \<in> range A" by auto thus "lmeasurable s" using as(1) by fastsimp
+ next fix m s t assume st:"s \<in> A ` {..<m}" "t \<in> A ` {..<m}" "s \<noteq> t"
+ from st(1-2) guess sa ta unfolding image_iff apply-by(erule bexE)+ note a=this
+ from st(3) have "sa \<noteq> ta" unfolding a by auto
+ thus "negligible (s \<inter> t)"
+ using as(2) unfolding disjoint_family_on_def a
+ apply(erule_tac x=sa in ballE,erule_tac x=ta in ballE) by auto
+ qed
+
+ have "\<And>m. lmeasurable (\<Union>A ` {..<m})" apply(rule lmeasurable_finite_unions)
+ apply(rule finite_imageI,rule) using as(1) by fastsimp
+ from this this show "(\<Sum>n<m. lmeasure (A n)) \<le> (\<Sum>n<n. lmeasure (A n))" unfolding *
+ apply(rule lmeasure_subset) apply(rule Union_mono) apply(rule image_mono) using mn by auto
+
+ next have *:"UNION UNIV A = \<Union>{A n |n. n \<in> UNIV}" by auto
+ show "(\<lambda>n. \<Sum>n<n. lmeasure (A n)) ----> lmeasure (UNION UNIV A)"
+ apply(rule LIMSEQ_imp_Suc) unfolding lessThan_Suc_atMost *
+ apply(rule has_lmeasure_countable_negligible_unions)
+ using as unfolding disjoint_family_on_def subset_eq by auto
+ qed
+
+next show "lmeasure {} = 0" by auto
+next fix A::"nat => _" assume as:"range A \<subseteq> sets lebesgue_space"
+ have *:"UNION UNIV A = (\<Union>{A n |n. n \<in> UNIV})" unfolding simple_image by auto
+ show "lmeasurable (UNION UNIV A)" unfolding * using as unfolding subset_eq
+ using lmeasurable_countable_unions_strong[of A] by auto
+qed(auto simp: lebesgue_space_def mem_def)
+
+
+
+lemma lmeasurbale_closed_interval[intro]:
+ "lmeasurable {a..b::'a::ordered_euclidean_space}"
+ unfolding lmeasurable_def cube_def inter_interval by auto
+
+lemma space_lebesgue_space[simp]:"space lebesgue_space = UNIV"
+ unfolding lebesgue_space_def by auto
+
+abbreviation "gintegral \<equiv> Integration.integral"
+
+lemma lebesgue_simple_function_indicator:
+ fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+ assumes f:"lebesgue.simple_function f"
+ shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
+ apply(rule,subst lebesgue.simple_function_indicator_representation[OF f]) by auto
+
+lemma lmeasure_gmeasure:
+ "gmeasurable s \<Longrightarrow> gmeasure s = real (lmeasure s)"
+ apply(subst gmeasure_lmeasure) by auto
+
+lemma lmeasure_finite: assumes "gmeasurable s" shows "lmeasure s \<noteq> \<omega>"
+ using gmeasure_lmeasure[OF assms] by auto
+
+lemma negligible_lmeasure: assumes "lmeasurable s"
+ shows "lmeasure s = 0 \<longleftrightarrow> negligible s" (is "?l = ?r")
+proof assume ?l
+ hence *:"gmeasurable s" using glmeasurable_finite[of s] assms by auto
+ show ?r unfolding gmeasurable_measure_eq_0[THEN sym,OF *]
+ unfolding lmeasure_gmeasure[OF *] using `?l` by auto
+next assume ?r
+ note g=negligible_img_gmeasurable[OF this] and measure_eq_0[OF this]
+ hence "real (lmeasure s) = 0" using lmeasure_gmeasure[of s] by auto
+ thus ?l using lmeasure_finite[OF g] apply- apply(rule real_0_imp_eq_0) by auto
+qed
+
+end
--- a/src/HOL/Probability/Measure.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Probability/Measure.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1,437 +1,120 @@
header {*Measures*}
theory Measure
- imports Caratheodory FuncSet
+ imports Caratheodory
begin
text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
-definition prod_sets where
- "prod_sets A B = {z. \<exists>x \<in> A. \<exists>y \<in> B. z = x \<times> y}"
-
-lemma prod_setsI: "x \<in> A \<Longrightarrow> y \<in> B \<Longrightarrow> (x \<times> y) \<in> prod_sets A B"
- by (auto simp add: prod_sets_def)
-
-lemma sigma_prod_sets_finite:
- assumes "finite A" and "finite B"
- shows "sets (sigma (A \<times> B) (prod_sets (Pow A) (Pow B))) = Pow (A \<times> B)"
-proof (simp add: sigma_def, safe)
- have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
-
- fix x assume subset: "x \<subseteq> A \<times> B"
- hence "finite x" using fin by (rule finite_subset)
- from this subset show "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))"
- (is "x \<in> sigma_sets ?prod ?sets")
- proof (induct x)
- case empty show ?case by (rule sigma_sets.Empty)
- next
- case (insert a x)
- hence "{a} \<in> sigma_sets ?prod ?sets" by (auto simp: prod_sets_def intro!: sigma_sets.Basic)
- moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
- ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
- qed
-next
- fix x a b
- assume "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))" and "(a, b) \<in> x"
- from sigma_sets_into_sp[OF _ this(1)] this(2)
- show "a \<in> A" and "b \<in> B"
- by (auto simp: prod_sets_def)
-qed
-
-
-definition
- closed_cdi where
- "closed_cdi M \<longleftrightarrow>
- sets M \<subseteq> Pow (space M) &
- (\<forall>s \<in> sets M. space M - s \<in> sets M) &
- (\<forall>A. (range A \<subseteq> sets M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
- (\<Union>i. A i) \<in> sets M) &
- (\<forall>A. (range A \<subseteq> sets M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
-
-
-inductive_set
- smallest_ccdi_sets :: "('a, 'b) algebra_scheme \<Rightarrow> 'a set set"
- for M
- where
- Basic [intro]:
- "a \<in> sets M \<Longrightarrow> a \<in> smallest_ccdi_sets M"
- | Compl [intro]:
- "a \<in> smallest_ccdi_sets M \<Longrightarrow> space M - a \<in> smallest_ccdi_sets M"
- | Inc:
- "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
- \<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets M"
- | Disj:
- "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> disjoint_family A
- \<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets M"
- monos Pow_mono
-
-
-definition
- smallest_closed_cdi where
- "smallest_closed_cdi M = (|space = space M, sets = smallest_ccdi_sets M|)"
-
-definition
- measurable where
- "measurable a b = {f . sigma_algebra a & sigma_algebra b &
- f \<in> (space a -> space b) &
- (\<forall>y \<in> sets b. (f -` y) \<inter> (space a) \<in> sets a)}"
+section {* Equations for the measure function @{text \<mu>} *}
-definition
- measure_preserving where
- "measure_preserving m1 m2 =
- measurable m1 m2 \<inter>
- {f . measure_space m1 & measure_space m2 &
- (\<forall>y \<in> sets m2. measure m1 ((f -` y) \<inter> (space m1)) = measure m2 y)}"
-
-lemma space_smallest_closed_cdi [simp]:
- "space (smallest_closed_cdi M) = space M"
- by (simp add: smallest_closed_cdi_def)
-
-
-lemma (in algebra) smallest_closed_cdi1: "sets M \<subseteq> sets (smallest_closed_cdi M)"
- by (auto simp add: smallest_closed_cdi_def)
-
-lemma (in algebra) smallest_ccdi_sets:
- "smallest_ccdi_sets M \<subseteq> Pow (space M)"
- apply (rule subsetI)
- apply (erule smallest_ccdi_sets.induct)
- apply (auto intro: range_subsetD dest: sets_into_space)
- done
-
-lemma (in algebra) smallest_closed_cdi2: "closed_cdi (smallest_closed_cdi M)"
- apply (auto simp add: closed_cdi_def smallest_closed_cdi_def smallest_ccdi_sets)
- apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) +
- done
-
-lemma (in algebra) smallest_closed_cdi3:
- "sets (smallest_closed_cdi M) \<subseteq> Pow (space M)"
- by (simp add: smallest_closed_cdi_def smallest_ccdi_sets)
-
-lemma closed_cdi_subset: "closed_cdi M \<Longrightarrow> sets M \<subseteq> Pow (space M)"
- by (simp add: closed_cdi_def)
-
-lemma closed_cdi_Compl: "closed_cdi M \<Longrightarrow> s \<in> sets M \<Longrightarrow> space M - s \<in> sets M"
- by (simp add: closed_cdi_def)
-
-lemma closed_cdi_Inc:
- "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow>
- (\<Union>i. A i) \<in> sets M"
- by (simp add: closed_cdi_def)
-
-lemma closed_cdi_Disj:
- "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
- by (simp add: closed_cdi_def)
-
-lemma closed_cdi_Un:
- assumes cdi: "closed_cdi M" and empty: "{} \<in> sets M"
- and A: "A \<in> sets M" and B: "B \<in> sets M"
- and disj: "A \<inter> B = {}"
- shows "A \<union> B \<in> sets M"
+lemma (in measure_space) measure_countably_additive:
+ assumes "range A \<subseteq> sets M" "disjoint_family A"
+ shows "psuminf (\<lambda>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
proof -
- have ra: "range (binaryset A B) \<subseteq> sets M"
- by (simp add: range_binaryset_eq empty A B)
- have di: "disjoint_family (binaryset A B)" using disj
- by (simp add: disjoint_family_on_def binaryset_def Int_commute)
- from closed_cdi_Disj [OF cdi ra di]
- show ?thesis
- by (simp add: UN_binaryset_eq)
-qed
-
-lemma (in algebra) smallest_ccdi_sets_Un:
- assumes A: "A \<in> smallest_ccdi_sets M" and B: "B \<in> smallest_ccdi_sets M"
- and disj: "A \<inter> B = {}"
- shows "A \<union> B \<in> smallest_ccdi_sets M"
-proof -
- have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets M)"
- by (simp add: range_binaryset_eq A B smallest_ccdi_sets.Basic)
- have di: "disjoint_family (binaryset A B)" using disj
- by (simp add: disjoint_family_on_def binaryset_def Int_commute)
- from Disj [OF ra di]
- show ?thesis
- by (simp add: UN_binaryset_eq)
+ have "(\<Union> i. A i) \<in> sets M" using assms(1) by (rule countable_UN)
+ with ca assms show ?thesis by (simp add: countably_additive_def)
qed
-
-
-lemma (in algebra) smallest_ccdi_sets_Int1:
- assumes a: "a \<in> sets M"
- shows "b \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M"
-proof (induct rule: smallest_ccdi_sets.induct)
- case (Basic x)
- thus ?case
- by (metis a Int smallest_ccdi_sets.Basic)
-next
- case (Compl x)
- have "a \<inter> (space M - x) = space M - ((space M - a) \<union> (a \<inter> x))"
- by blast
- also have "... \<in> smallest_ccdi_sets M"
- by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2
- Diff_disjoint Int_Diff Int_empty_right Un_commute
- smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl
- smallest_ccdi_sets_Un)
- finally show ?case .
-next
- case (Inc A)
- have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
- by blast
- have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Inc
- by blast
- moreover have "(\<lambda>i. a \<inter> A i) 0 = {}"
- by (simp add: Inc)
- moreover have "!!n. (\<lambda>i. a \<inter> A i) n \<subseteq> (\<lambda>i. a \<inter> A i) (Suc n)" using Inc
- by blast
- ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M"
- by (rule smallest_ccdi_sets.Inc)
- show ?case
- by (metis 1 2)
-next
- case (Disj A)
- have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
- by blast
- have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Disj
- by blast
- moreover have "disjoint_family (\<lambda>i. a \<inter> A i)" using Disj
- by (auto simp add: disjoint_family_on_def)
- ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M"
- by (rule smallest_ccdi_sets.Disj)
- show ?case
- by (metis 1 2)
-qed
-
-
-lemma (in algebra) smallest_ccdi_sets_Int:
- assumes b: "b \<in> smallest_ccdi_sets M"
- shows "a \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M"
-proof (induct rule: smallest_ccdi_sets.induct)
- case (Basic x)
- thus ?case
- by (metis b smallest_ccdi_sets_Int1)
-next
- case (Compl x)
- have "(space M - x) \<inter> b = space M - (x \<inter> b \<union> (space M - b))"
- by blast
- also have "... \<in> smallest_ccdi_sets M"
- by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b
- smallest_ccdi_sets.Compl smallest_ccdi_sets_Un)
- finally show ?case .
-next
- case (Inc A)
- have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
- by blast
- have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Inc
- by blast
- moreover have "(\<lambda>i. A i \<inter> b) 0 = {}"
- by (simp add: Inc)
- moreover have "!!n. (\<lambda>i. A i \<inter> b) n \<subseteq> (\<lambda>i. A i \<inter> b) (Suc n)" using Inc
- by blast
- ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M"
- by (rule smallest_ccdi_sets.Inc)
- show ?case
- by (metis 1 2)
-next
- case (Disj A)
- have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
- by blast
- have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Disj
- by blast
- moreover have "disjoint_family (\<lambda>i. A i \<inter> b)" using Disj
- by (auto simp add: disjoint_family_on_def)
- ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M"
- by (rule smallest_ccdi_sets.Disj)
- show ?case
- by (metis 1 2)
+lemma (in measure_space) measure_space_cong:
+ assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A"
+ shows "measure_space M \<nu>"
+proof
+ show "\<nu> {} = 0" using assms by auto
+ show "countably_additive M \<nu>" unfolding countably_additive_def
+ proof safe
+ fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A"
+ then have "\<And>i. A i \<in> sets M" "(UNION UNIV A) \<in> sets M" by auto
+ from this[THEN assms] measure_countably_additive[OF A]
+ show "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (UNION UNIV A)" by simp
+ qed
qed
-lemma (in algebra) sets_smallest_closed_cdi_Int:
- "a \<in> sets (smallest_closed_cdi M) \<Longrightarrow> b \<in> sets (smallest_closed_cdi M)
- \<Longrightarrow> a \<inter> b \<in> sets (smallest_closed_cdi M)"
- by (simp add: smallest_ccdi_sets_Int smallest_closed_cdi_def)
-
-lemma algebra_iff_Int:
- "algebra M \<longleftrightarrow>
- sets M \<subseteq> Pow (space M) & {} \<in> sets M &
- (\<forall>a \<in> sets M. space M - a \<in> sets M) &
- (\<forall>a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)"
-proof (auto simp add: algebra.Int, auto simp add: algebra_def)
- fix a b
- assume ab: "sets M \<subseteq> Pow (space M)"
- "\<forall>a\<in>sets M. space M - a \<in> sets M"
- "\<forall>a\<in>sets M. \<forall>b\<in>sets M. a \<inter> b \<in> sets M"
- "a \<in> sets M" "b \<in> sets M"
- hence "a \<union> b = space M - ((space M - a) \<inter> (space M - b))"
- by blast
- also have "... \<in> sets M"
- by (metis ab)
- finally show "a \<union> b \<in> sets M" .
+lemma (in measure_space) additive: "additive M \<mu>"
+proof (rule algebra.countably_additive_additive [OF _ _ ca])
+ show "algebra M" by default
+ show "positive \<mu>" by (simp add: positive_def)
qed
-lemma (in algebra) sigma_property_disjoint_lemma:
- assumes sbC: "sets M \<subseteq> C"
- and ccdi: "closed_cdi (|space = space M, sets = C|)"
- shows "sigma_sets (space M) (sets M) \<subseteq> C"
-proof -
- have "smallest_ccdi_sets M \<in> {B . sets M \<subseteq> B \<and> sigma_algebra (|space = space M, sets = B|)}"
- apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int
- smallest_ccdi_sets_Int)
- apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets)
- apply (blast intro: smallest_ccdi_sets.Disj)
- done
- hence "sigma_sets (space M) (sets M) \<subseteq> smallest_ccdi_sets M"
- by clarsimp
- (drule sigma_algebra.sigma_sets_subset [where a="sets M"], auto)
- also have "... \<subseteq> C"
- proof
- fix x
- assume x: "x \<in> smallest_ccdi_sets M"
- thus "x \<in> C"
- proof (induct rule: smallest_ccdi_sets.induct)
- case (Basic x)
- thus ?case
- by (metis Basic subsetD sbC)
- next
- case (Compl x)
- thus ?case
- by (blast intro: closed_cdi_Compl [OF ccdi, simplified])
- next
- case (Inc A)
- thus ?case
- by (auto intro: closed_cdi_Inc [OF ccdi, simplified])
- next
- case (Disj A)
- thus ?case
- by (auto intro: closed_cdi_Disj [OF ccdi, simplified])
- qed
- qed
- finally show ?thesis .
-qed
-
-lemma (in algebra) sigma_property_disjoint:
- assumes sbC: "sets M \<subseteq> C"
- and compl: "!!s. s \<in> C \<inter> sigma_sets (space M) (sets M) \<Longrightarrow> space M - s \<in> C"
- and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M)
- \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n))
- \<Longrightarrow> (\<Union>i. A i) \<in> C"
- and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M)
- \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> C"
- shows "sigma_sets (space M) (sets M) \<subseteq> C"
-proof -
- have "sigma_sets (space M) (sets M) \<subseteq> C \<inter> sigma_sets (space M) (sets M)"
- proof (rule sigma_property_disjoint_lemma)
- show "sets M \<subseteq> C \<inter> sigma_sets (space M) (sets M)"
- by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic)
- next
- show "closed_cdi \<lparr>space = space M, sets = C \<inter> sigma_sets (space M) (sets M)\<rparr>"
- by (simp add: closed_cdi_def compl inc disj)
- (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed
- IntE sigma_sets.Compl range_subsetD sigma_sets.Union)
- qed
- thus ?thesis
- by blast
-qed
-
-
-(* or just countably_additiveD [OF measure_space.ca] *)
-lemma (in measure_space) measure_countably_additive:
- "range A \<subseteq> sets M
- \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> sets M
- \<Longrightarrow> (measure M o A) sums measure M (\<Union>i. A i)"
- by (insert ca) (simp add: countably_additive_def o_def)
-
-lemma (in measure_space) additive:
- "additive M (measure M)"
-proof (rule algebra.countably_additive_additive [OF _ _ ca])
- show "algebra M"
- by (blast intro: sigma_algebra.axioms local.sigma_algebra_axioms)
- show "positive M (measure M)"
- by (simp add: positive_def positive)
-qed
-
lemma (in measure_space) measure_additive:
- "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {}
- \<Longrightarrow> measure M a + measure M b = measure M (a \<union> b)"
+ "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {}
+ \<Longrightarrow> \<mu> a + \<mu> b = \<mu> (a \<union> b)"
by (metis additiveD additive)
-lemma (in measure_space) measure_compl:
- assumes s: "s \<in> sets M"
- shows "measure M (space M - s) = measure M (space M) - measure M s"
-proof -
- have "measure M (space M) = measure M (s \<union> (space M - s))" using s
- by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
- also have "... = measure M s + measure M (space M - s)"
- by (rule additiveD [OF additive]) (auto simp add: s)
- finally have "measure M (space M) = measure M s + measure M (space M - s)" .
- thus ?thesis
- by arith
-qed
-
lemma (in measure_space) measure_mono:
assumes "a \<subseteq> b" "a \<in> sets M" "b \<in> sets M"
- shows "measure M a \<le> measure M b"
+ shows "\<mu> a \<le> \<mu> b"
proof -
have "b = a \<union> (b - a)" using assms by auto
moreover have "{} = a \<inter> (b - a)" by auto
- ultimately have "measure M b = measure M a + measure M (b - a)"
+ ultimately have "\<mu> b = \<mu> a + \<mu> (b - a)"
using measure_additive[of a "b - a"] local.Diff[of b a] assms by auto
- moreover have "measure M (b - a) \<ge> 0" using positive assms by auto
- ultimately show "measure M a \<le> measure M b" by auto
+ moreover have "\<mu> (b - a) \<ge> 0" using assms by auto
+ ultimately show "\<mu> a \<le> \<mu> b" by auto
qed
-lemma disjoint_family_Suc:
- assumes Suc: "!!n. A n \<subseteq> A (Suc n)"
- shows "disjoint_family (\<lambda>i. A (Suc i) - A i)"
+lemma (in measure_space) measure_compl:
+ assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<omega>"
+ shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
proof -
- {
- fix m
- have "!!n. A n \<subseteq> A (m+n)"
- proof (induct m)
- case 0 show ?case by simp
- next
- case (Suc m) thus ?case
- by (metis Suc_eq_plus1 assms nat_add_commute nat_add_left_commute subset_trans)
- qed
- }
- hence "!!m n. m < n \<Longrightarrow> A m \<subseteq> A n"
- by (metis add_commute le_add_diff_inverse nat_less_le)
+ have s_less_space: "\<mu> s \<le> \<mu> (space M)"
+ using s by (auto intro!: measure_mono sets_into_space)
+
+ have "\<mu> (space M) = \<mu> (s \<union> (space M - s))" using s
+ by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
+ also have "... = \<mu> s + \<mu> (space M - s)"
+ by (rule additiveD [OF additive]) (auto simp add: s)
+ finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" .
thus ?thesis
- by (auto simp add: disjoint_family_on_def)
- (metis insert_absorb insert_subset le_SucE le_antisym not_leE)
+ unfolding minus_pinfreal_eq2[OF s_less_space fin]
+ by (simp add: ac_simps)
qed
+lemma (in measure_space) measure_Diff:
+ assumes finite: "\<mu> B \<noteq> \<omega>"
+ and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
+ shows "\<mu> (A - B) = \<mu> A - \<mu> B"
+proof -
+ have *: "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
+
+ have "\<mu> ((A - B) \<union> B) = \<mu> (A - B) + \<mu> B"
+ using measurable by (rule_tac measure_additive[symmetric]) auto
+ thus ?thesis unfolding * using `\<mu> B \<noteq> \<omega>`
+ by (simp add: pinfreal_cancel_plus_minus)
+qed
lemma (in measure_space) measure_countable_increasing:
assumes A: "range A \<subseteq> sets M"
and A0: "A 0 = {}"
- and ASuc: "!!n. A n \<subseteq> A (Suc n)"
- shows "(measure M o A) ----> measure M (\<Union>i. A i)"
+ and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
+ shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
proof -
{
fix n
- have "(measure M \<circ> A) n =
- setsum (measure M \<circ> (\<lambda>i. A (Suc i) - A i)) {0..<n}"
+ have "\<mu> (A n) =
+ setsum (\<mu> \<circ> (\<lambda>i. A (Suc i) - A i)) {..<n}"
proof (induct n)
case 0 thus ?case by (auto simp add: A0)
next
- case (Suc m)
+ case (Suc m)
have "A (Suc m) = A m \<union> (A (Suc m) - A m)"
by (metis ASuc Un_Diff_cancel Un_absorb1)
- hence "measure M (A (Suc m)) =
- measure M (A m) + measure M (A (Suc m) - A m)"
- by (subst measure_additive)
- (auto simp add: measure_additive range_subsetD [OF A])
+ hence "\<mu> (A (Suc m)) =
+ \<mu> (A m) + \<mu> (A (Suc m) - A m)"
+ by (subst measure_additive)
+ (auto simp add: measure_additive range_subsetD [OF A])
with Suc show ?case
by simp
- qed
- }
- hence Meq: "measure M o A = (\<lambda>n. setsum (measure M o (\<lambda>i. A(Suc i) - A i)) {0..<n})"
- by (blast intro: ext)
+ qed }
+ note Meq = this
have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)"
- proof (rule UN_finite2_eq [where k=1], simp)
+ proof (rule UN_finite2_eq [where k=1], simp)
fix i
show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
proof (induct i)
case 0 thus ?case by (simp add: A0)
next
- case (Suc i)
+ case (Suc i)
thus ?case
by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc])
qed
@@ -440,454 +123,133 @@
by (metis A Diff range_subsetD)
have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
by (blast intro: range_subsetD [OF A])
- have "(measure M o (\<lambda>i. A (Suc i) - A i)) sums measure M (\<Union>i. A (Suc i) - A i)"
- by (rule measure_countably_additive)
+ have "psuminf ( (\<lambda>i. \<mu> (A (Suc i) - A i))) = \<mu> (\<Union>i. A (Suc i) - A i)"
+ by (rule measure_countably_additive)
(auto simp add: disjoint_family_Suc ASuc A1 A2)
- also have "... = measure M (\<Union>i. A i)"
- by (simp add: Aeq)
- finally have "(measure M o (\<lambda>i. A (Suc i) - A i)) sums measure M (\<Union>i. A i)" .
+ also have "... = \<mu> (\<Union>i. A i)"
+ by (simp add: Aeq)
+ finally have "psuminf (\<lambda>i. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" .
thus ?thesis
- by (auto simp add: sums_iff Meq dest: summable_sumr_LIMSEQ_suminf)
-qed
-
-lemma (in measure_space) monotone_convergence:
- assumes A: "range A \<subseteq> sets M"
- and ASuc: "!!n. A n \<subseteq> A (Suc n)"
- shows "(measure M \<circ> A) ----> measure M (\<Union>i. A i)"
-proof -
- have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
- by (auto simp add: split: nat.splits)
- have meq: "measure M \<circ> A = (\<lambda>n. (measure M \<circ> nat_case {} A) (Suc n))"
- by (rule ext) simp
- have "(measure M \<circ> nat_case {} A) ----> measure M (\<Union>i. nat_case {} A i)"
- by (rule measure_countable_increasing)
- (auto simp add: range_subsetD [OF A] subsetD [OF ASuc] split: nat.splits)
- also have "measure M (\<Union>i. nat_case {} A i) = measure M (\<Union>i. A i)"
- by (simp add: ueq)
- finally have "(measure M \<circ> nat_case {} A) ----> measure M (\<Union>i. A i)" .
- thus ?thesis using meq
- by (metis LIMSEQ_Suc)
+ by (auto simp add: Meq psuminf_def)
qed
-lemma measurable_sigma_preimages:
- assumes a: "sigma_algebra a" and b: "sigma_algebra b"
- and f: "f \<in> space a -> space b"
- and ba: "!!y. y \<in> sets b ==> f -` y \<in> sets a"
- shows "sigma_algebra (|space = space a, sets = (vimage f) ` (sets b) |)"
-proof (simp add: sigma_algebra_iff2, intro conjI)
- show "op -` f ` sets b \<subseteq> Pow (space a)"
- by auto (metis IntE a algebra.Int_space_eq1 ba sigma_algebra_iff vimageI)
-next
- show "{} \<in> op -` f ` sets b"
- by (metis algebra.empty_sets b image_iff sigma_algebra_iff vimage_empty)
-next
- { fix y
- assume y: "y \<in> sets b"
- with a b f
- have "space a - f -` y = f -` (space b - y)"
- by (auto, clarsimp simp add: sigma_algebra_iff2)
- (drule ba, blast intro: ba)
- hence "space a - (f -` y) \<in> (vimage f) ` sets b"
- by auto
- (metis b y subsetD equalityE imageI sigma_algebra.sigma_sets_eq
- sigma_sets.Compl)
- }
- thus "\<forall>s\<in>sets b. space a - f -` s \<in> (vimage f) ` sets b"
- by blast
-next
- {
- fix A:: "nat \<Rightarrow> 'a set"
- assume A: "range A \<subseteq> (vimage f) ` (sets b)"
- have "(\<Union>i. A i) \<in> (vimage f) ` (sets b)"
- proof -
- def B \<equiv> "\<lambda>i. @v. v \<in> sets b \<and> f -` v = A i"
- {
- fix i
- have "A i \<in> (vimage f) ` (sets b)" using A
- by blast
- hence "\<exists>v. v \<in> sets b \<and> f -` v = A i"
- by blast
- hence "B i \<in> sets b \<and> f -` B i = A i"
- by (simp add: B_def) (erule someI_ex)
- } note B = this
- show ?thesis
- proof
- show "(\<Union>i. A i) = f -` (\<Union>i. B i)"
- by (simp add: vimage_UN B)
- show "(\<Union>i. B i) \<in> sets b" using B
- by (auto intro: sigma_algebra.countable_UN [OF b])
- qed
- qed
- }
- thus "\<forall>A::nat \<Rightarrow> 'a set.
- range A \<subseteq> (vimage f) ` sets b \<longrightarrow> (\<Union>i. A i) \<in> (vimage f) ` sets b"
- by blast
-qed
+lemma (in measure_space) continuity_from_below:
+ assumes A: "range A \<subseteq> sets M"
+ and ASuc: "!!n. A n \<subseteq> A (Suc n)"
+ shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
+proof -
+ have *: "(SUP n. \<mu> (nat_case {} A (Suc n))) = (SUP n. \<mu> (nat_case {} A n))"
+ apply (rule Sup_mono_offset_Suc)
+ apply (rule measure_mono)
+ using assms by (auto split: nat.split)
-lemma (in sigma_algebra) measurable_sigma:
- assumes B: "B \<subseteq> Pow X"
- and f: "f \<in> space M -> X"
- and ba: "!!y. y \<in> B ==> (f -` y) \<inter> space M \<in> sets M"
- shows "f \<in> measurable M (sigma X B)"
-proof -
- have "sigma_sets X B \<subseteq> {y . (f -` y) \<inter> space M \<in> sets M & y \<subseteq> X}"
- proof clarify
- fix x
- assume "x \<in> sigma_sets X B"
- thus "f -` x \<inter> space M \<in> sets M \<and> x \<subseteq> X"
- proof induct
- case (Basic a)
- thus ?case
- by (auto simp add: ba) (metis B subsetD PowD)
- next
- case Empty
- thus ?case
- by auto
- next
- case (Compl a)
- have [simp]: "f -` X \<inter> space M = space M"
- by (auto simp add: funcset_mem [OF f])
- thus ?case
- by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl)
- next
- case (Union a)
- thus ?case
- by (simp add: vimage_UN, simp only: UN_extend_simps(4))
- (blast intro: countable_UN)
- qed
- qed
- thus ?thesis
- by (simp add: measurable_def sigma_algebra_axioms sigma_algebra_sigma B f)
- (auto simp add: sigma_def)
+ have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
+ by (auto simp add: split: nat.splits)
+ have meq: "\<And>n. \<mu> (A n) = (\<mu> \<circ> nat_case {} A) (Suc n)"
+ by simp
+ have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. nat_case {} A i)"
+ by (rule measure_countable_increasing)
+ (auto simp add: range_subsetD [OF A] subsetD [OF ASuc] split: nat.splits)
+ also have "\<mu> (\<Union>i. nat_case {} A i) = \<mu> (\<Union>i. A i)"
+ by (simp add: ueq)
+ finally have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. A i)" .
+ thus ?thesis unfolding meq * comp_def .
qed
-lemma measurable_subset:
- "measurable a b \<subseteq> measurable a (sigma (space b) (sets b))"
- apply clarify
- apply (rule sigma_algebra.measurable_sigma)
- apply (auto simp add: measurable_def)
- apply (metis algebra.sets_into_space subsetD sigma_algebra_iff)
- done
+lemma (in measure_space) measure_up:
+ assumes "P \<in> sets M" "\<And>i. B i \<in> sets M" "B \<up> P"
+ shows "(\<lambda>i. \<mu> (B i)) \<up> \<mu> P"
+ using assms unfolding isoton_def
+ by (auto intro!: measure_mono continuity_from_below)
-lemma measurable_eqI:
- "space m1 = space m1' \<Longrightarrow> space m2 = space m2'
- \<Longrightarrow> sets m1 = sets m1' \<Longrightarrow> sets m2 = sets m2'
- \<Longrightarrow> measurable m1 m2 = measurable m1' m2'"
- by (simp add: measurable_def sigma_algebra_iff2)
-lemma measure_preserving_lift:
- fixes f :: "'a1 \<Rightarrow> 'a2"
- and m1 :: "('a1, 'b1) measure_space_scheme"
- and m2 :: "('a2, 'b2) measure_space_scheme"
- assumes m1: "measure_space m1" and m2: "measure_space m2"
- defines "m x \<equiv> (|space = space m2, sets = x, measure = measure m2, ... = more m2|)"
- assumes setsm2: "sets m2 = sigma_sets (space m2) a"
- and fmp: "f \<in> measure_preserving m1 (m a)"
- shows "f \<in> measure_preserving m1 m2"
+lemma (in measure_space) continuity_from_above:
+ assumes A: "range A \<subseteq> sets M"
+ and mono_Suc: "\<And>n. A (Suc n) \<subseteq> A n"
+ and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ shows "(INF n. \<mu> (A n)) = \<mu> (\<Inter>i. A i)"
proof -
- have [simp]: "!!x. space (m x) = space m2 & sets (m x) = x & measure (m x) = measure m2"
- by (simp add: m_def)
- have sa1: "sigma_algebra m1" using m1
- by (simp add: measure_space_def)
- show ?thesis using fmp
- proof (clarsimp simp add: measure_preserving_def m1 m2)
- assume fm: "f \<in> measurable m1 (m a)"
- and mam: "measure_space (m a)"
- and meq: "\<forall>y\<in>a. measure m1 (f -` y \<inter> space m1) = measure m2 y"
- have "f \<in> measurable m1 (sigma (space (m a)) (sets (m a)))"
- by (rule subsetD [OF measurable_subset fm])
- also have "... = measurable m1 m2"
- by (rule measurable_eqI) (simp_all add: m_def setsm2 sigma_def)
- finally have f12: "f \<in> measurable m1 m2" .
- hence ffn: "f \<in> space m1 \<rightarrow> space m2"
- by (simp add: measurable_def)
- have "space m1 \<subseteq> f -` (space m2)"
- by auto (metis PiE ffn)
- hence fveq [simp]: "(f -` (space m2)) \<inter> space m1 = space m1"
- by blast
- {
- fix y
- assume y: "y \<in> sets m2"
- have ama: "algebra (m a)" using mam
- by (simp add: measure_space_def sigma_algebra_iff)
- have spa: "space m2 \<in> a" using algebra.top [OF ama]
- by (simp add: m_def)
- have "sigma_sets (space m2) a = sigma_sets (space (m a)) (sets (m a))"
- by (simp add: m_def)
- also have "... \<subseteq> {s . measure m1 ((f -` s) \<inter> space m1) = measure m2 s}"
- proof (rule algebra.sigma_property_disjoint, auto simp add: ama)
- fix x
- assume x: "x \<in> a"
- thus "measure m1 (f -` x \<inter> space m1) = measure m2 x"
- by (simp add: meq)
- next
- fix s
- assume eq: "measure m1 (f -` s \<inter> space m1) = measure m2 s"
- and s: "s \<in> sigma_sets (space m2) a"
- hence s2: "s \<in> sets m2"
- by (simp add: setsm2)
- thus "measure m1 (f -` (space m2 - s) \<inter> space m1) =
- measure m2 (space m2 - s)" using f12
- by (simp add: vimage_Diff Diff_Int_distrib2 eq m1 m2
- measure_space.measure_compl measurable_def)
- (metis fveq meq spa)
- next
- fix A
- assume A0: "A 0 = {}"
- and ASuc: "!!n. A n \<subseteq> A (Suc n)"
- and rA1: "range A \<subseteq>
- {s. measure m1 (f -` s \<inter> space m1) = measure m2 s}"
- and "range A \<subseteq> sigma_sets (space m2) a"
- hence rA2: "range A \<subseteq> sets m2" by (metis setsm2)
- have eq1: "!!i. measure m1 (f -` A i \<inter> space m1) = measure m2 (A i)"
- using rA1 by blast
- have "(measure m2 \<circ> A) = measure m1 o (\<lambda>i. (f -` A i \<inter> space m1))"
- by (simp add: o_def eq1)
- also have "... ----> measure m1 (\<Union>i. f -` A i \<inter> space m1)"
- proof (rule measure_space.measure_countable_increasing [OF m1])
- show "range (\<lambda>i. f -` A i \<inter> space m1) \<subseteq> sets m1"
- using f12 rA2 by (auto simp add: measurable_def)
- show "f -` A 0 \<inter> space m1 = {}" using A0
- by blast
- show "\<And>n. f -` A n \<inter> space m1 \<subseteq> f -` A (Suc n) \<inter> space m1"
- using ASuc by auto
- qed
- also have "measure m1 (\<Union>i. f -` A i \<inter> space m1) = measure m1 (f -` (\<Union>i. A i) \<inter> space m1)"
- by (simp add: vimage_UN)
- finally have "(measure m2 \<circ> A) ----> measure m1 (f -` (\<Union>i. A i) \<inter> space m1)" .
- moreover
- have "(measure m2 \<circ> A) ----> measure m2 (\<Union>i. A i)"
- by (rule measure_space.measure_countable_increasing
- [OF m2 rA2, OF A0 ASuc])
- ultimately
- show "measure m1 (f -` (\<Union>i. A i) \<inter> space m1) =
- measure m2 (\<Union>i. A i)"
- by (rule LIMSEQ_unique)
- next
- fix A :: "nat => 'a2 set"
- assume dA: "disjoint_family A"
- and rA1: "range A \<subseteq>
- {s. measure m1 (f -` s \<inter> space m1) = measure m2 s}"
- and "range A \<subseteq> sigma_sets (space m2) a"
- hence rA2: "range A \<subseteq> sets m2" by (metis setsm2)
- hence Um2: "(\<Union>i. A i) \<in> sets m2"
- by (metis range_subsetD setsm2 sigma_sets.Union)
- have "!!i. measure m1 (f -` A i \<inter> space m1) = measure m2 (A i)"
- using rA1 by blast
- hence "measure m2 o A = measure m1 o (\<lambda>i. f -` A i \<inter> space m1)"
- by (simp add: o_def)
- also have "... sums measure m1 (\<Union>i. f -` A i \<inter> space m1)"
- proof (rule measure_space.measure_countably_additive [OF m1] )
- show "range (\<lambda>i. f -` A i \<inter> space m1) \<subseteq> sets m1"
- using f12 rA2 by (auto simp add: measurable_def)
- show "disjoint_family (\<lambda>i. f -` A i \<inter> space m1)" using dA
- by (auto simp add: disjoint_family_on_def)
- show "(\<Union>i. f -` A i \<inter> space m1) \<in> sets m1"
- proof (rule sigma_algebra.countable_UN [OF sa1])
- show "range (\<lambda>i. f -` A i \<inter> space m1) \<subseteq> sets m1" using f12 rA2
- by (auto simp add: measurable_def)
- qed
- qed
- finally have "(measure m2 o A) sums measure m1 (\<Union>i. f -` A i \<inter> space m1)" .
- with measure_space.measure_countably_additive [OF m2 rA2 dA Um2]
- have "measure m1 (\<Union>i. f -` A i \<inter> space m1) = measure m2 (\<Union>i. A i)"
- by (metis sums_unique)
+ { fix n have "A n \<subseteq> A 0" using mono_Suc by (induct n) auto }
+ note mono = this
+
+ have le_MI: "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
+ using A by (auto intro!: measure_mono)
+ hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<omega>" using finite[of 0] by auto
+
+ have le_IM: "(INF n. \<mu> (A n)) \<le> \<mu> (A 0)"
+ by (rule INF_leI) simp
- moreover have "measure m1 (f -` (\<Union>i. A i) \<inter> space m1) = measure m1 (\<Union>i. f -` A i \<inter> space m1)"
- by (simp add: vimage_UN)
- ultimately show "measure m1 (f -` (\<Union>i. A i) \<inter> space m1) =
- measure m2 (\<Union>i. A i)"
- by metis
- qed
- finally have "sigma_sets (space m2) a
- \<subseteq> {s . measure m1 ((f -` s) \<inter> space m1) = measure m2 s}" .
- hence "measure m1 (f -` y \<inter> space m1) = measure m2 y" using y
- by (force simp add: setsm2)
- }
- thus "f \<in> measurable m1 m2 \<and>
- (\<forall>y\<in>sets m2.
- measure m1 (f -` y \<inter> space m1) = measure m2 y)"
- by (blast intro: f12)
- qed
+ have "\<mu> (A 0) - (INF n. \<mu> (A n)) = (SUP n. \<mu> (A 0 - A n))"
+ unfolding pinfreal_SUP_minus[symmetric]
+ using mono A finite by (subst measure_Diff) auto
+ also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)"
+ proof (rule continuity_from_below)
+ show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
+ using A by auto
+ show "\<And>n. A 0 - A n \<subseteq> A 0 - A (Suc n)"
+ using mono_Suc by auto
+ qed
+ also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)"
+ using mono A finite * by (simp, subst measure_Diff) auto
+ finally show ?thesis
+ by (rule pinfreal_diff_eq_diff_imp_eq[OF finite[of 0] le_IM le_MI])
qed
-lemma measurable_ident:
- "sigma_algebra M ==> id \<in> measurable M M"
- apply (simp add: measurable_def)
- apply (auto simp add: sigma_algebra_def algebra.Int algebra.top)
- done
-
-lemma measurable_comp:
- fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
- shows "f \<in> measurable a b \<Longrightarrow> g \<in> measurable b c \<Longrightarrow> (g o f) \<in> measurable a c"
- apply (auto simp add: measurable_def vimage_compose)
- apply (subgoal_tac "f -` g -` y \<inter> space a = f -` (g -` y \<inter> space b) \<inter> space a")
- apply force+
- done
-
- lemma measurable_strong:
- fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
- assumes f: "f \<in> measurable a b" and g: "g \<in> (space b -> space c)"
- and c: "sigma_algebra c"
- and t: "f ` (space a) \<subseteq> t"
- and cb: "\<And>s. s \<in> sets c \<Longrightarrow> (g -` s) \<inter> t \<in> sets b"
- shows "(g o f) \<in> measurable a c"
+lemma (in measure_space) measure_down:
+ assumes "P \<in> sets M" "\<And>i. B i \<in> sets M" "B \<down> P"
+ and finite: "\<And>i. \<mu> (B i) \<noteq> \<omega>"
+ shows "(\<lambda>i. \<mu> (B i)) \<down> \<mu> P"
+ using assms unfolding antiton_def
+ by (auto intro!: measure_mono continuity_from_above)
+lemma (in measure_space) measure_insert:
+ assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
+ shows "\<mu> (insert x A) = \<mu> {x} + \<mu> A"
proof -
- have a: "sigma_algebra a" and b: "sigma_algebra b"
- and fab: "f \<in> (space a -> space b)"
- and ba: "\<And>y. y \<in> sets b \<Longrightarrow> (f -` y) \<inter> (space a) \<in> sets a" using f
- by (auto simp add: measurable_def)
- have eq: "f -` g -` y \<inter> space a = f -` (g -` y \<inter> t) \<inter> space a" using t
- by force
- show ?thesis
- apply (auto simp add: measurable_def vimage_compose a c)
- apply (metis funcset_mem fab g)
- apply (subst eq, metis ba cb)
- done
-qed
-
-lemma measurable_mono1:
- "a \<subseteq> b \<Longrightarrow> sigma_algebra \<lparr>space = X, sets = b\<rparr>
- \<Longrightarrow> measurable \<lparr>space = X, sets = a\<rparr> c \<subseteq> measurable \<lparr>space = X, sets = b\<rparr> c"
- by (auto simp add: measurable_def)
-
-lemma measurable_up_sigma:
- "measurable a b \<subseteq> measurable (sigma (space a) (sets a)) b"
- apply (auto simp add: measurable_def)
- apply (metis sigma_algebra_iff2 sigma_algebra_sigma)
- apply (metis algebra.simps(2) sigma_algebra.sigma_sets_eq sigma_def)
- done
-
-lemma measure_preserving_up:
- "f \<in> measure_preserving \<lparr>space = space m1, sets = a, measure = measure m1\<rparr> m2 \<Longrightarrow>
- measure_space m1 \<Longrightarrow> sigma_algebra m1 \<Longrightarrow> a \<subseteq> sets m1
- \<Longrightarrow> f \<in> measure_preserving m1 m2"
- by (auto simp add: measure_preserving_def measurable_def)
-
-lemma measure_preserving_up_sigma:
- "measure_space m1 \<Longrightarrow> (sets m1 = sets (sigma (space m1) a))
- \<Longrightarrow> measure_preserving \<lparr>space = space m1, sets = a, measure = measure m1\<rparr> m2
- \<subseteq> measure_preserving m1 m2"
- apply (rule subsetI)
- apply (rule measure_preserving_up)
- apply (simp_all add: measure_space_def sigma_def)
- apply (metis sigma_algebra.sigma_sets_eq subsetI sigma_sets.Basic)
- done
-
-lemma (in sigma_algebra) measurable_prod_sigma:
- assumes 1: "(fst o f) \<in> measurable M a1" and 2: "(snd o f) \<in> measurable M a2"
- shows "f \<in> measurable M (sigma ((space a1) \<times> (space a2))
- (prod_sets (sets a1) (sets a2)))"
-proof -
- from 1 have sa1: "sigma_algebra a1" and fn1: "fst \<circ> f \<in> space M \<rightarrow> space a1"
- and q1: "\<forall>y\<in>sets a1. (fst \<circ> f) -` y \<inter> space M \<in> sets M"
- by (auto simp add: measurable_def)
- from 2 have sa2: "sigma_algebra a2" and fn2: "snd \<circ> f \<in> space M \<rightarrow> space a2"
- and q2: "\<forall>y\<in>sets a2. (snd \<circ> f) -` y \<inter> space M \<in> sets M"
- by (auto simp add: measurable_def)
- show ?thesis
- proof (rule measurable_sigma)
- show "prod_sets (sets a1) (sets a2) \<subseteq> Pow (space a1 \<times> space a2)" using sa1 sa2
- by (auto simp add: prod_sets_def sigma_algebra_iff dest: algebra.space_closed)
- next
- show "f \<in> space M \<rightarrow> space a1 \<times> space a2"
- by (rule prod_final [OF fn1 fn2])
- next
- fix z
- assume z: "z \<in> prod_sets (sets a1) (sets a2)"
- thus "f -` z \<inter> space M \<in> sets M"
- proof (auto simp add: prod_sets_def vimage_Times)
- fix x y
- assume x: "x \<in> sets a1" and y: "y \<in> sets a2"
- have "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M =
- ((fst \<circ> f) -` x \<inter> space M) \<inter> ((snd \<circ> f) -` y \<inter> space M)"
- by blast
- also have "... \<in> sets M" using x y q1 q2
- by blast
- finally show "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M \<in> sets M" .
- qed
- qed
+ have "{x} \<inter> A = {}" using `x \<notin> A` by auto
+ from measure_additive[OF sets this] show ?thesis by simp
qed
-
-lemma (in measure_space) measurable_range_reduce:
- "f \<in> measurable M \<lparr>space = s, sets = Pow s\<rparr> \<Longrightarrow>
- s \<noteq> {}
- \<Longrightarrow> f \<in> measurable M \<lparr>space = s \<inter> (f ` space M), sets = Pow (s \<inter> (f ` space M))\<rparr>"
- by (simp add: measurable_def sigma_algebra_Pow del: Pow_Int_eq) blast
-
-lemma (in measure_space) measurable_Pow_to_Pow:
- "(sets M = Pow (space M)) \<Longrightarrow> f \<in> measurable M \<lparr>space = UNIV, sets = Pow UNIV\<rparr>"
- by (auto simp add: measurable_def sigma_algebra_def sigma_algebra_axioms_def algebra_def)
-
-lemma (in measure_space) measurable_Pow_to_Pow_image:
- "sets M = Pow (space M)
- \<Longrightarrow> f \<in> measurable M \<lparr>space = f ` space M, sets = Pow (f ` space M)\<rparr>"
- by (simp add: measurable_def sigma_algebra_Pow) intro_locales
+lemma (in measure_space) measure_finite_singleton:
+ assumes fin: "finite S"
+ and ssets: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+ shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})"
+using assms proof induct
+ case (insert x S)
+ have *: "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})" "{x} \<in> sets M"
+ using insert.prems by (blast intro!: insert.hyps(3))+
-lemma (in measure_space) measure_real_sum_image:
- assumes s: "s \<in> sets M"
- and ssets: "\<And>x. x \<in> s ==> {x} \<in> sets M"
- and fin: "finite s"
- shows "measure M s = (\<Sum>x\<in>s. measure M {x})"
-proof -
- {
- fix u
- assume u: "u \<subseteq> s & u \<in> sets M"
- hence "finite u"
- by (metis fin finite_subset)
- hence "measure M u = (\<Sum>x\<in>u. measure M {x})" using u
- proof (induct u)
- case empty
- thus ?case by simp
- next
- case (insert x t)
- hence x: "x \<in> s" and t: "t \<subseteq> s"
- by simp_all
- hence ts: "t \<in> sets M" using insert
- by (metis Diff_insert_absorb Diff ssets)
- have "measure M (insert x t) = measure M ({x} \<union> t)"
- by simp
- also have "... = measure M {x} + measure M t"
- by (simp add: measure_additive insert ssets x ts
- del: Un_insert_left)
- also have "... = (\<Sum>x\<in>insert x t. measure M {x})"
- by (simp add: x t ts insert)
- finally show ?case .
- qed
- }
- thus ?thesis
- by (metis subset_refl s)
-qed
+ have "(\<Union>x\<in>S. {x}) \<in> sets M"
+ using insert.prems `finite S` by (blast intro!: finite_UN)
+ hence "S \<in> sets M" by auto
+ from measure_insert[OF `{x} \<in> sets M` this `x \<notin> S`]
+ show ?case using `x \<notin> S` `finite S` * by simp
+qed simp
lemma (in measure_space) measure_finitely_additive':
- assumes "f \<in> ({0 :: nat ..< n} \<rightarrow> sets M)"
+ assumes "f \<in> ({..< n :: nat} \<rightarrow> sets M)"
assumes "\<And> a b. \<lbrakk>a < n ; b < n ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
- assumes "s = \<Union> (f ` {0 ..< n})"
- shows "(\<Sum> i \<in> {0 ..< n}. (measure M \<circ> f) i) = measure M s"
+ assumes "s = \<Union> (f ` {..< n})"
+ shows "(\<Sum>i<n. (\<mu> \<circ> f) i) = \<mu> s"
proof -
def f' == "\<lambda> i. (if i < n then f i else {})"
- have rf: "range f' \<subseteq> sets M" unfolding f'_def
+ have rf: "range f' \<subseteq> sets M" unfolding f'_def
using assms empty_sets by auto
- have df: "disjoint_family f'" unfolding f'_def disjoint_family_on_def
+ have df: "disjoint_family f'" unfolding f'_def disjoint_family_on_def
using assms by simp
- have "\<Union> range f' = (\<Union> i \<in> {0 ..< n}. f i)"
+ have "\<Union> range f' = (\<Union> i \<in> {..< n}. f i)"
unfolding f'_def by auto
- then have "measure M s = measure M (\<Union> range f')"
+ then have "\<mu> s = \<mu> (\<Union> range f')"
using assms by simp
- then have part1: "(\<lambda> i. measure M (f' i)) sums measure M s"
+ then have part1: "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = \<mu> s"
using df rf ca[unfolded countably_additive_def, rule_format, of f']
by auto
- have "(\<lambda> i. measure M (f' i)) sums (\<Sum> i \<in> {0 ..< n}. measure M (f' i))"
- using series_zero[of "n" "\<lambda> i. measure M (f' i)", rule_format]
- unfolding f'_def by simp
- also have "\<dots> = (\<Sum> i \<in> {0 ..< n}. measure M (f i))"
+ have "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = (\<Sum> i< n. \<mu> (f' i))"
+ by (rule psuminf_finite) (simp add: f'_def)
+ also have "\<dots> = (\<Sum>i<n. \<mu> (f i))"
unfolding f'_def by auto
- finally have part2: "(\<lambda> i. measure M (f' i)) sums (\<Sum> i \<in> {0 ..< n}. measure M (f i))" by simp
- show ?thesis
- using sums_unique[OF part1]
- sums_unique[OF part2] by auto
+ finally have part2: "(\<Sum>\<^isub>\<infinity> i. \<mu> (f' i)) = (\<Sum>i<n. \<mu> (f i))" by simp
+ show ?thesis using part1 part2 by auto
qed
@@ -895,14 +257,14 @@
assumes "finite r"
assumes "r \<subseteq> sets M"
assumes d: "\<And> a b. \<lbrakk>a \<in> r ; b \<in> r ; a \<noteq> b\<rbrakk> \<Longrightarrow> a \<inter> b = {}"
- shows "(\<Sum> i \<in> r. measure M i) = measure M (\<Union> r)"
+ shows "(\<Sum> i \<in> r. \<mu> i) = \<mu> (\<Union> r)"
using assms
proof -
(* counting the elements in r is enough *)
- let ?R = "{0 ..< card r}"
+ let ?R = "{..<card r}"
obtain f where f: "f ` ?R = r" "inj_on f ?R"
using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite r`]
- by auto
+ unfolding atLeast0LessThan by auto
hence f_into_sets: "f \<in> ?R \<rightarrow> sets M" using assms by auto
have disj: "\<And> a b. \<lbrakk>a \<in> ?R ; b \<in> ?R ; a \<noteq> b\<rbrakk> \<Longrightarrow> f a \<inter> f b = {}"
proof -
@@ -913,11 +275,11 @@
qed
have "(\<Union> r) = (\<Union> i \<in> ?R. f i)"
using f by auto
- hence "measure M (\<Union> r)= measure M (\<Union> i \<in> ?R. f i)" by simp
- also have "\<dots> = (\<Sum> i \<in> ?R. measure M (f i))"
+ hence "\<mu> (\<Union> r)= \<mu> (\<Union> i \<in> ?R. f i)" by simp
+ also have "\<dots> = (\<Sum> i \<in> ?R. \<mu> (f i))"
using measure_finitely_additive'[OF f_into_sets disj] by simp
- also have "\<dots> = (\<Sum> a \<in> r. measure M a)"
- using f[rule_format] setsum_reindex[of f ?R "\<lambda> a. measure M a"] by auto
+ also have "\<dots> = (\<Sum> a \<in> r. \<mu> a)"
+ using f[rule_format] setsum_reindex[of f ?R "\<lambda> a. \<mu> a"] by auto
finally show ?thesis by simp
qed
@@ -925,14 +287,14 @@
assumes "finite s"
assumes "\<And> i. i \<in> s \<Longrightarrow> a i \<in> sets M"
assumes d: "disjoint_family_on a s"
- shows "(\<Sum> i \<in> s. measure M (a i)) = measure M (\<Union> i \<in> s. a i)"
+ shows "(\<Sum> i \<in> s. \<mu> (a i)) = \<mu> (\<Union> i \<in> s. a i)"
using assms
proof -
(* counting the elements in s is enough *)
- let ?R = "{0 ..< card s}"
+ let ?R = "{..< card s}"
obtain f where f: "f ` ?R = s" "inj_on f ?R"
using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite s`]
- by auto
+ unfolding atLeast0LessThan by auto
hence f_into_sets: "a \<circ> f \<in> ?R \<rightarrow> sets M" using assms unfolding o_def by auto
have disj: "\<And> i j. \<lbrakk>i \<in> ?R ; j \<in> ?R ; i \<noteq> j\<rbrakk> \<Longrightarrow> (a \<circ> f) i \<inter> (a \<circ> f) j = {}"
proof -
@@ -944,105 +306,88 @@
qed
have "(\<Union> i \<in> s. a i) = (\<Union> i \<in> f ` ?R. a i)" using f by auto
hence "(\<Union> i \<in> s. a i) = (\<Union> i \<in> ?R. a (f i))" by auto
- hence "measure M (\<Union> i \<in> s. a i) = (\<Sum> i \<in> ?R. measure M (a (f i)))"
+ hence "\<mu> (\<Union> i \<in> s. a i) = (\<Sum> i \<in> ?R. \<mu> (a (f i)))"
using measure_finitely_additive'[OF f_into_sets disj] by simp
- also have "\<dots> = (\<Sum> i \<in> s. measure M (a i))"
- using f[rule_format] setsum_reindex[of f ?R "\<lambda> i. measure M (a i)"] by auto
+ also have "\<dots> = (\<Sum> i \<in> s. \<mu> (a i))"
+ using f[rule_format] setsum_reindex[of f ?R "\<lambda> i. \<mu> (a i)"] by auto
finally show ?thesis by simp
qed
-lemma (in sigma_algebra) sigma_algebra_extend:
- "sigma_algebra \<lparr>space = space M, sets = sets M, measure_space.measure = f\<rparr>"
- by unfold_locales (auto intro!: space_closed)
-
lemma (in sigma_algebra) finite_additivity_sufficient:
- assumes fin: "finite (space M)"
- and posf: "positive M f" and addf: "additive M f"
- defines "Mf \<equiv> \<lparr>space = space M, sets = sets M, measure = f\<rparr>"
- shows "measure_space Mf"
-proof -
- have [simp]: "f {} = 0" using posf
- by (simp add: positive_def)
- have "countably_additive Mf f"
- proof (auto simp add: countably_additive_def positive_def)
+ assumes fin: "finite (space M)" and pos: "positive \<mu>" and add: "additive M \<mu>"
+ shows "measure_space M \<mu>"
+proof
+ show [simp]: "\<mu> {} = 0" using pos by (simp add: positive_def)
+ show "countably_additive M \<mu>"
+ proof (auto simp add: countably_additive_def)
fix A :: "nat \<Rightarrow> 'a set"
- assume A: "range A \<subseteq> sets Mf"
+ assume A: "range A \<subseteq> sets M"
and disj: "disjoint_family A"
- and UnA: "(\<Union>i. A i) \<in> sets Mf"
+ and UnA: "(\<Union>i. A i) \<in> sets M"
def I \<equiv> "{i. A i \<noteq> {}}"
have "Union (A ` I) \<subseteq> space M" using A
- by (auto simp add: Mf_def) (metis range_subsetD subsetD sets_into_space)
+ by auto (metis range_subsetD subsetD sets_into_space)
hence "finite (A ` I)"
- by (metis finite_UnionD finite_subset fin)
+ by (metis finite_UnionD finite_subset fin)
moreover have "inj_on A I" using disj
- by (auto simp add: I_def disjoint_family_on_def inj_on_def)
+ by (auto simp add: I_def disjoint_family_on_def inj_on_def)
ultimately have finI: "finite I"
by (metis finite_imageD)
hence "\<exists>N. \<forall>m\<ge>N. A m = {}"
proof (cases "I = {}")
- case True thus ?thesis by (simp add: I_def)
+ case True thus ?thesis by (simp add: I_def)
next
case False
hence "\<forall>i\<in>I. i < Suc(Max I)"
- by (simp add: Max_less_iff [symmetric] finI)
+ by (simp add: Max_less_iff [symmetric] finI)
hence "\<forall>m \<ge> Suc(Max I). A m = {}"
- by (simp add: I_def) (metis less_le_not_le)
+ by (simp add: I_def) (metis less_le_not_le)
thus ?thesis
by blast
qed
then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast
- hence "\<forall>m\<ge>N. (f o A) m = 0"
- by simp
- hence "(\<lambda>n. f (A n)) sums setsum (f o A) {0..<N}"
- by (simp add: series_zero o_def)
- also have "... = f (\<Union>i<N. A i)"
+ then have "\<forall>m\<ge>N. \<mu> (A m) = 0" by simp
+ then have "(\<Sum>\<^isub>\<infinity> n. \<mu> (A n)) = setsum (\<lambda>m. \<mu> (A m)) {..<N}"
+ by (simp add: psuminf_finite)
+ also have "... = \<mu> (\<Union>i<N. A i)"
proof (induct N)
case 0 thus ?case by simp
next
- case (Suc n)
- have "f (A n \<union> (\<Union> x<n. A x)) = f (A n) + f (\<Union> i<n. A i)"
- proof (rule Caratheodory.additiveD [OF addf])
+ case (Suc n)
+ have "\<mu> (A n \<union> (\<Union> x<n. A x)) = \<mu> (A n) + \<mu> (\<Union> i<n. A i)"
+ proof (rule Caratheodory.additiveD [OF add])
show "A n \<inter> (\<Union> x<n. A x) = {}" using disj
by (auto simp add: disjoint_family_on_def nat_less_le) blast
- show "A n \<in> sets M" using A
- by (force simp add: Mf_def)
+ show "A n \<in> sets M" using A
+ by force
show "(\<Union>i<n. A i) \<in> sets M"
proof (induct n)
case 0 thus ?case by simp
next
case (Suc n) thus ?case using A
- by (simp add: lessThan_Suc Mf_def Un range_subsetD)
+ by (simp add: lessThan_Suc Un range_subsetD)
qed
qed
thus ?case using Suc
- by (simp add: lessThan_Suc)
+ by (simp add: lessThan_Suc)
qed
- also have "... = f (\<Union>i. A i)"
+ also have "... = \<mu> (\<Union>i. A i)"
proof -
have "(\<Union> i<N. A i) = (\<Union>i. A i)" using N
by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE)
thus ?thesis by simp
qed
- finally show "(\<lambda>n. f (A n)) sums f (\<Union>i. A i)" .
+ finally show "(\<Sum>\<^isub>\<infinity> n. \<mu> (A n)) = \<mu> (\<Union>i. A i)" .
qed
- thus ?thesis using posf
- by (simp add: Mf_def measure_space_def measure_space_axioms_def sigma_algebra_extend positive_def)
qed
-lemma finite_additivity_sufficient:
- "sigma_algebra M
- \<Longrightarrow> finite (space M)
- \<Longrightarrow> positive M (measure M) \<Longrightarrow> additive M (measure M)
- \<Longrightarrow> measure_space M"
- by (rule measure_down [OF sigma_algebra.finite_additivity_sufficient], auto)
-
lemma (in measure_space) measure_setsum_split:
assumes "finite r" and "a \<in> sets M" and br_in_M: "b ` r \<subseteq> sets M"
assumes "(\<Union>i \<in> r. b i) = space M"
assumes "disjoint_family_on b r"
- shows "(measure M) a = (\<Sum> i \<in> r. measure M (a \<inter> (b i)))"
+ shows "\<mu> a = (\<Sum> i \<in> r. \<mu> (a \<inter> (b i)))"
proof -
- have *: "measure M a = measure M (\<Union>i \<in> r. a \<inter> b i)"
+ have *: "\<mu> a = \<mu> (\<Union>i \<in> r. a \<inter> b i)"
using assms by auto
show ?thesis unfolding *
proof (rule measure_finitely_additive''[symmetric])
@@ -1057,12 +402,456 @@
qed
qed
+lemma (in measure_space) measure_subadditive:
+ assumes measurable: "A \<in> sets M" "B \<in> sets M"
+ shows "\<mu> (A \<union> B) \<le> \<mu> A + \<mu> B"
+proof -
+ from measure_additive[of A "B - A"]
+ have "\<mu> (A \<union> B) = \<mu> A + \<mu> (B - A)"
+ using assms by (simp add: Diff)
+ also have "\<dots> \<le> \<mu> A + \<mu> B"
+ using assms by (auto intro!: add_left_mono measure_mono)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) measurable_countably_subadditive:
+ assumes "range f \<subseteq> sets M"
+ shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
+proof -
+ have "\<mu> (\<Union>i. f i) = \<mu> (\<Union>i. disjointed f i)"
+ unfolding UN_disjointed_eq ..
+ also have "\<dots> = (\<Sum>\<^isub>\<infinity> i. \<mu> (disjointed f i))"
+ using range_disjointed_sets[OF assms] measure_countably_additive
+ by (simp add: disjoint_family_disjointed comp_def)
+ also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
+ proof (rule psuminf_le, rule measure_mono)
+ fix i show "disjointed f i \<subseteq> f i" by (rule disjointed_subset)
+ show "f i \<in> sets M" "disjointed f i \<in> sets M"
+ using assms range_disjointed_sets[OF assms] by auto
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) restricted_measure_space:
+ assumes "S \<in> sets M"
+ shows "measure_space (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+ (is "measure_space ?r \<mu>")
+ unfolding measure_space_def measure_space_axioms_def
+proof safe
+ show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] .
+ show "\<mu> {} = 0" by simp
+ show "countably_additive ?r \<mu>"
+ unfolding countably_additive_def
+ proof safe
+ fix A :: "nat \<Rightarrow> 'a set"
+ assume *: "range A \<subseteq> sets ?r" and **: "disjoint_family A"
+ from restriction_in_sets[OF assms *[simplified]] **
+ show "(\<Sum>\<^isub>\<infinity> n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
+ using measure_countably_additive by simp
+ qed
+qed
+
+section "@{text \<sigma>}-finite Measures"
+
+locale sigma_finite_measure = measure_space +
+ assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
+
+lemma (in sigma_finite_measure) restricted_sigma_finite_measure:
+ assumes "S \<in> sets M"
+ shows "sigma_finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+ (is "sigma_finite_measure ?r _")
+ unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def
+proof safe
+ show "measure_space ?r \<mu>" using restricted_measure_space[OF assms] .
+next
+ obtain A :: "nat \<Rightarrow> 'a set" where
+ "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ using sigma_finite by auto
+ show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
+ proof (safe intro!: exI[of _ "\<lambda>i. A i \<inter> S"] del: notI)
+ fix i
+ show "A i \<inter> S \<in> sets ?r"
+ using `range A \<subseteq> sets M` `S \<in> sets M` by auto
+ next
+ fix x i assume "x \<in> S" thus "x \<in> space ?r" by simp
+ next
+ fix x assume "x \<in> space ?r" thus "x \<in> (\<Union>i. A i \<inter> S)"
+ using `(\<Union>i. A i) = space M` `S \<in> sets M` by auto
+ next
+ fix i
+ have "\<mu> (A i \<inter> S) \<le> \<mu> (A i)"
+ using `range A \<subseteq> sets M` `S \<in> sets M` by (auto intro!: measure_mono)
+ also have "\<dots> < \<omega>" using `\<mu> (A i) \<noteq> \<omega>` by (auto simp: pinfreal_less_\<omega>)
+ finally show "\<mu> (A i \<inter> S) \<noteq> \<omega>" by (auto simp: pinfreal_less_\<omega>)
+ qed
+qed
+
+section "Real measure values"
+
+lemma (in measure_space) real_measure_Union:
+ assumes finite: "\<mu> A \<noteq> \<omega>" "\<mu> B \<noteq> \<omega>"
+ and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
+ shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
+ unfolding measure_additive[symmetric, OF measurable]
+ using finite by (auto simp: real_of_pinfreal_add)
+
+lemma (in measure_space) real_measure_finite_Union:
+ assumes measurable:
+ "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" "disjoint_family_on A S"
+ assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<omega>"
+ shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
+ using real_of_pinfreal_setsum[of S, OF finite]
+ measure_finitely_additive''[symmetric, OF measurable]
+ by simp
+
+lemma (in measure_space) real_measure_Diff:
+ assumes finite: "\<mu> A \<noteq> \<omega>"
+ and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
+ shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
+proof -
+ have "\<mu> (A - B) \<le> \<mu> A"
+ "\<mu> B \<le> \<mu> A"
+ using measurable by (auto intro!: measure_mono)
+ hence "real (\<mu> ((A - B) \<union> B)) = real (\<mu> (A - B)) + real (\<mu> B)"
+ using measurable finite by (rule_tac real_measure_Union) auto
+ thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2)
+qed
+
+lemma (in measure_space) real_measure_UNION:
+ assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
+ assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<omega>"
+ shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
+proof -
+ have *: "(\<Sum>\<^isub>\<infinity> i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
+ using measure_countably_additive[OF measurable] by (simp add: comp_def)
+ hence "(\<Sum>\<^isub>\<infinity> i. \<mu> (A i)) \<noteq> \<omega>" using finite by simp
+ from psuminf_imp_suminf[OF this]
+ show ?thesis using * by simp
+qed
+
+lemma (in measure_space) real_measure_subadditive:
+ assumes measurable: "A \<in> sets M" "B \<in> sets M"
+ and fin: "\<mu> A \<noteq> \<omega>" "\<mu> B \<noteq> \<omega>"
+ shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
+proof -
+ have "real (\<mu> (A \<union> B)) \<le> real (\<mu> A + \<mu> B)"
+ using measure_subadditive[OF measurable] fin by (auto intro!: real_of_pinfreal_mono)
+ also have "\<dots> = real (\<mu> A) + real (\<mu> B)"
+ using fin by (simp add: real_of_pinfreal_add)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) real_measurable_countably_subadditive:
+ assumes "range f \<subseteq> sets M" and "(\<Sum>\<^isub>\<infinity> i. \<mu> (f i)) \<noteq> \<omega>"
+ shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
+proof -
+ have "real (\<mu> (\<Union>i. f i)) \<le> real (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
+ using assms by (auto intro!: real_of_pinfreal_mono measurable_countably_subadditive)
+ also have "\<dots> = (\<Sum> i. real (\<mu> (f i)))"
+ using assms by (auto intro!: sums_unique psuminf_imp_suminf)
+ finally show ?thesis .
+qed
+
+lemma (in measure_space) real_measure_setsum_singleton:
+ assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+ and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
+ shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
+ using measure_finite_singleton[OF S] fin by (simp add: real_of_pinfreal_setsum)
+
+lemma (in measure_space) real_continuity_from_below:
+ assumes A: "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)" and "\<mu> (\<Union>i. A i) \<noteq> \<omega>"
+ shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
+proof (rule SUP_eq_LIMSEQ[THEN iffD1])
+ { fix n have "\<mu> (A n) \<le> \<mu> (\<Union>i. A i)"
+ using A by (auto intro!: measure_mono)
+ hence "\<mu> (A n) \<noteq> \<omega>" using assms by auto }
+ note this[simp]
+
+ show "mono (\<lambda>i. real (\<mu> (A i)))" unfolding mono_iff_le_Suc using A
+ by (auto intro!: real_of_pinfreal_mono measure_mono)
+
+ show "(SUP n. Real (real (\<mu> (A n)))) = Real (real (\<mu> (\<Union>i. A i)))"
+ using continuity_from_below[OF A(1), OF A(2)]
+ using assms by (simp add: Real_real)
+qed simp_all
+
+lemma (in measure_space) real_continuity_from_above:
+ assumes A: "range A \<subseteq> sets M"
+ and mono_Suc: "\<And>n. A (Suc n) \<subseteq> A n"
+ and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
+proof (rule INF_eq_LIMSEQ[THEN iffD1])
+ { fix n have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A n)"
+ using A by (auto intro!: measure_mono)
+ hence "\<mu> (\<Inter>i. A i) \<noteq> \<omega>" using assms by auto }
+ note this[simp]
+
+ show "mono (\<lambda>i. - real (\<mu> (A i)))" unfolding mono_iff_le_Suc using assms
+ by (auto intro!: real_of_pinfreal_mono measure_mono)
+
+ show "(INF n. Real (real (\<mu> (A n)))) =
+ Real (real (\<mu> (\<Inter>i. A i)))"
+ using continuity_from_above[OF A, OF mono_Suc finite]
+ using assms by (simp add: Real_real)
+qed simp_all
+
+locale finite_measure = measure_space +
+ assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<omega>"
+
+sublocale finite_measure < sigma_finite_measure
+proof
+ show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<omega>)"
+ using finite_measure_of_space by (auto intro!: exI[of _ "\<lambda>x. space M"])
+qed
+
+lemma (in finite_measure) finite_measure:
+ assumes "A \<in> sets M"
+ shows "\<mu> A \<noteq> \<omega>"
+proof -
+ from `A \<in> sets M` have "A \<subseteq> space M"
+ using sets_into_space by blast
+ hence "\<mu> A \<le> \<mu> (space M)"
+ using assms top by (rule measure_mono)
+ also have "\<dots> < \<omega>"
+ using finite_measure_of_space unfolding pinfreal_less_\<omega> .
+ finally show ?thesis unfolding pinfreal_less_\<omega> .
+qed
+
+lemma (in finite_measure) restricted_finite_measure:
+ assumes "S \<in> sets M"
+ shows "finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+ (is "finite_measure ?r _")
+ unfolding finite_measure_def finite_measure_axioms_def
+proof (safe del: notI)
+ show "measure_space ?r \<mu>" using restricted_measure_space[OF assms] .
+next
+ show "\<mu> (space ?r) \<noteq> \<omega>" using finite_measure[OF `S \<in> sets M`] by auto
+qed
+
+lemma (in finite_measure) real_finite_measure_Diff:
+ assumes "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
+ shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
+ using finite_measure[OF `A \<in> sets M`] by (rule real_measure_Diff[OF _ assms])
+
+lemma (in finite_measure) real_finite_measure_Union:
+ assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
+ shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
+ using sets by (auto intro!: real_measure_Union[OF _ _ assms] finite_measure)
+
+lemma (in finite_measure) real_finite_measure_finite_Union:
+ assumes "finite S" and dis: "disjoint_family_on A S"
+ and *: "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
+ shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
+proof (rule real_measure_finite_Union[OF `finite S` _ dis])
+ fix i assume "i \<in> S" from *[OF this] show "A i \<in> sets M" .
+ from finite_measure[OF this] show "\<mu> (A i) \<noteq> \<omega>" .
+qed
+
+lemma (in finite_measure) real_finite_measure_UNION:
+ assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
+ shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
+proof (rule real_measure_UNION[OF assms])
+ have "(\<Union>i. A i) \<in> sets M" using measurable(1) by (rule countable_UN)
+ thus "\<mu> (\<Union>i. A i) \<noteq> \<omega>" by (rule finite_measure)
+qed
+
+lemma (in finite_measure) real_measure_mono:
+ "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<subseteq> B \<Longrightarrow> real (\<mu> A) \<le> real (\<mu> B)"
+ by (auto intro!: measure_mono real_of_pinfreal_mono finite_measure)
+
+lemma (in finite_measure) real_finite_measure_subadditive:
+ assumes measurable: "A \<in> sets M" "B \<in> sets M"
+ shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
+ using measurable measurable[THEN finite_measure] by (rule real_measure_subadditive)
+
+lemma (in finite_measure) real_finite_measurable_countably_subadditive:
+ assumes "range f \<subseteq> sets M" and "summable (\<lambda>i. real (\<mu> (f i)))"
+ shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
+proof (rule real_measurable_countably_subadditive[OF assms(1)])
+ have *: "\<And>i. f i \<in> sets M" using assms by auto
+ have "(\<lambda>i. real (\<mu> (f i))) sums (\<Sum>i. real (\<mu> (f i)))"
+ using assms(2) by (rule summable_sums)
+ from suminf_imp_psuminf[OF this]
+ have "(\<Sum>\<^isub>\<infinity>i. \<mu> (f i)) = Real (\<Sum>i. real (\<mu> (f i)))"
+ using * by (simp add: Real_real finite_measure)
+ thus "(\<Sum>\<^isub>\<infinity>i. \<mu> (f i)) \<noteq> \<omega>" by simp
+qed
+
+lemma (in finite_measure) real_finite_measure_finite_singelton:
+ assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+ shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
+proof (rule real_measure_setsum_singleton[OF `finite S`])
+ fix x assume "x \<in> S" thus "{x} \<in> sets M" by (rule *)
+ with finite_measure show "\<mu> {x} \<noteq> \<omega>" .
+qed
+
+lemma (in finite_measure) real_finite_continuity_from_below:
+ assumes "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)"
+ shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
+ using real_continuity_from_below[OF assms(1), OF assms(2) finite_measure] assms by auto
+
+lemma (in finite_measure) real_finite_continuity_from_above:
+ assumes A: "range A \<subseteq> sets M"
+ and mono_Suc: "\<And>n. A (Suc n) \<subseteq> A n"
+ shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
+ using real_continuity_from_above[OF A, OF mono_Suc finite_measure] A
+ by auto
+
+lemma (in finite_measure) real_finite_measure_finite_Union':
+ assumes "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S"
+ shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
+ using assms real_finite_measure_finite_Union[of S A] by auto
+
+lemma (in finite_measure) finite_measure_compl:
+ assumes s: "s \<in> sets M"
+ shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
+ using measure_compl[OF s, OF finite_measure, OF s] .
+
+section {* Measure preserving *}
+
+definition "measure_preserving A \<mu> B \<nu> =
+ {f \<in> measurable A B. (\<forall>y \<in> sets B. \<mu> (f -` y \<inter> space A) = \<nu> y)}"
+
+lemma (in finite_measure) measure_preserving_lift:
+ fixes f :: "'a \<Rightarrow> 'a2" and A :: "('a2, 'b2) algebra_scheme"
+ assumes "algebra A"
+ assumes "finite_measure (sigma (space A) (sets A)) \<nu>" (is "finite_measure ?sA _")
+ assumes fmp: "f \<in> measure_preserving M \<mu> A \<nu>"
+ shows "f \<in> measure_preserving M \<mu> (sigma (space A) (sets A)) \<nu>"
+proof -
+ interpret sA: finite_measure ?sA \<nu> by fact
+ interpret A: algebra A by fact
+ show ?thesis using fmp
+ proof (clarsimp simp add: measure_preserving_def)
+ assume fm: "f \<in> measurable M A"
+ and meq: "\<forall>y\<in>sets A. \<mu> (f -` y \<inter> space M) = \<nu> y"
+ have f12: "f \<in> measurable M ?sA"
+ using measurable_subset[OF A.sets_into_space] fm by auto
+ hence ffn: "f \<in> space M \<rightarrow> space A"
+ by (simp add: measurable_def)
+ have "space M \<subseteq> f -` (space A)"
+ by auto (metis PiE ffn)
+ hence fveq [simp]: "(f -` (space A)) \<inter> space M = space M"
+ by blast
+ {
+ fix y
+ assume y: "y \<in> sets ?sA"
+ have "sets ?sA = sigma_sets (space A) (sets A)" (is "_ = ?A") by (auto simp: sigma_def)
+ also have "\<dots> \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = \<nu> s}"
+ proof (rule A.sigma_property_disjoint, auto)
+ fix x assume "x \<in> sets A" then show "\<mu> (f -` x \<inter> space M) = \<nu> x" by (simp add: meq)
+ next
+ fix s
+ assume eq: "\<mu> (f -` s \<inter> space M) = \<nu> s" and s: "s \<in> ?A"
+ then have s': "s \<in> sets ?sA" by (simp add: sigma_def)
+ show "\<mu> (f -` (space A - s) \<inter> space M) = \<nu> (space A - s)"
+ using sA.finite_measure_compl[OF s']
+ using measurable_sets[OF f12 s'] meq[THEN bspec, OF A.top]
+ by (simp add: vimage_Diff Diff_Int_distrib2 finite_measure_compl eq)
+ next
+ fix S
+ assume S0: "S 0 = {}"
+ and SSuc: "\<And>n. S n \<subseteq> S (Suc n)"
+ and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = \<nu> s}"
+ and "range S \<subseteq> ?A"
+ hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
+ have eq1: "\<And>i. \<mu> (f -` S i \<inter> space M) = \<nu> (S i)"
+ using rS1 by blast
+ have *: "(\<lambda>n. \<nu> (S n)) = (\<lambda>n. \<mu> (f -` S n \<inter> space M))"
+ by (simp add: eq1)
+ have "(SUP n. ... n) = \<mu> (\<Union>i. f -` S i \<inter> space M)"
+ proof (rule measure_countable_increasing)
+ show "range (\<lambda>i. f -` S i \<inter> space M) \<subseteq> sets M"
+ using f12 rS2 by (auto simp add: measurable_def)
+ show "f -` S 0 \<inter> space M = {}" using S0
+ by blast
+ show "\<And>n. f -` S n \<inter> space M \<subseteq> f -` S (Suc n) \<inter> space M"
+ using SSuc by auto
+ qed
+ also have "\<mu> (\<Union>i. f -` S i \<inter> space M) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)"
+ by (simp add: vimage_UN)
+ finally have "(SUP n. \<nu> (S n)) = \<mu> (f -` (\<Union>i. S i) \<inter> space M)" unfolding * .
+ moreover
+ have "(SUP n. \<nu> (S n)) = \<nu> (\<Union>i. S i)"
+ by (rule sA.measure_countable_increasing[OF rS2, OF S0 SSuc])
+ ultimately
+ show "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = \<nu> (\<Union>i. S i)" by simp
+ next
+ fix S :: "nat => 'a2 set"
+ assume dS: "disjoint_family S"
+ and rS1: "range S \<subseteq> {s. \<mu> (f -` s \<inter> space M) = \<nu> s}"
+ and "range S \<subseteq> ?A"
+ hence rS2: "range S \<subseteq> sets ?sA" by (simp add: sigma_def)
+ have "\<And>i. \<mu> (f -` S i \<inter> space M) = \<nu> (S i)"
+ using rS1 by blast
+ hence *: "(\<lambda>i. \<nu> (S i)) = (\<lambda>n. \<mu> (f -` S n \<inter> space M))"
+ by simp
+ have "psuminf ... = \<mu> (\<Union>i. f -` S i \<inter> space M)"
+ proof (rule measure_countably_additive)
+ show "range (\<lambda>i. f -` S i \<inter> space M) \<subseteq> sets M"
+ using f12 rS2 by (auto simp add: measurable_def)
+ show "disjoint_family (\<lambda>i. f -` S i \<inter> space M)" using dS
+ by (auto simp add: disjoint_family_on_def)
+ qed
+ hence "(\<Sum>\<^isub>\<infinity> i. \<nu> (S i)) = \<mu> (\<Union>i. f -` S i \<inter> space M)" unfolding * .
+ with sA.measure_countably_additive [OF rS2 dS]
+ have "\<mu> (\<Union>i. f -` S i \<inter> space M) = \<nu> (\<Union>i. S i)"
+ by simp
+ moreover have "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = \<mu> (\<Union>i. f -` S i \<inter> space M)"
+ by (simp add: vimage_UN)
+ ultimately show "\<mu> (f -` (\<Union>i. S i) \<inter> space M) = \<nu> (\<Union>i. S i)"
+ by metis
+ qed
+ finally have "sets ?sA \<subseteq> {s . \<mu> ((f -` s) \<inter> space M) = \<nu> s}" .
+ hence "\<mu> (f -` y \<inter> space M) = \<nu> y" using y by force
+ }
+ thus "f \<in> measurable M ?sA \<and> (\<forall>y\<in>sets ?sA. \<mu> (f -` y \<inter> space M) = \<nu> y)"
+ by (blast intro: f12)
+ qed
+qed
+
+section "Finite spaces"
+
locale finite_measure_space = measure_space +
assumes finite_space: "finite (space M)"
and sets_eq_Pow: "sets M = Pow (space M)"
+ and finite_single_measure: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
-lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. measure M {x}) = measure M (space M)"
+lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
using measure_finitely_additive''[of "space M" "\<lambda>i. {i}"]
by (simp add: sets_eq_Pow disjoint_family_on_def finite_space)
+sublocale finite_measure_space < finite_measure
+proof
+ show "\<mu> (space M) \<noteq> \<omega>"
+ unfolding sum_over_space[symmetric] setsum_\<omega>
+ using finite_space finite_single_measure by auto
+qed
+
+lemma psuminf_cmult_indicator:
+ assumes "disjoint_family A" "x \<in> A i"
+ shows "(\<Sum>\<^isub>\<infinity> n. f n * indicator (A n) x) = f i"
+proof -
+ have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: pinfreal)"
+ using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
+ then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: pinfreal)"
+ by (auto simp: setsum_cases)
+ moreover have "(SUP n. if i < n then f i else 0) = (f i :: pinfreal)"
+ proof (rule pinfreal_SUPI)
+ fix y :: pinfreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
+ from this[of "Suc i"] show "f i \<le> y" by auto
+ qed simp
+ ultimately show ?thesis using `x \<in> A i` unfolding psuminf_def by auto
+qed
+
+lemma psuminf_indicator:
+ assumes "disjoint_family A"
+ shows "(\<Sum>\<^isub>\<infinity> n. indicator (A n) x) = indicator (\<Union>i. A i) x"
+proof cases
+ assume *: "x \<in> (\<Union>i. A i)"
+ then obtain i where "x \<in> A i" by auto
+ from psuminf_cmult_indicator[OF assms, OF this, of "\<lambda>i. 1"]
+ show ?thesis using * by simp
+qed simp
+
end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Positive_Infinite_Real.thy Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,2723 @@
+(* Author: Johannes Hoelzl, TU Muenchen *)
+
+header {* A type for positive real numbers with infinity *}
+
+theory Positive_Infinite_Real
+ imports Complex_Main Nat_Bijection Multivariate_Analysis
+begin
+
+lemma less_Sup_iff:
+ fixes a :: "'x\<Colon>{complete_lattice,linorder}"
+ shows "a < Sup S \<longleftrightarrow> (\<exists> x \<in> S. a < x)"
+ unfolding not_le[symmetric] Sup_le_iff by auto
+
+lemma Inf_less_iff:
+ fixes a :: "'x\<Colon>{complete_lattice,linorder}"
+ shows "Inf S < a \<longleftrightarrow> (\<exists> x \<in> S. x < a)"
+ unfolding not_le[symmetric] le_Inf_iff by auto
+
+lemma SUPR_fun_expand: "(SUP y:A. f y) = (\<lambda>x. SUP y:A. f y x)"
+ unfolding SUPR_def expand_fun_eq Sup_fun_def
+ by (auto intro!: arg_cong[where f=Sup])
+
+lemma real_Suc_natfloor: "r < real (Suc (natfloor r))"
+proof cases
+ assume "0 \<le> r"
+ from floor_correct[of r]
+ have "r < real (\<lfloor>r\<rfloor> + 1)" by auto
+ also have "\<dots> = real (Suc (natfloor r))"
+ using `0 \<le> r` by (auto simp: real_of_nat_Suc natfloor_def)
+ finally show ?thesis .
+next
+ assume "\<not> 0 \<le> r"
+ hence "r < 0" by auto
+ also have "0 < real (Suc (natfloor r))" by auto
+ finally show ?thesis .
+qed
+
+lemma (in complete_lattice) Sup_mono:
+ assumes "\<And>a. a \<in> A \<Longrightarrow> \<exists>b\<in>B. a \<le> b"
+ shows "Sup A \<le> Sup B"
+proof (rule Sup_least)
+ fix a assume "a \<in> A"
+ with assms obtain b where "b \<in> B" and "a \<le> b" by auto
+ hence "b \<le> Sup B" by (auto intro: Sup_upper)
+ with `a \<le> b` show "a \<le> Sup B" by auto
+qed
+
+lemma (in complete_lattice) Inf_mono:
+ assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b"
+ shows "Inf A \<le> Inf B"
+proof (rule Inf_greatest)
+ fix b assume "b \<in> B"
+ with assms obtain a where "a \<in> A" and "a \<le> b" by auto
+ hence "Inf A \<le> a" by (auto intro: Inf_lower)
+ with `a \<le> b` show "Inf A \<le> b" by auto
+qed
+
+lemma (in complete_lattice) Sup_mono_offset:
+ fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
+ assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y" and "0 \<le> k"
+ shows "(SUP n . f (k + n)) = (SUP n. f n)"
+proof (rule antisym)
+ show "(SUP n. f (k + n)) \<le> (SUP n. f n)"
+ by (auto intro!: Sup_mono simp: SUPR_def)
+ { fix n :: 'b
+ have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
+ with * have "f n \<le> f (k + n)" by simp }
+ thus "(SUP n. f n) \<le> (SUP n. f (k + n))"
+ by (auto intro!: Sup_mono exI simp: SUPR_def)
+qed
+
+lemma (in complete_lattice) Sup_mono_offset_Suc:
+ assumes *: "\<And>x. f x \<le> f (Suc x)"
+ shows "(SUP n . f (Suc n)) = (SUP n. f n)"
+ unfolding Suc_eq_plus1
+ apply (subst add_commute)
+ apply (rule Sup_mono_offset)
+ by (auto intro!: order.lift_Suc_mono_le[of "op \<le>" "op <" f, OF _ *]) default
+
+lemma (in complete_lattice) Inf_start:
+ assumes *: "\<And>x. f 0 \<le> f x"
+ shows "(INF n. f n) = f 0"
+proof (rule antisym)
+ show "(INF n. f n) \<le> f 0" by (rule INF_leI) simp
+ show "f 0 \<le> (INF n. f n)" by (rule le_INFI[OF *])
+qed
+
+lemma (in complete_lattice) isotone_converge:
+ fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y "
+ shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
+proof -
+ have "\<And>n. (SUP m. f (n + m)) = (SUP n. f n)"
+ apply (rule Sup_mono_offset)
+ apply (rule assms)
+ by simp_all
+ moreover
+ { fix n have "(INF m. f (n + m)) = f n"
+ using Inf_start[of "\<lambda>m. f (n + m)"] assms by simp }
+ ultimately show ?thesis by simp
+qed
+
+lemma (in complete_lattice) Inf_mono_offset:
+ fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
+ assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x" and "0 \<le> k"
+ shows "(INF n . f (k + n)) = (INF n. f n)"
+proof (rule antisym)
+ show "(INF n. f n) \<le> (INF n. f (k + n))"
+ by (auto intro!: Inf_mono simp: INFI_def)
+ { fix n :: 'b
+ have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
+ with * have "f (k + n) \<le> f n" by simp }
+ thus "(INF n. f (k + n)) \<le> (INF n. f n)"
+ by (auto intro!: Inf_mono exI simp: INFI_def)
+qed
+
+lemma (in complete_lattice) Sup_start:
+ assumes *: "\<And>x. f x \<le> f 0"
+ shows "(SUP n. f n) = f 0"
+proof (rule antisym)
+ show "f 0 \<le> (SUP n. f n)" by (rule le_SUPI) auto
+ show "(SUP n. f n) \<le> f 0" by (rule SUP_leI[OF *])
+qed
+
+lemma (in complete_lattice) antitone_converges:
+ fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x"
+ shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
+proof -
+ have "\<And>n. (INF m. f (n + m)) = (INF n. f n)"
+ apply (rule Inf_mono_offset)
+ apply (rule assms)
+ by simp_all
+ moreover
+ { fix n have "(SUP m. f (n + m)) = f n"
+ using Sup_start[of "\<lambda>m. f (n + m)"] assms by simp }
+ ultimately show ?thesis by simp
+qed
+
+text {*
+
+We introduce the the positive real numbers as needed for measure theory.
+
+*}
+
+typedef pinfreal = "(Some ` {0::real..}) \<union> {None}"
+ by (rule exI[of _ None]) simp
+
+subsection "Introduce @{typ pinfreal} similar to a datatype"
+
+definition "Real x = Abs_pinfreal (Some (sup 0 x))"
+definition "\<omega> = Abs_pinfreal None"
+
+definition "pinfreal_case f i x = (if x = \<omega> then i else f (THE r. 0 \<le> r \<and> x = Real r))"
+
+definition "of_pinfreal = pinfreal_case (\<lambda>x. x) 0"
+
+defs (overloaded)
+ real_of_pinfreal_def [code_unfold]: "real == of_pinfreal"
+
+lemma pinfreal_Some[simp]: "0 \<le> x \<Longrightarrow> Some x \<in> pinfreal"
+ unfolding pinfreal_def by simp
+
+lemma pinfreal_Some_sup[simp]: "Some (sup 0 x) \<in> pinfreal"
+ by (simp add: sup_ge1)
+
+lemma pinfreal_None[simp]: "None \<in> pinfreal"
+ unfolding pinfreal_def by simp
+
+lemma Real_inj[simp]:
+ assumes "0 \<le> x" and "0 \<le> y"
+ shows "Real x = Real y \<longleftrightarrow> x = y"
+ unfolding Real_def assms[THEN sup_absorb2]
+ using assms by (simp add: Abs_pinfreal_inject)
+
+lemma Real_neq_\<omega>[simp]:
+ "Real x = \<omega> \<longleftrightarrow> False"
+ "\<omega> = Real x \<longleftrightarrow> False"
+ by (simp_all add: Abs_pinfreal_inject \<omega>_def Real_def)
+
+lemma Real_neg: "x < 0 \<Longrightarrow> Real x = Real 0"
+ unfolding Real_def by (auto simp add: Abs_pinfreal_inject intro!: sup_absorb1)
+
+lemma pinfreal_cases[case_names preal infinite, cases type: pinfreal]:
+ assumes preal: "\<And>r. x = Real r \<Longrightarrow> 0 \<le> r \<Longrightarrow> P" and inf: "x = \<omega> \<Longrightarrow> P"
+ shows P
+proof (cases x rule: pinfreal.Abs_pinfreal_cases)
+ case (Abs_pinfreal y)
+ hence "y = None \<or> (\<exists>x \<ge> 0. y = Some x)"
+ unfolding pinfreal_def by auto
+ thus P
+ proof (rule disjE)
+ assume "\<exists>x\<ge>0. y = Some x" then guess x ..
+ thus P by (simp add: preal[of x] Real_def Abs_pinfreal(1) sup_absorb2)
+ qed (simp add: \<omega>_def Abs_pinfreal(1) inf)
+qed
+
+lemma pinfreal_case_\<omega>[simp]: "pinfreal_case f i \<omega> = i"
+ unfolding pinfreal_case_def by simp
+
+lemma pinfreal_case_Real[simp]: "pinfreal_case f i (Real x) = (if 0 \<le> x then f x else f 0)"
+proof (cases "0 \<le> x")
+ case True thus ?thesis unfolding pinfreal_case_def by (auto intro: theI2)
+next
+ case False
+ moreover have "(THE r. 0 \<le> r \<and> Real 0 = Real r) = 0"
+ by (auto intro!: the_equality)
+ ultimately show ?thesis unfolding pinfreal_case_def by (simp add: Real_neg)
+qed
+
+lemma pinfreal_case_cancel[simp]: "pinfreal_case (\<lambda>c. i) i x = i"
+ by (cases x) simp_all
+
+lemma pinfreal_case_split:
+ "P (pinfreal_case f i x) = ((x = \<omega> \<longrightarrow> P i) \<and> (\<forall>r\<ge>0. x = Real r \<longrightarrow> P (f r)))"
+ by (cases x) simp_all
+
+lemma pinfreal_case_split_asm:
+ "P (pinfreal_case f i x) = (\<not> (x = \<omega> \<and> \<not> P i \<or> (\<exists>r. r \<ge> 0 \<and> x = Real r \<and> \<not> P (f r))))"
+ by (cases x) auto
+
+lemma pinfreal_case_cong[cong]:
+ assumes eq: "x = x'" "i = i'" and cong: "\<And>r. 0 \<le> r \<Longrightarrow> f r = f' r"
+ shows "pinfreal_case f i x = pinfreal_case f' i' x'"
+ unfolding eq using cong by (cases x') simp_all
+
+lemma real_Real[simp]: "real (Real x) = (if 0 \<le> x then x else 0)"
+ unfolding real_of_pinfreal_def of_pinfreal_def by simp
+
+lemma Real_real_image:
+ assumes "\<omega> \<notin> A" shows "Real ` real ` A = A"
+proof safe
+ fix x assume "x \<in> A"
+ hence *: "x = Real (real x)"
+ using `\<omega> \<notin> A` by (cases x) auto
+ show "x \<in> Real ` real ` A"
+ using `x \<in> A` by (subst *) (auto intro!: imageI)
+next
+ fix x assume "x \<in> A"
+ thus "Real (real x) \<in> A"
+ using `\<omega> \<notin> A` by (cases x) auto
+qed
+
+lemma real_pinfreal_nonneg[simp, intro]: "0 \<le> real (x :: pinfreal)"
+ unfolding real_of_pinfreal_def of_pinfreal_def
+ by (cases x) auto
+
+lemma real_\<omega>[simp]: "real \<omega> = 0"
+ unfolding real_of_pinfreal_def of_pinfreal_def by simp
+
+lemma pinfreal_noteq_omega_Ex: "X \<noteq> \<omega> \<longleftrightarrow> (\<exists>r\<ge>0. X = Real r)" by (cases X) auto
+
+subsection "@{typ pinfreal} is a monoid for addition"
+
+instantiation pinfreal :: comm_monoid_add
+begin
+
+definition "0 = Real 0"
+definition "x + y = pinfreal_case (\<lambda>r. pinfreal_case (\<lambda>p. Real (r + p)) \<omega> y) \<omega> x"
+
+lemma pinfreal_plus[simp]:
+ "Real r + Real p = (if 0 \<le> r then if 0 \<le> p then Real (r + p) else Real r else Real p)"
+ "x + 0 = x"
+ "0 + x = x"
+ "x + \<omega> = \<omega>"
+ "\<omega> + x = \<omega>"
+ by (simp_all add: plus_pinfreal_def Real_neg zero_pinfreal_def split: pinfreal_case_split)
+
+lemma \<omega>_neq_0[simp]:
+ "\<omega> = 0 \<longleftrightarrow> False"
+ "0 = \<omega> \<longleftrightarrow> False"
+ by (simp_all add: zero_pinfreal_def)
+
+lemma Real_eq_0[simp]:
+ "Real r = 0 \<longleftrightarrow> r \<le> 0"
+ "0 = Real r \<longleftrightarrow> r \<le> 0"
+ by (auto simp add: Abs_pinfreal_inject zero_pinfreal_def Real_def sup_real_def)
+
+lemma Real_0[simp]: "Real 0 = 0" by (simp add: zero_pinfreal_def)
+
+instance
+proof
+ fix a :: pinfreal
+ show "0 + a = a" by (cases a) simp_all
+
+ fix b show "a + b = b + a"
+ by (cases a, cases b) simp_all
+
+ fix c show "a + b + c = a + (b + c)"
+ by (cases a, cases b, cases c) simp_all
+qed
+end
+
+lemma pinfreal_plus_eq_\<omega>[simp]: "(a :: pinfreal) + b = \<omega> \<longleftrightarrow> a = \<omega> \<or> b = \<omega>"
+ by (cases a, cases b) auto
+
+lemma pinfreal_add_cancel_left:
+ "a + b = a + c \<longleftrightarrow> (a = \<omega> \<or> b = c)"
+ by (cases a, cases b, cases c, simp_all, cases c, simp_all)
+
+lemma pinfreal_add_cancel_right:
+ "b + a = c + a \<longleftrightarrow> (a = \<omega> \<or> b = c)"
+ by (cases a, cases b, cases c, simp_all, cases c, simp_all)
+
+lemma Real_eq_Real:
+ "Real a = Real b \<longleftrightarrow> (a = b \<or> (a \<le> 0 \<and> b \<le> 0))"
+proof (cases "a \<le> 0 \<or> b \<le> 0")
+ case False with Real_inj[of a b] show ?thesis by auto
+next
+ case True
+ thus ?thesis
+ proof
+ assume "a \<le> 0"
+ hence *: "Real a = 0" by simp
+ show ?thesis using `a \<le> 0` unfolding * by auto
+ next
+ assume "b \<le> 0"
+ hence *: "Real b = 0" by simp
+ show ?thesis using `b \<le> 0` unfolding * by auto
+ qed
+qed
+
+lemma real_pinfreal_0[simp]: "real (0 :: pinfreal) = 0"
+ unfolding zero_pinfreal_def real_Real by simp
+
+lemma real_of_pinfreal_eq_0: "real X = 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
+ by (cases X) auto
+
+lemma real_of_pinfreal_eq: "real X = real Y \<longleftrightarrow>
+ (X = Y \<or> (X = 0 \<and> Y = \<omega>) \<or> (Y = 0 \<and> X = \<omega>))"
+ by (cases X, cases Y) (auto simp add: real_of_pinfreal_eq_0)
+
+lemma real_of_pinfreal_add: "real X + real Y =
+ (if X = \<omega> then real Y else if Y = \<omega> then real X else real (X + Y))"
+ by (auto simp: pinfreal_noteq_omega_Ex)
+
+subsection "@{typ pinfreal} is a monoid for multiplication"
+
+instantiation pinfreal :: comm_monoid_mult
+begin
+
+definition "1 = Real 1"
+definition "x * y = (if x = 0 \<or> y = 0 then 0 else
+ pinfreal_case (\<lambda>r. pinfreal_case (\<lambda>p. Real (r * p)) \<omega> y) \<omega> x)"
+
+lemma pinfreal_times[simp]:
+ "Real r * Real p = (if 0 \<le> r \<and> 0 \<le> p then Real (r * p) else 0)"
+ "\<omega> * x = (if x = 0 then 0 else \<omega>)"
+ "x * \<omega> = (if x = 0 then 0 else \<omega>)"
+ "0 * x = 0"
+ "x * 0 = 0"
+ "1 = \<omega> \<longleftrightarrow> False"
+ "\<omega> = 1 \<longleftrightarrow> False"
+ by (auto simp add: times_pinfreal_def one_pinfreal_def)
+
+lemma pinfreal_one_mult[simp]:
+ "Real x + 1 = (if 0 \<le> x then Real (x + 1) else 1)"
+ "1 + Real x = (if 0 \<le> x then Real (1 + x) else 1)"
+ unfolding one_pinfreal_def by simp_all
+
+instance
+proof
+ fix a :: pinfreal show "1 * a = a"
+ by (cases a) (simp_all add: one_pinfreal_def)
+
+ fix b show "a * b = b * a"
+ by (cases a, cases b) (simp_all add: mult_nonneg_nonneg)
+
+ fix c show "a * b * c = a * (b * c)"
+ apply (cases a, cases b, cases c)
+ apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
+ apply (cases b, cases c)
+ apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
+ done
+qed
+end
+
+lemma pinfreal_mult_cancel_left:
+ "a * b = a * c \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
+ by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
+
+lemma pinfreal_mult_cancel_right:
+ "b * a = c * a \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
+ by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
+
+lemma Real_1[simp]: "Real 1 = 1" by (simp add: one_pinfreal_def)
+
+lemma real_pinfreal_1[simp]: "real (1 :: pinfreal) = 1"
+ unfolding one_pinfreal_def real_Real by simp
+
+lemma real_of_pinfreal_mult: "real X * real Y = real (X * Y :: pinfreal)"
+ by (cases X, cases Y) (auto simp: zero_le_mult_iff)
+
+subsection "@{typ pinfreal} is a linear order"
+
+instantiation pinfreal :: linorder
+begin
+
+definition "x < y \<longleftrightarrow> pinfreal_case (\<lambda>i. pinfreal_case (\<lambda>j. i < j) True y) False x"
+definition "x \<le> y \<longleftrightarrow> pinfreal_case (\<lambda>j. pinfreal_case (\<lambda>i. i \<le> j) False x) True y"
+
+lemma pinfreal_less[simp]:
+ "Real r < \<omega>"
+ "Real r < Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r < p else 0 < p)"
+ "\<omega> < x \<longleftrightarrow> False"
+ "0 < \<omega>"
+ "0 < Real r \<longleftrightarrow> 0 < r"
+ "x < 0 \<longleftrightarrow> False"
+ "0 < (1::pinfreal)"
+ by (simp_all add: less_pinfreal_def zero_pinfreal_def one_pinfreal_def del: Real_0 Real_1)
+
+lemma pinfreal_less_eq[simp]:
+ "x \<le> \<omega>"
+ "Real r \<le> Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r \<le> p else r \<le> 0)"
+ "0 \<le> x"
+ by (simp_all add: less_eq_pinfreal_def zero_pinfreal_def del: Real_0)
+
+lemma pinfreal_\<omega>_less_eq[simp]:
+ "\<omega> \<le> x \<longleftrightarrow> x = \<omega>"
+ by (cases x) (simp_all add: not_le less_eq_pinfreal_def)
+
+lemma pinfreal_less_eq_zero[simp]:
+ "(x::pinfreal) \<le> 0 \<longleftrightarrow> x = 0"
+ by (cases x) (simp_all add: zero_pinfreal_def del: Real_0)
+
+instance
+proof
+ fix x :: pinfreal
+ show "x \<le> x" by (cases x) simp_all
+ fix y
+ show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
+ by (cases x, cases y) auto
+ show "x \<le> y \<or> y \<le> x "
+ by (cases x, cases y) auto
+ { assume "x \<le> y" "y \<le> x" thus "x = y"
+ by (cases x, cases y) auto }
+ { fix z assume "x \<le> y" "y \<le> z"
+ thus "x \<le> z" by (cases x, cases y, cases z) auto }
+qed
+end
+
+lemma pinfreal_zero_lessI[intro]:
+ "(a :: pinfreal) \<noteq> 0 \<Longrightarrow> 0 < a"
+ by (cases a) auto
+
+lemma pinfreal_less_omegaI[intro, simp]:
+ "a \<noteq> \<omega> \<Longrightarrow> a < \<omega>"
+ by (cases a) auto
+
+lemma pinfreal_plus_eq_0[simp]: "(a :: pinfreal) + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
+ by (cases a, cases b) auto
+
+lemma pinfreal_le_add1[simp, intro]: "n \<le> n + (m::pinfreal)"
+ by (cases n, cases m) simp_all
+
+lemma pinfreal_le_add2: "(n::pinfreal) + m \<le> k \<Longrightarrow> m \<le> k"
+ by (cases n, cases m, cases k) simp_all
+
+lemma pinfreal_le_add3: "(n::pinfreal) + m \<le> k \<Longrightarrow> n \<le> k"
+ by (cases n, cases m, cases k) simp_all
+
+lemma pinfreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>"
+ by (cases x) auto
+
+subsection {* @{text "x - y"} on @{typ pinfreal} *}
+
+instantiation pinfreal :: minus
+begin
+definition "x - y = (if y < x then THE d. x = y + d else 0 :: pinfreal)"
+
+lemma minus_pinfreal_eq:
+ "(x - y = (z :: pinfreal)) \<longleftrightarrow> (if y < x then x = y + z else z = 0)"
+ (is "?diff \<longleftrightarrow> ?if")
+proof
+ assume ?diff
+ thus ?if
+ proof (cases "y < x")
+ case True
+ then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
+
+ show ?thesis unfolding `?diff`[symmetric] if_P[OF True] minus_pinfreal_def
+ proof (rule theI2[where Q="\<lambda>d. x = y + d"])
+ show "x = y + pinfreal_case (\<lambda>r. Real (r - real y)) \<omega> x" (is "x = y + ?d")
+ using `y < x` p by (cases x) simp_all
+
+ fix d assume "x = y + d"
+ thus "d = ?d" using `y < x` p by (cases d, cases x) simp_all
+ qed simp
+ qed (simp add: minus_pinfreal_def)
+next
+ assume ?if
+ thus ?diff
+ proof (cases "y < x")
+ case True
+ then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
+
+ from True `?if` have "x = y + z" by simp
+
+ show ?thesis unfolding minus_pinfreal_def if_P[OF True] unfolding `x = y + z`
+ proof (rule the_equality)
+ fix d :: pinfreal assume "y + z = y + d"
+ thus "d = z" using `y < x` p
+ by (cases d, cases z) simp_all
+ qed simp
+ qed (simp add: minus_pinfreal_def)
+qed
+
+instance ..
+end
+
+lemma pinfreal_minus[simp]:
+ "Real r - Real p = (if 0 \<le> r \<and> p < r then if 0 \<le> p then Real (r - p) else Real r else 0)"
+ "(A::pinfreal) - A = 0"
+ "\<omega> - Real r = \<omega>"
+ "Real r - \<omega> = 0"
+ "A - 0 = A"
+ "0 - A = 0"
+ by (auto simp: minus_pinfreal_eq not_less)
+
+lemma pinfreal_le_epsilon:
+ fixes x y :: pinfreal
+ assumes "\<And>e. 0 < e \<Longrightarrow> x \<le> y + e"
+ shows "x \<le> y"
+proof (cases y)
+ case (preal r)
+ then obtain p where x: "x = Real p" "0 \<le> p"
+ using assms[of 1] by (cases x) auto
+ { fix e have "0 < e \<Longrightarrow> p \<le> r + e"
+ using assms[of "Real e"] preal x by auto }
+ hence "p \<le> r" by (rule field_le_epsilon)
+ thus ?thesis using preal x by auto
+qed simp
+
+instance pinfreal :: "{ordered_comm_semiring, comm_semiring_1}"
+proof
+ show "0 \<noteq> (1::pinfreal)" unfolding zero_pinfreal_def one_pinfreal_def
+ by (simp del: Real_1 Real_0)
+
+ fix a :: pinfreal
+ show "0 * a = 0" "a * 0 = 0" by simp_all
+
+ fix b c :: pinfreal
+ show "(a + b) * c = a * c + b * c"
+ by (cases c, cases a, cases b)
+ (auto intro!: arg_cong[where f=Real] simp: field_simps not_le mult_le_0_iff mult_less_0_iff)
+
+ { assume "a \<le> b" thus "c + a \<le> c + b"
+ by (cases c, cases a, cases b) auto }
+
+ assume "a \<le> b" "0 \<le> c"
+ thus "c * a \<le> c * b"
+ apply (cases c, cases a, cases b)
+ by (auto simp: mult_left_mono mult_le_0_iff mult_less_0_iff not_le)
+qed
+
+lemma mult_\<omega>[simp]: "x * y = \<omega> \<longleftrightarrow> (x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0"
+ by (cases x, cases y) auto
+
+lemma \<omega>_mult[simp]: "(\<omega> = x * y) = ((x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0)"
+ by (cases x, cases y) auto
+
+lemma pinfreal_mult_0[simp]: "x * y = 0 \<longleftrightarrow> x = 0 \<or> (y::pinfreal) = 0"
+ by (cases x, cases y) (auto simp: mult_le_0_iff)
+
+lemma pinfreal_mult_cancel:
+ fixes x y z :: pinfreal
+ assumes "y \<le> z"
+ shows "x * y \<le> x * z"
+ using assms
+ by (cases x, cases y, cases z)
+ (auto simp: mult_le_cancel_left mult_le_0_iff mult_less_0_iff not_le)
+
+lemma Real_power[simp]:
+ "Real x ^ n = (if x \<le> 0 then (if n = 0 then 1 else 0) else Real (x ^ n))"
+ by (induct n) auto
+
+lemma Real_power_\<omega>[simp]:
+ "\<omega> ^ n = (if n = 0 then 1 else \<omega>)"
+ by (induct n) auto
+
+lemma pinfreal_of_nat[simp]: "of_nat m = Real (real m)"
+ by (induct m) (auto simp: real_of_nat_Suc one_pinfreal_def simp del: Real_1)
+
+lemma real_of_pinfreal_mono:
+ fixes a b :: pinfreal
+ assumes "b \<noteq> \<omega>" "a \<le> b"
+ shows "real a \<le> real b"
+using assms by (cases b, cases a) auto
+
+instance pinfreal :: "semiring_char_0"
+proof
+ fix m n
+ show "inj (of_nat::nat\<Rightarrow>pinfreal)" by (auto intro!: inj_onI)
+qed
+
+subsection "@{typ pinfreal} is a complete lattice"
+
+instantiation pinfreal :: lattice
+begin
+definition [simp]: "sup x y = (max x y :: pinfreal)"
+definition [simp]: "inf x y = (min x y :: pinfreal)"
+instance proof qed simp_all
+end
+
+instantiation pinfreal :: complete_lattice
+begin
+
+definition "bot = Real 0"
+definition "top = \<omega>"
+
+definition "Sup S = (LEAST z. \<forall>x\<in>S. x \<le> z :: pinfreal)"
+definition "Inf S = (GREATEST z. \<forall>x\<in>S. z \<le> x :: pinfreal)"
+
+lemma pinfreal_complete_Sup:
+ fixes S :: "pinfreal set" assumes "S \<noteq> {}"
+ shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
+proof (cases "\<exists>x\<ge>0. \<forall>a\<in>S. a \<le> Real x")
+ case False
+ hence *: "\<And>x. x\<ge>0 \<Longrightarrow> \<exists>a\<in>S. \<not>a \<le> Real x" by simp
+ show ?thesis
+ proof (safe intro!: exI[of _ \<omega>])
+ fix y assume **: "\<forall>z\<in>S. z \<le> y"
+ show "\<omega> \<le> y" unfolding pinfreal_\<omega>_less_eq
+ proof (rule ccontr)
+ assume "y \<noteq> \<omega>"
+ then obtain x where [simp]: "y = Real x" and "0 \<le> x" by (cases y) auto
+ from *[OF `0 \<le> x`] show False using ** by auto
+ qed
+ qed simp
+next
+ case True then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> Real y" and "0 \<le> y" by auto
+ from y[of \<omega>] have "\<omega> \<notin> S" by auto
+
+ with `S \<noteq> {}` obtain x where "x \<in> S" and "x \<noteq> \<omega>" by auto
+
+ have bound: "\<forall>x\<in>real ` S. x \<le> y"
+ proof
+ fix z assume "z \<in> real ` S" then guess a ..
+ with y[of a] `\<omega> \<notin> S` `0 \<le> y` show "z \<le> y" by (cases a) auto
+ qed
+ with reals_complete2[of "real ` S"] `x \<in> S`
+ obtain s where s: "\<forall>y\<in>S. real y \<le> s" "\<forall>z. ((\<forall>y\<in>S. real y \<le> z) \<longrightarrow> s \<le> z)"
+ by auto
+
+ show ?thesis
+ proof (safe intro!: exI[of _ "Real s"])
+ fix z assume "z \<in> S" thus "z \<le> Real s"
+ using s `\<omega> \<notin> S` by (cases z) auto
+ next
+ fix z assume *: "\<forall>y\<in>S. y \<le> z"
+ show "Real s \<le> z"
+ proof (cases z)
+ case (preal u)
+ { fix v assume "v \<in> S"
+ hence "v \<le> Real u" using * preal by auto
+ hence "real v \<le> u" using `\<omega> \<notin> S` `0 \<le> u` by (cases v) auto }
+ hence "s \<le> u" using s(2) by auto
+ thus "Real s \<le> z" using preal by simp
+ qed simp
+ qed
+qed
+
+lemma pinfreal_complete_Inf:
+ fixes S :: "pinfreal set" assumes "S \<noteq> {}"
+ shows "\<exists>x. (\<forall>y\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> x)"
+proof (cases "S = {\<omega>}")
+ case True thus ?thesis by (auto intro!: exI[of _ \<omega>])
+next
+ case False with `S \<noteq> {}` have "S - {\<omega>} \<noteq> {}" by auto
+ hence not_empty: "\<exists>x. x \<in> uminus ` real ` (S - {\<omega>})" by auto
+ have bounds: "\<exists>x. \<forall>y\<in>uminus ` real ` (S - {\<omega>}). y \<le> x" by (auto intro!: exI[of _ 0])
+ from reals_complete2[OF not_empty bounds]
+ obtain s where s: "\<And>y. y\<in>S - {\<omega>} \<Longrightarrow> - real y \<le> s" "\<forall>z. ((\<forall>y\<in>S - {\<omega>}. - real y \<le> z) \<longrightarrow> s \<le> z)"
+ by auto
+
+ show ?thesis
+ proof (safe intro!: exI[of _ "Real (-s)"])
+ fix z assume "z \<in> S"
+ show "Real (-s) \<le> z"
+ proof (cases z)
+ case (preal r)
+ with s `z \<in> S` have "z \<in> S - {\<omega>}" by simp
+ hence "- r \<le> s" using preal s(1)[of z] by auto
+ hence "- s \<le> r" by (subst neg_le_iff_le[symmetric]) simp
+ thus ?thesis using preal by simp
+ qed simp
+ next
+ fix z assume *: "\<forall>y\<in>S. z \<le> y"
+ show "z \<le> Real (-s)"
+ proof (cases z)
+ case (preal u)
+ { fix v assume "v \<in> S-{\<omega>}"
+ hence "Real u \<le> v" using * preal by auto
+ hence "- real v \<le> - u" using `0 \<le> u` `v \<in> S - {\<omega>}` by (cases v) auto }
+ hence "u \<le> - s" using s(2) by (subst neg_le_iff_le[symmetric]) auto
+ thus "z \<le> Real (-s)" using preal by simp
+ next
+ case infinite
+ with * have "S = {\<omega>}" using `S \<noteq> {}` by auto
+ with `S - {\<omega>} \<noteq> {}` show ?thesis by simp
+ qed
+ qed
+qed
+
+instance
+proof
+ fix x :: pinfreal and A
+
+ show "bot \<le> x" by (cases x) (simp_all add: bot_pinfreal_def)
+ show "x \<le> top" by (simp add: top_pinfreal_def)
+
+ { assume "x \<in> A"
+ with pinfreal_complete_Sup[of A]
+ obtain s where s: "\<forall>y\<in>A. y \<le> s" "\<forall>z. (\<forall>y\<in>A. y \<le> z) \<longrightarrow> s \<le> z" by auto
+ hence "x \<le> s" using `x \<in> A` by auto
+ also have "... = Sup A" using s unfolding Sup_pinfreal_def
+ by (auto intro!: Least_equality[symmetric])
+ finally show "x \<le> Sup A" . }
+
+ { assume "x \<in> A"
+ with pinfreal_complete_Inf[of A]
+ obtain i where i: "\<forall>y\<in>A. i \<le> y" "\<forall>z. (\<forall>y\<in>A. z \<le> y) \<longrightarrow> z \<le> i" by auto
+ hence "Inf A = i" unfolding Inf_pinfreal_def
+ by (auto intro!: Greatest_equality)
+ also have "i \<le> x" using i `x \<in> A` by auto
+ finally show "Inf A \<le> x" . }
+
+ { assume *: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x"
+ show "Sup A \<le> x"
+ proof (cases "A = {}")
+ case True
+ hence "Sup A = 0" unfolding Sup_pinfreal_def
+ by (auto intro!: Least_equality)
+ thus "Sup A \<le> x" by simp
+ next
+ case False
+ with pinfreal_complete_Sup[of A]
+ obtain s where s: "\<forall>y\<in>A. y \<le> s" "\<forall>z. (\<forall>y\<in>A. y \<le> z) \<longrightarrow> s \<le> z" by auto
+ hence "Sup A = s"
+ unfolding Sup_pinfreal_def by (auto intro!: Least_equality)
+ also have "s \<le> x" using * s by auto
+ finally show "Sup A \<le> x" .
+ qed }
+
+ { assume *: "\<And>z. z \<in> A \<Longrightarrow> x \<le> z"
+ show "x \<le> Inf A"
+ proof (cases "A = {}")
+ case True
+ hence "Inf A = \<omega>" unfolding Inf_pinfreal_def
+ by (auto intro!: Greatest_equality)
+ thus "x \<le> Inf A" by simp
+ next
+ case False
+ with pinfreal_complete_Inf[of A]
+ obtain i where i: "\<forall>y\<in>A. i \<le> y" "\<forall>z. (\<forall>y\<in>A. z \<le> y) \<longrightarrow> z \<le> i" by auto
+ have "x \<le> i" using * i by auto
+ also have "i = Inf A" using i
+ unfolding Inf_pinfreal_def by (auto intro!: Greatest_equality[symmetric])
+ finally show "x \<le> Inf A" .
+ qed }
+qed
+end
+
+lemma Inf_pinfreal_iff:
+ fixes z :: pinfreal
+ shows "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> (\<exists>x\<in>X. x<y) \<longleftrightarrow> Inf X < y"
+ by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
+ order_less_le_trans)
+
+lemma Inf_greater:
+ fixes z :: pinfreal assumes "Inf X < z"
+ shows "\<exists>x \<in> X. x < z"
+proof -
+ have "X \<noteq> {}" using assms by (auto simp: Inf_empty top_pinfreal_def)
+ with assms show ?thesis
+ by (metis Inf_pinfreal_iff mem_def not_leE)
+qed
+
+lemma Inf_close:
+ fixes e :: pinfreal assumes "Inf X \<noteq> \<omega>" "0 < e"
+ shows "\<exists>x \<in> X. x < Inf X + e"
+proof (rule Inf_greater)
+ show "Inf X < Inf X + e" using assms
+ by (cases "Inf X", cases e) auto
+qed
+
+lemma pinfreal_SUPI:
+ fixes x :: pinfreal
+ assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<le> x"
+ assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> x \<le> y"
+ shows "(SUP i:A. f i) = x"
+ unfolding SUPR_def Sup_pinfreal_def
+ using assms by (auto intro!: Least_equality)
+
+lemma Sup_pinfreal_iff:
+ fixes z :: pinfreal
+ shows "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> (\<exists>x\<in>X. y<x) \<longleftrightarrow> y < Sup X"
+ by (metis complete_lattice_class.Sup_least complete_lattice_class.Sup_upper less_le_not_le linear
+ order_less_le_trans)
+
+lemma Sup_lesser:
+ fixes z :: pinfreal assumes "z < Sup X"
+ shows "\<exists>x \<in> X. z < x"
+proof -
+ have "X \<noteq> {}" using assms by (auto simp: Sup_empty bot_pinfreal_def)
+ with assms show ?thesis
+ by (metis Sup_pinfreal_iff mem_def not_leE)
+qed
+
+lemma Sup_eq_\<omega>: "\<omega> \<in> S \<Longrightarrow> Sup S = \<omega>"
+ unfolding Sup_pinfreal_def
+ by (auto intro!: Least_equality)
+
+lemma Sup_close:
+ assumes "0 < e" and S: "Sup S \<noteq> \<omega>" "S \<noteq> {}"
+ shows "\<exists>X\<in>S. Sup S < X + e"
+proof cases
+ assume "Sup S = 0"
+ moreover obtain X where "X \<in> S" using `S \<noteq> {}` by auto
+ ultimately show ?thesis using `0 < e` by (auto intro!: bexI[OF _ `X\<in>S`])
+next
+ assume "Sup S \<noteq> 0"
+ have "\<exists>X\<in>S. Sup S - e < X"
+ proof (rule Sup_lesser)
+ show "Sup S - e < Sup S" using `0 < e` `Sup S \<noteq> 0` `Sup S \<noteq> \<omega>`
+ by (cases e) (auto simp: pinfreal_noteq_omega_Ex)
+ qed
+ then guess X .. note X = this
+ with `Sup S \<noteq> \<omega>` Sup_eq_\<omega> have "X \<noteq> \<omega>" by auto
+ thus ?thesis using `Sup S \<noteq> \<omega>` X unfolding pinfreal_noteq_omega_Ex
+ by (cases e) (auto intro!: bexI[OF _ `X\<in>S`] simp: split: split_if_asm)
+qed
+
+lemma Sup_\<omega>: "(SUP i::nat. Real (real i)) = \<omega>"
+proof (rule pinfreal_SUPI)
+ fix y assume *: "\<And>i::nat. i \<in> UNIV \<Longrightarrow> Real (real i) \<le> y"
+ thus "\<omega> \<le> y"
+ proof (cases y)
+ case (preal r)
+ then obtain k :: nat where "r < real k"
+ using ex_less_of_nat by (auto simp: real_eq_of_nat)
+ with *[of k] preal show ?thesis by auto
+ qed simp
+qed simp
+
+subsubsection {* Equivalence between @{text "f ----> x"} and @{text SUP} on @{typ pinfreal} *}
+
+lemma monoseq_monoI: "mono f \<Longrightarrow> monoseq f"
+ unfolding mono_def monoseq_def by auto
+
+lemma incseq_mono: "mono f \<longleftrightarrow> incseq f"
+ unfolding mono_def incseq_def by auto
+
+lemma SUP_eq_LIMSEQ:
+ assumes "mono f" and "\<And>n. 0 \<le> f n" and "0 \<le> x"
+ shows "(SUP n. Real (f n)) = Real x \<longleftrightarrow> f ----> x"
+proof
+ assume x: "(SUP n. Real (f n)) = Real x"
+ { fix n
+ have "Real (f n) \<le> Real x" using x[symmetric] by (auto intro: le_SUPI)
+ hence "f n \<le> x" using assms by simp }
+ show "f ----> x"
+ proof (rule LIMSEQ_I)
+ fix r :: real assume "0 < r"
+ show "\<exists>no. \<forall>n\<ge>no. norm (f n - x) < r"
+ proof (rule ccontr)
+ assume *: "\<not> ?thesis"
+ { fix N
+ from * obtain n where "N \<le> n" "r \<le> x - f n"
+ using `\<And>n. f n \<le> x` by (auto simp: not_less)
+ hence "f N \<le> f n" using `mono f` by (auto dest: monoD)
+ hence "f N \<le> x - r" using `r \<le> x - f n` by auto
+ hence "Real (f N) \<le> Real (x - r)" and "r \<le> x" using `0 \<le> f N` by auto }
+ hence "(SUP n. Real (f n)) \<le> Real (x - r)"
+ and "Real (x - r) < Real x" using `0 < r` by (auto intro: SUP_leI)
+ hence "(SUP n. Real (f n)) < Real x" by (rule le_less_trans)
+ thus False using x by auto
+ qed
+ qed
+next
+ assume "f ----> x"
+ show "(SUP n. Real (f n)) = Real x"
+ proof (rule pinfreal_SUPI)
+ fix n
+ from incseq_le[of f x] `mono f` `f ----> x`
+ show "Real (f n) \<le> Real x" using assms incseq_mono by auto
+ next
+ fix y assume *: "\<And>n. n\<in>UNIV \<Longrightarrow> Real (f n) \<le> y"
+ show "Real x \<le> y"
+ proof (cases y)
+ case (preal r)
+ with * have "\<exists>N. \<forall>n\<ge>N. f n \<le> r" using assms by fastsimp
+ from LIMSEQ_le_const2[OF `f ----> x` this]
+ show "Real x \<le> y" using `0 \<le> x` preal by auto
+ qed simp
+ qed
+qed
+
+lemma SUPR_bound:
+ assumes "\<forall>N. f N \<le> x"
+ shows "(SUP n. f n) \<le> x"
+ using assms by (simp add: SUPR_def Sup_le_iff)
+
+lemma pinfreal_less_eq_diff_eq_sum:
+ fixes x y z :: pinfreal
+ assumes "y \<le> x" and "x \<noteq> \<omega>"
+ shows "z \<le> x - y \<longleftrightarrow> z + y \<le> x"
+ using assms
+ apply (cases z, cases y, cases x)
+ by (simp_all add: field_simps minus_pinfreal_eq)
+
+lemma Real_diff_less_omega: "Real r - x < \<omega>" by (cases x) auto
+
+subsubsection {* Numbers on @{typ pinfreal} *}
+
+instantiation pinfreal :: number
+begin
+definition [simp]: "number_of x = Real (number_of x)"
+instance proof qed
+end
+
+subsubsection {* Division on @{typ pinfreal} *}
+
+instantiation pinfreal :: inverse
+begin
+
+definition "inverse x = pinfreal_case (\<lambda>x. if x = 0 then \<omega> else Real (inverse x)) 0 x"
+definition [simp]: "x / y = x * inverse (y :: pinfreal)"
+
+instance proof qed
+end
+
+lemma pinfreal_inverse[simp]:
+ "inverse 0 = \<omega>"
+ "inverse (Real x) = (if x \<le> 0 then \<omega> else Real (inverse x))"
+ "inverse \<omega> = 0"
+ "inverse (1::pinfreal) = 1"
+ "inverse (inverse x) = x"
+ by (simp_all add: inverse_pinfreal_def one_pinfreal_def split: pinfreal_case_split del: Real_1)
+
+lemma pinfreal_inverse_le_eq:
+ assumes "x \<noteq> 0" "x \<noteq> \<omega>"
+ shows "y \<le> z / x \<longleftrightarrow> x * y \<le> (z :: pinfreal)"
+proof -
+ from assms obtain r where r: "x = Real r" "0 < r" by (cases x) auto
+ { fix p q :: real assume "0 \<le> p" "0 \<le> q"
+ have "p \<le> q * inverse r \<longleftrightarrow> p \<le> q / r" by (simp add: divide_inverse)
+ also have "... \<longleftrightarrow> p * r \<le> q" using `0 < r` by (auto simp: field_simps)
+ finally have "p \<le> q * inverse r \<longleftrightarrow> p * r \<le> q" . }
+ with r show ?thesis
+ by (cases y, cases z, auto simp: zero_le_mult_iff field_simps)
+qed
+
+lemma inverse_antimono_strict:
+ fixes x y :: pinfreal
+ assumes "x < y" shows "inverse y < inverse x"
+ using assms by (cases x, cases y) auto
+
+lemma inverse_antimono:
+ fixes x y :: pinfreal
+ assumes "x \<le> y" shows "inverse y \<le> inverse x"
+ using assms by (cases x, cases y) auto
+
+lemma pinfreal_inverse_\<omega>_iff[simp]: "inverse x = \<omega> \<longleftrightarrow> x = 0"
+ by (cases x) auto
+
+subsection "Infinite sum over @{typ pinfreal}"
+
+text {*
+
+The infinite sum over @{typ pinfreal} has the nice property that it is always
+defined.
+
+*}
+
+definition psuminf :: "(nat \<Rightarrow> pinfreal) \<Rightarrow> pinfreal" (binder "\<Sum>\<^isub>\<infinity>" 10) where
+ "(\<Sum>\<^isub>\<infinity> x. f x) = (SUP n. \<Sum>i<n. f i)"
+
+subsubsection {* Equivalence between @{text "\<Sum> n. f n"} and @{text "\<Sum>\<^isub>\<infinity> n. f n"} *}
+
+lemma setsum_Real:
+ assumes "\<forall>x\<in>A. 0 \<le> f x"
+ shows "(\<Sum>x\<in>A. Real (f x)) = Real (\<Sum>x\<in>A. f x)"
+proof (cases "finite A")
+ case True
+ thus ?thesis using assms
+ proof induct case (insert x s)
+ hence "0 \<le> setsum f s" apply-apply(rule setsum_nonneg) by auto
+ thus ?case using insert by auto
+ qed auto
+qed simp
+
+lemma setsum_Real':
+ assumes "\<forall>x. 0 \<le> f x"
+ shows "(\<Sum>x\<in>A. Real (f x)) = Real (\<Sum>x\<in>A. f x)"
+ apply(rule setsum_Real) using assms by auto
+
+lemma setsum_\<omega>:
+ "(\<Sum>x\<in>P. f x) = \<omega> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<omega>))"
+proof safe
+ assume *: "setsum f P = \<omega>"
+ show "finite P"
+ proof (rule ccontr) assume "infinite P" with * show False by auto qed
+ show "\<exists>i\<in>P. f i = \<omega>"
+ proof (rule ccontr)
+ assume "\<not> ?thesis" hence "\<And>i. i\<in>P \<Longrightarrow> f i \<noteq> \<omega>" by auto
+ from `finite P` this have "setsum f P \<noteq> \<omega>"
+ by induct auto
+ with * show False by auto
+ qed
+next
+ fix i assume "finite P" "i \<in> P" "f i = \<omega>"
+ thus "setsum f P = \<omega>"
+ proof induct
+ case (insert x A)
+ show ?case using insert by (cases "x = i") auto
+ qed simp
+qed
+
+lemma real_of_pinfreal_setsum:
+ assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<noteq> \<omega>"
+ shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
+proof cases
+ assume "finite S"
+ from this assms show ?thesis
+ by induct (simp_all add: real_of_pinfreal_add setsum_\<omega>)
+qed simp
+
+lemma setsum_0:
+ fixes f :: "'a \<Rightarrow> pinfreal" assumes "finite A"
+ shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)"
+ using assms by induct auto
+
+lemma suminf_imp_psuminf:
+ assumes "f sums x" and "\<forall>n. 0 \<le> f n"
+ shows "(\<Sum>\<^isub>\<infinity> x. Real (f x)) = Real x"
+ unfolding psuminf_def setsum_Real'[OF assms(2)]
+proof (rule SUP_eq_LIMSEQ[THEN iffD2])
+ show "mono (\<lambda>n. setsum f {..<n})" (is "mono ?S")
+ unfolding mono_iff_le_Suc using assms by simp
+
+ { fix n show "0 \<le> ?S n"
+ using setsum_nonneg[of "{..<n}" f] assms by auto }
+
+ thus "0 \<le> x" "?S ----> x"
+ using `f sums x` LIMSEQ_le_const
+ by (auto simp: atLeast0LessThan sums_def)
+qed
+
+lemma psuminf_equality:
+ assumes "\<And>n. setsum f {..<n} \<le> x"
+ and "\<And>y. y \<noteq> \<omega> \<Longrightarrow> (\<And>n. setsum f {..<n} \<le> y) \<Longrightarrow> x \<le> y"
+ shows "psuminf f = x"
+ unfolding psuminf_def
+proof (safe intro!: pinfreal_SUPI)
+ fix n show "setsum f {..<n} \<le> x" using assms(1) .
+next
+ fix y assume *: "\<forall>n. n \<in> UNIV \<longrightarrow> setsum f {..<n} \<le> y"
+ show "x \<le> y"
+ proof (cases "y = \<omega>")
+ assume "y \<noteq> \<omega>" from assms(2)[OF this] *
+ show "x \<le> y" by auto
+ qed simp
+qed
+
+lemma psuminf_\<omega>:
+ assumes "f i = \<omega>"
+ shows "(\<Sum>\<^isub>\<infinity> x. f x) = \<omega>"
+proof (rule psuminf_equality)
+ fix y assume *: "\<And>n. setsum f {..<n} \<le> y"
+ have "setsum f {..<Suc i} = \<omega>"
+ using assms by (simp add: setsum_\<omega>)
+ thus "\<omega> \<le> y" using *[of "Suc i"] by auto
+qed simp
+
+lemma psuminf_imp_suminf:
+ assumes "(\<Sum>\<^isub>\<infinity> x. f x) \<noteq> \<omega>"
+ shows "(\<lambda>x. real (f x)) sums real (\<Sum>\<^isub>\<infinity> x. f x)"
+proof -
+ have "\<forall>i. \<exists>r. f i = Real r \<and> 0 \<le> r"
+ proof
+ fix i show "\<exists>r. f i = Real r \<and> 0 \<le> r" using psuminf_\<omega> assms by (cases "f i") auto
+ qed
+ from choice[OF this] obtain r where f: "f = (\<lambda>i. Real (r i))"
+ and pos: "\<forall>i. 0 \<le> r i"
+ by (auto simp: expand_fun_eq)
+ hence [simp]: "\<And>i. real (f i) = r i" by auto
+
+ have "mono (\<lambda>n. setsum r {..<n})" (is "mono ?S")
+ unfolding mono_iff_le_Suc using pos by simp
+
+ { fix n have "0 \<le> ?S n"
+ using setsum_nonneg[of "{..<n}" r] pos by auto }
+
+ from assms obtain p where *: "(\<Sum>\<^isub>\<infinity> x. f x) = Real p" and "0 \<le> p"
+ by (cases "(\<Sum>\<^isub>\<infinity> x. f x)") auto
+ show ?thesis unfolding * using * pos `0 \<le> p` SUP_eq_LIMSEQ[OF `mono ?S` `\<And>n. 0 \<le> ?S n` `0 \<le> p`]
+ by (simp add: f atLeast0LessThan sums_def psuminf_def setsum_Real'[OF pos] f)
+qed
+
+lemma psuminf_bound:
+ assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x"
+ shows "(\<Sum>\<^isub>\<infinity> n. f n) \<le> x"
+ using assms by (simp add: psuminf_def SUPR_def Sup_le_iff)
+
+lemma psuminf_bound_add:
+ assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x"
+ shows "(\<Sum>\<^isub>\<infinity> n. f n) + y \<le> x"
+proof (cases "x = \<omega>")
+ have "y \<le> x" using assms by (auto intro: pinfreal_le_add2)
+ assume "x \<noteq> \<omega>"
+ note move_y = pinfreal_less_eq_diff_eq_sum[OF `y \<le> x` this]
+
+ have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y" using assms by (simp add: move_y)
+ hence "(\<Sum>\<^isub>\<infinity> n. f n) \<le> x - y" by (rule psuminf_bound)
+ thus ?thesis by (simp add: move_y)
+qed simp
+
+lemma psuminf_finite:
+ assumes "\<forall>N\<ge>n. f N = 0"
+ shows "(\<Sum>\<^isub>\<infinity> n. f n) = (\<Sum>N<n. f N)"
+proof (rule psuminf_equality)
+ fix N
+ show "setsum f {..<N} \<le> setsum f {..<n}"
+ proof (cases rule: linorder_cases)
+ assume "N < n" thus ?thesis by (auto intro!: setsum_mono3)
+ next
+ assume "n < N"
+ hence *: "{..<N} = {..<n} \<union> {n..<N}" by auto
+ moreover have "setsum f {n..<N} = 0"
+ using assms by (auto intro!: setsum_0')
+ ultimately show ?thesis unfolding *
+ by (subst setsum_Un_disjoint) auto
+ qed simp
+qed simp
+
+lemma psuminf_upper:
+ shows "(\<Sum>n<N. f n) \<le> (\<Sum>\<^isub>\<infinity> n. f n)"
+ unfolding psuminf_def SUPR_def
+ by (auto intro: complete_lattice_class.Sup_upper image_eqI)
+
+lemma psuminf_le:
+ assumes "\<And>N. f N \<le> g N"
+ shows "psuminf f \<le> psuminf g"
+proof (safe intro!: psuminf_bound)
+ fix n
+ have "setsum f {..<n} \<le> setsum g {..<n}" using assms by (auto intro: setsum_mono)
+ also have "... \<le> psuminf g" by (rule psuminf_upper)
+ finally show "setsum f {..<n} \<le> psuminf g" .
+qed
+
+lemma psuminf_const[simp]: "psuminf (\<lambda>n. c) = (if c = 0 then 0 else \<omega>)" (is "_ = ?if")
+proof (rule psuminf_equality)
+ fix y assume *: "\<And>n :: nat. (\<Sum>n<n. c) \<le> y" and "y \<noteq> \<omega>"
+ then obtain r p where
+ y: "y = Real r" "0 \<le> r" and
+ c: "c = Real p" "0 \<le> p"
+ using *[of 1] by (cases c, cases y) auto
+ show "(if c = 0 then 0 else \<omega>) \<le> y"
+ proof (cases "p = 0")
+ assume "p = 0" with c show ?thesis by simp
+ next
+ assume "p \<noteq> 0"
+ with * c y have **: "\<And>n :: nat. real n \<le> r / p"
+ by (auto simp: zero_le_mult_iff field_simps)
+ from ex_less_of_nat[of "r / p"] guess n ..
+ with **[of n] show ?thesis by (simp add: real_eq_of_nat)
+ qed
+qed (cases "c = 0", simp_all)
+
+lemma psuminf_add[simp]: "psuminf (\<lambda>n. f n + g n) = psuminf f + psuminf g"
+proof (rule psuminf_equality)
+ fix n
+ from psuminf_upper[of f n] psuminf_upper[of g n]
+ show "(\<Sum>n<n. f n + g n) \<le> psuminf f + psuminf g"
+ by (auto simp add: setsum_addf intro!: add_mono)
+next
+ fix y assume *: "\<And>n. (\<Sum>n<n. f n + g n) \<le> y" and "y \<noteq> \<omega>"
+ { fix n m
+ have **: "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> y"
+ proof (cases rule: linorder_le_cases)
+ assume "n \<le> m"
+ hence "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> (\<Sum>n<m. f n) + (\<Sum>n<m. g n)"
+ by (auto intro!: add_right_mono setsum_mono3)
+ also have "... \<le> y"
+ using * by (simp add: setsum_addf)
+ finally show ?thesis .
+ next
+ assume "m \<le> n"
+ hence "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> (\<Sum>n<n. f n) + (\<Sum>n<n. g n)"
+ by (auto intro!: add_left_mono setsum_mono3)
+ also have "... \<le> y"
+ using * by (simp add: setsum_addf)
+ finally show ?thesis .
+ qed }
+ hence "\<And>m. \<forall>n. (\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> y" by simp
+ from psuminf_bound_add[OF this]
+ have "\<forall>m. (\<Sum>n<m. g n) + psuminf f \<le> y" by (simp add: ac_simps)
+ from psuminf_bound_add[OF this]
+ show "psuminf f + psuminf g \<le> y" by (simp add: ac_simps)
+qed
+
+lemma psuminf_0: "psuminf f = 0 \<longleftrightarrow> (\<forall>i. f i = 0)"
+proof safe
+ assume "\<forall>i. f i = 0"
+ hence "f = (\<lambda>i. 0)" by (simp add: expand_fun_eq)
+ thus "psuminf f = 0" using psuminf_const by simp
+next
+ fix i assume "psuminf f = 0"
+ hence "(\<Sum>n<Suc i. f n) = 0" using psuminf_upper[of f "Suc i"] by simp
+ thus "f i = 0" by simp
+qed
+
+lemma psuminf_cmult_right[simp]: "psuminf (\<lambda>n. c * f n) = c * psuminf f"
+proof (rule psuminf_equality)
+ fix n show "(\<Sum>n<n. c * f n) \<le> c * psuminf f"
+ by (auto simp: setsum_right_distrib[symmetric] intro: mult_left_mono psuminf_upper)
+next
+ fix y
+ assume "\<And>n. (\<Sum>n<n. c * f n) \<le> y"
+ hence *: "\<And>n. c * (\<Sum>n<n. f n) \<le> y" by (auto simp add: setsum_right_distrib)
+ thus "c * psuminf f \<le> y"
+ proof (cases "c = \<omega> \<or> c = 0")
+ assume "c = \<omega> \<or> c = 0"
+ thus ?thesis
+ using * by (fastsimp simp add: psuminf_0 setsum_0 split: split_if_asm)
+ next
+ assume "\<not> (c = \<omega> \<or> c = 0)"
+ hence "c \<noteq> 0" "c \<noteq> \<omega>" by auto
+ note rewrite_div = pinfreal_inverse_le_eq[OF this, of _ y]
+ hence "\<forall>n. (\<Sum>n<n. f n) \<le> y / c" using * by simp
+ hence "psuminf f \<le> y / c" by (rule psuminf_bound)
+ thus ?thesis using rewrite_div by simp
+ qed
+qed
+
+lemma psuminf_cmult_left[simp]: "psuminf (\<lambda>n. f n * c) = psuminf f * c"
+ using psuminf_cmult_right[of c f] by (simp add: ac_simps)
+
+lemma psuminf_half_series: "psuminf (\<lambda>n. (1/2)^Suc n) = 1"
+ using suminf_imp_psuminf[OF power_half_series] by auto
+
+lemma setsum_pinfsum: "(\<Sum>\<^isub>\<infinity> n. \<Sum>m\<in>A. f n m) = (\<Sum>m\<in>A. (\<Sum>\<^isub>\<infinity> n. f n m))"
+proof (cases "finite A")
+ assume "finite A"
+ thus ?thesis by induct simp_all
+qed simp
+
+lemma psuminf_reindex:
+ fixes f:: "nat \<Rightarrow> nat" assumes "bij f"
+ shows "psuminf (g \<circ> f) = psuminf g"
+proof -
+ have [intro, simp]: "\<And>A. inj_on f A" using `bij f` unfolding bij_def by (auto intro: subset_inj_on)
+ have f[intro, simp]: "\<And>x. f (inv f x) = x"
+ using `bij f` unfolding bij_def by (auto intro: surj_f_inv_f)
+
+ show ?thesis
+ proof (rule psuminf_equality)
+ fix n
+ have "setsum (g \<circ> f) {..<n} = setsum g (f ` {..<n})"
+ by (simp add: setsum_reindex)
+ also have "\<dots> \<le> setsum g {..Max (f ` {..<n})}"
+ by (rule setsum_mono3) auto
+ also have "\<dots> \<le> psuminf g" unfolding lessThan_Suc_atMost[symmetric] by (rule psuminf_upper)
+ finally show "setsum (g \<circ> f) {..<n} \<le> psuminf g" .
+ next
+ fix y assume *: "\<And>n. setsum (g \<circ> f) {..<n} \<le> y"
+ show "psuminf g \<le> y"
+ proof (safe intro!: psuminf_bound)
+ fix N
+ have "setsum g {..<N} \<le> setsum g (f ` {..Max (inv f ` {..<N})})"
+ by (rule setsum_mono3) (auto intro!: image_eqI[where f="f", OF f[symmetric]])
+ also have "\<dots> = setsum (g \<circ> f) {..Max (inv f ` {..<N})}"
+ by (simp add: setsum_reindex)
+ also have "\<dots> \<le> y" unfolding lessThan_Suc_atMost[symmetric] by (rule *)
+ finally show "setsum g {..<N} \<le> y" .
+ qed
+ qed
+qed
+
+lemma psuminf_2dimen:
+ fixes f:: "nat * nat \<Rightarrow> pinfreal"
+ assumes fsums: "\<And>m. g m = (\<Sum>\<^isub>\<infinity> n. f (m,n))"
+ shows "psuminf (f \<circ> prod_decode) = psuminf g"
+proof (rule psuminf_equality)
+ fix n :: nat
+ let ?P = "prod_decode ` {..<n}"
+ have "setsum (f \<circ> prod_decode) {..<n} = setsum f ?P"
+ by (auto simp: setsum_reindex inj_prod_decode)
+ also have "\<dots> \<le> setsum f ({..Max (fst ` ?P)} \<times> {..Max (snd ` ?P)})"
+ proof (safe intro!: setsum_mono3 Max_ge image_eqI)
+ fix a b x assume "(a, b) = prod_decode x"
+ from this[symmetric] show "a = fst (prod_decode x)" "b = snd (prod_decode x)"
+ by simp_all
+ qed simp_all
+ also have "\<dots> = (\<Sum>m\<le>Max (fst ` ?P). (\<Sum>n\<le>Max (snd ` ?P). f (m,n)))"
+ unfolding setsum_cartesian_product by simp
+ also have "\<dots> \<le> (\<Sum>m\<le>Max (fst ` ?P). g m)"
+ by (auto intro!: setsum_mono psuminf_upper simp del: setsum_lessThan_Suc
+ simp: fsums lessThan_Suc_atMost[symmetric])
+ also have "\<dots> \<le> psuminf g"
+ by (auto intro!: psuminf_upper simp del: setsum_lessThan_Suc
+ simp: lessThan_Suc_atMost[symmetric])
+ finally show "setsum (f \<circ> prod_decode) {..<n} \<le> psuminf g" .
+next
+ fix y assume *: "\<And>n. setsum (f \<circ> prod_decode) {..<n} \<le> y"
+ have g: "g = (\<lambda>m. \<Sum>\<^isub>\<infinity> n. f (m,n))" unfolding fsums[symmetric] ..
+ show "psuminf g \<le> y" unfolding g
+ proof (rule psuminf_bound, unfold setsum_pinfsum[symmetric], safe intro!: psuminf_bound)
+ fix N M :: nat
+ let ?P = "{..<N} \<times> {..<M}"
+ let ?M = "Max (prod_encode ` ?P)"
+ have "(\<Sum>n<M. \<Sum>m<N. f (m, n)) \<le> (\<Sum>(m, n)\<in>?P. f (m, n))"
+ unfolding setsum_commute[of _ _ "{..<M}"] unfolding setsum_cartesian_product ..
+ also have "\<dots> \<le> (\<Sum>(m,n)\<in>(prod_decode ` {..?M}). f (m, n))"
+ by (auto intro!: setsum_mono3 image_eqI[where f=prod_decode, OF prod_encode_inverse[symmetric]])
+ also have "\<dots> \<le> y" using *[of "Suc ?M"]
+ by (simp add: lessThan_Suc_atMost[symmetric] setsum_reindex
+ inj_prod_decode del: setsum_lessThan_Suc)
+ finally show "(\<Sum>n<M. \<Sum>m<N. f (m, n)) \<le> y" .
+ qed
+qed
+
+lemma pinfreal_mult_less_right:
+ assumes "b * a < c * a" "0 < a" "a < \<omega>"
+ shows "b < c"
+ using assms
+ by (cases a, cases b, cases c) (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
+
+lemma pinfreal_\<omega>_eq_plus[simp]: "\<omega> = a + b \<longleftrightarrow> (a = \<omega> \<or> b = \<omega>)"
+ by (cases a, cases b) auto
+
+lemma pinfreal_of_nat_le_iff:
+ "(of_nat k :: pinfreal) \<le> of_nat m \<longleftrightarrow> k \<le> m" by auto
+
+lemma pinfreal_of_nat_less_iff:
+ "(of_nat k :: pinfreal) < of_nat m \<longleftrightarrow> k < m" by auto
+
+lemma pinfreal_bound_add:
+ assumes "\<forall>N. f N + y \<le> (x::pinfreal)"
+ shows "(SUP n. f n) + y \<le> x"
+proof (cases "x = \<omega>")
+ have "y \<le> x" using assms by (auto intro: pinfreal_le_add2)
+ assume "x \<noteq> \<omega>"
+ note move_y = pinfreal_less_eq_diff_eq_sum[OF `y \<le> x` this]
+
+ have "\<forall>N. f N \<le> x - y" using assms by (simp add: move_y)
+ hence "(SUP n. f n) \<le> x - y" by (rule SUPR_bound)
+ thus ?thesis by (simp add: move_y)
+qed simp
+
+lemma SUPR_pinfreal_add:
+ fixes f g :: "nat \<Rightarrow> pinfreal"
+ assumes f: "\<forall>n. f n \<le> f (Suc n)" and g: "\<forall>n. g n \<le> g (Suc n)"
+ shows "(SUP n. f n + g n) = (SUP n. f n) + (SUP n. g n)"
+proof (rule pinfreal_SUPI)
+ fix n :: nat from le_SUPI[of n UNIV f] le_SUPI[of n UNIV g]
+ show "f n + g n \<le> (SUP n. f n) + (SUP n. g n)"
+ by (auto intro!: add_mono)
+next
+ fix y assume *: "\<And>n. n \<in> UNIV \<Longrightarrow> f n + g n \<le> y"
+ { fix n m
+ have "f n + g m \<le> y"
+ proof (cases rule: linorder_le_cases)
+ assume "n \<le> m"
+ hence "f n + g m \<le> f m + g m"
+ using f lift_Suc_mono_le by (auto intro!: add_right_mono)
+ also have "\<dots> \<le> y" using * by simp
+ finally show ?thesis .
+ next
+ assume "m \<le> n"
+ hence "f n + g m \<le> f n + g n"
+ using g lift_Suc_mono_le by (auto intro!: add_left_mono)
+ also have "\<dots> \<le> y" using * by simp
+ finally show ?thesis .
+ qed }
+ hence "\<And>m. \<forall>n. f n + g m \<le> y" by simp
+ from pinfreal_bound_add[OF this]
+ have "\<forall>m. (g m) + (SUP n. f n) \<le> y" by (simp add: ac_simps)
+ from pinfreal_bound_add[OF this]
+ show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
+qed
+
+lemma SUPR_pinfreal_setsum:
+ fixes f :: "'x \<Rightarrow> nat \<Rightarrow> pinfreal"
+ assumes "\<And>i. i \<in> P \<Longrightarrow> \<forall>n. f i n \<le> f i (Suc n)"
+ shows "(SUP n. \<Sum>i\<in>P. f i n) = (\<Sum>i\<in>P. SUP n. f i n)"
+proof cases
+ assume "finite P" from this assms show ?thesis
+ proof induct
+ case (insert i P)
+ thus ?case
+ apply simp
+ apply (subst SUPR_pinfreal_add)
+ by (auto intro!: setsum_mono)
+ qed simp
+qed simp
+
+lemma Real_max:
+ assumes "x \<ge> 0" "y \<ge> 0"
+ shows "Real (max x y) = max (Real x) (Real y)"
+ using assms unfolding max_def by (auto simp add:not_le)
+
+lemma Real_real: "Real (real x) = (if x = \<omega> then 0 else x)"
+ using assms by (cases x) auto
+
+lemma inj_on_real: "inj_on real (UNIV - {\<omega>})"
+proof (rule inj_onI)
+ fix x y assume mem: "x \<in> UNIV - {\<omega>}" "y \<in> UNIV - {\<omega>}" and "real x = real y"
+ thus "x = y" by (cases x, cases y) auto
+qed
+
+lemma inj_on_Real: "inj_on Real {0..}"
+ by (auto intro!: inj_onI)
+
+lemma range_Real[simp]: "range Real = UNIV - {\<omega>}"
+proof safe
+ fix x assume "x \<notin> range Real"
+ thus "x = \<omega>" by (cases x) auto
+qed auto
+
+lemma image_Real[simp]: "Real ` {0..} = UNIV - {\<omega>}"
+proof safe
+ fix x assume "x \<notin> Real ` {0..}"
+ thus "x = \<omega>" by (cases x) auto
+qed auto
+
+lemma pinfreal_SUP_cmult:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ shows "(SUP i : R. z * f i) = z * (SUP i : R. f i)"
+proof (rule pinfreal_SUPI)
+ fix i assume "i \<in> R"
+ from le_SUPI[OF this]
+ show "z * f i \<le> z * (SUP i:R. f i)" by (rule pinfreal_mult_cancel)
+next
+ fix y assume "\<And>i. i\<in>R \<Longrightarrow> z * f i \<le> y"
+ hence *: "\<And>i. i\<in>R \<Longrightarrow> z * f i \<le> y" by auto
+ show "z * (SUP i:R. f i) \<le> y"
+ proof (cases "\<forall>i\<in>R. f i = 0")
+ case True
+ show ?thesis
+ proof cases
+ assume "R \<noteq> {}" hence "f ` R = {0}" using True by auto
+ thus ?thesis by (simp add: SUPR_def)
+ qed (simp add: SUPR_def Sup_empty bot_pinfreal_def)
+ next
+ case False then obtain i where i: "i \<in> R" and f0: "f i \<noteq> 0" by auto
+ show ?thesis
+ proof (cases "z = 0 \<or> z = \<omega>")
+ case True with f0 *[OF i] show ?thesis by auto
+ next
+ case False hence z: "z \<noteq> 0" "z \<noteq> \<omega>" by auto
+ note div = pinfreal_inverse_le_eq[OF this, symmetric]
+ hence "\<And>i. i\<in>R \<Longrightarrow> f i \<le> y / z" using * by auto
+ thus ?thesis unfolding div SUP_le_iff by simp
+ qed
+ qed
+qed
+
+instantiation pinfreal :: topological_space
+begin
+
+definition "open A \<longleftrightarrow>
+ (\<exists>T. open T \<and> (Real ` (T\<inter>{0..}) = A - {\<omega>})) \<and> (\<omega> \<in> A \<longrightarrow> (\<exists>x\<ge>0. {Real x <..} \<subseteq> A))"
+
+lemma open_omega: "open A \<Longrightarrow> \<omega> \<in> A \<Longrightarrow> (\<exists>x\<ge>0. {Real x<..} \<subseteq> A)"
+ unfolding open_pinfreal_def by auto
+
+lemma open_omegaD: assumes "open A" "\<omega> \<in> A" obtains x where "x\<ge>0" "{Real x<..} \<subseteq> A"
+ using open_omega[OF assms] by auto
+
+lemma pinfreal_openE: assumes "open A" obtains A' x where
+ "open A'" "Real ` (A' \<inter> {0..}) = A - {\<omega>}"
+ "x \<ge> 0" "\<omega> \<in> A \<Longrightarrow> {Real x<..} \<subseteq> A"
+ using assms open_pinfreal_def by auto
+
+instance
+proof
+ let ?U = "UNIV::pinfreal set"
+ show "open ?U" unfolding open_pinfreal_def
+ by (auto intro!: exI[of _ "UNIV"] exI[of _ 0])
+next
+ fix S T::"pinfreal set" assume "open S" and "open T"
+ from `open S`[THEN pinfreal_openE] guess S' xS . note S' = this
+ from `open T`[THEN pinfreal_openE] guess T' xT . note T' = this
+
+ from S'(1-3) T'(1-3)
+ show "open (S \<inter> T)" unfolding open_pinfreal_def
+ proof (safe intro!: exI[of _ "S' \<inter> T'"] exI[of _ "max xS xT"])
+ fix x assume *: "Real (max xS xT) < x" and "\<omega> \<in> S" "\<omega> \<in> T"
+ from `\<omega> \<in> S`[THEN S'(4)] * show "x \<in> S"
+ by (cases x, auto simp: max_def split: split_if_asm)
+ from `\<omega> \<in> T`[THEN T'(4)] * show "x \<in> T"
+ by (cases x, auto simp: max_def split: split_if_asm)
+ next
+ fix x assume x: "x \<notin> Real ` (S' \<inter> T' \<inter> {0..})"
+ have *: "S' \<inter> T' \<inter> {0..} = (S' \<inter> {0..}) \<inter> (T' \<inter> {0..})" by auto
+ assume "x \<in> T" "x \<in> S"
+ with S'(2) T'(2) show "x = \<omega>"
+ using x[unfolded *] inj_on_image_Int[OF inj_on_Real] by auto
+ qed auto
+next
+ fix K assume openK: "\<forall>S \<in> K. open (S:: pinfreal set)"
+ hence "\<forall>S\<in>K. \<exists>T. open T \<and> Real ` (T \<inter> {0..}) = S - {\<omega>}" by (auto simp: open_pinfreal_def)
+ from bchoice[OF this] guess T .. note T = this[rule_format]
+
+ show "open (\<Union>K)" unfolding open_pinfreal_def
+ proof (safe intro!: exI[of _ "\<Union>(T ` K)"])
+ fix x S assume "0 \<le> x" "x \<in> T S" "S \<in> K"
+ with T[OF `S \<in> K`] show "Real x \<in> \<Union>K" by auto
+ next
+ fix x S assume x: "x \<notin> Real ` (\<Union>T ` K \<inter> {0..})" "S \<in> K" "x \<in> S"
+ hence "x \<notin> Real ` (T S \<inter> {0..})"
+ by (auto simp: image_UN UN_simps[symmetric] simp del: UN_simps)
+ thus "x = \<omega>" using T[OF `S \<in> K`] `x \<in> S` by auto
+ next
+ fix S assume "\<omega> \<in> S" "S \<in> K"
+ from openK[rule_format, OF `S \<in> K`, THEN pinfreal_openE] guess S' x .
+ from this(3, 4) `\<omega> \<in> S`
+ show "\<exists>x\<ge>0. {Real x<..} \<subseteq> \<Union>K"
+ by (auto intro!: exI[of _ x] bexI[OF _ `S \<in> K`])
+ next
+ from T[THEN conjunct1] show "open (\<Union>T`K)" by auto
+ qed auto
+qed
+end
+
+lemma open_pinfreal_lessThan[simp]:
+ "open {..< a :: pinfreal}"
+proof (cases a)
+ case (preal x) thus ?thesis unfolding open_pinfreal_def
+ proof (safe intro!: exI[of _ "{..< x}"])
+ fix y assume "y < Real x"
+ moreover assume "y \<notin> Real ` ({..<x} \<inter> {0..})"
+ ultimately have "y \<noteq> Real (real y)" using preal by (cases y) auto
+ thus "y = \<omega>" by (auto simp: Real_real split: split_if_asm)
+ qed auto
+next
+ case infinite thus ?thesis
+ unfolding open_pinfreal_def by (auto intro!: exI[of _ UNIV])
+qed
+
+lemma open_pinfreal_greaterThan[simp]:
+ "open {a :: pinfreal <..}"
+proof (cases a)
+ case (preal x) thus ?thesis unfolding open_pinfreal_def
+ proof (safe intro!: exI[of _ "{x <..}"])
+ fix y assume "Real x < y"
+ moreover assume "y \<notin> Real ` ({x<..} \<inter> {0..})"
+ ultimately have "y \<noteq> Real (real y)" using preal by (cases y) auto
+ thus "y = \<omega>" by (auto simp: Real_real split: split_if_asm)
+ qed auto
+next
+ case infinite thus ?thesis
+ unfolding open_pinfreal_def by (auto intro!: exI[of _ "{}"])
+qed
+
+lemma pinfreal_open_greaterThanLessThan[simp]: "open {a::pinfreal <..< b}"
+ unfolding greaterThanLessThan_def by auto
+
+lemma closed_pinfreal_atLeast[simp, intro]: "closed {a :: pinfreal ..}"
+proof -
+ have "- {a ..} = {..< a}" by auto
+ then show "closed {a ..}"
+ unfolding closed_def using open_pinfreal_lessThan by auto
+qed
+
+lemma closed_pinfreal_atMost[simp, intro]: "closed {.. b :: pinfreal}"
+proof -
+ have "- {.. b} = {b <..}" by auto
+ then show "closed {.. b}"
+ unfolding closed_def using open_pinfreal_greaterThan by auto
+qed
+
+lemma closed_pinfreal_atLeastAtMost[simp, intro]:
+ shows "closed {a :: pinfreal .. b}"
+ unfolding atLeastAtMost_def by auto
+
+lemma pinfreal_dense:
+ fixes x y :: pinfreal assumes "x < y"
+ shows "\<exists>z. x < z \<and> z < y"
+proof -
+ from `x < y` obtain p where p: "x = Real p" "0 \<le> p" by (cases x) auto
+ show ?thesis
+ proof (cases y)
+ case (preal r) with p `x < y` have "p < r" by auto
+ with dense obtain z where "p < z" "z < r" by auto
+ thus ?thesis using preal p by (auto intro!: exI[of _ "Real z"])
+ next
+ case infinite thus ?thesis using `x < y` p
+ by (auto intro!: exI[of _ "Real p + 1"])
+ qed
+qed
+
+instance pinfreal :: t2_space
+proof
+ fix x y :: pinfreal assume "x \<noteq> y"
+ let "?P x (y::pinfreal)" = "\<exists> U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+
+ { fix x y :: pinfreal assume "x < y"
+ from pinfreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
+ have "?P x y"
+ apply (rule exI[of _ "{..<z}"])
+ apply (rule exI[of _ "{z<..}"])
+ using z by auto }
+ note * = this
+
+ from `x \<noteq> y`
+ show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
+ proof (cases rule: linorder_cases)
+ assume "x = y" with `x \<noteq> y` show ?thesis by simp
+ next assume "x < y" from *[OF this] show ?thesis by auto
+ next assume "y < x" from *[OF this] show ?thesis by auto
+ qed
+qed
+
+definition (in complete_lattice) isoton :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<up>" 50) where
+ "A \<up> X \<longleftrightarrow> (\<forall>i. A i \<le> A (Suc i)) \<and> (SUP i. A i) = X"
+
+definition (in complete_lattice) antiton (infix "\<down>" 50) where
+ "A \<down> X \<longleftrightarrow> (\<forall>i. A i \<ge> A (Suc i)) \<and> (INF i. A i) = X"
+
+lemma range_const[simp]: "range (\<lambda>x. c) = {c}" by auto
+
+lemma isoton_cmult_right:
+ assumes "f \<up> (x::pinfreal)"
+ shows "(\<lambda>i. c * f i) \<up> (c * x)"
+ using assms unfolding isoton_def pinfreal_SUP_cmult
+ by (auto intro: pinfreal_mult_cancel)
+
+lemma isoton_cmult_left:
+ "f \<up> (x::pinfreal) \<Longrightarrow> (\<lambda>i. f i * c) \<up> (x * c)"
+ by (subst (1 2) mult_commute) (rule isoton_cmult_right)
+
+lemma isoton_add:
+ assumes "f \<up> (x::pinfreal)" and "g \<up> y"
+ shows "(\<lambda>i. f i + g i) \<up> (x + y)"
+ using assms unfolding isoton_def
+ by (auto intro: pinfreal_mult_cancel add_mono simp: SUPR_pinfreal_add)
+
+lemma isoton_fun_expand:
+ "f \<up> x \<longleftrightarrow> (\<forall>i. (\<lambda>j. f j i) \<up> (x i))"
+proof -
+ have "\<And>i. {y. \<exists>f'\<in>range f. y = f' i} = range (\<lambda>j. f j i)"
+ by auto
+ with assms show ?thesis
+ by (auto simp add: isoton_def le_fun_def Sup_fun_def SUPR_def)
+qed
+
+lemma isoton_indicator:
+ assumes "f \<up> g"
+ shows "(\<lambda>i x. f i x * indicator A x) \<up> (\<lambda>x. g x * indicator A x :: pinfreal)"
+ using assms unfolding isoton_fun_expand by (auto intro!: isoton_cmult_left)
+
+lemma pinfreal_ord_one[simp]:
+ "Real p < 1 \<longleftrightarrow> p < 1"
+ "Real p \<le> 1 \<longleftrightarrow> p \<le> 1"
+ "1 < Real p \<longleftrightarrow> 1 < p"
+ "1 \<le> Real p \<longleftrightarrow> 1 \<le> p"
+ by (simp_all add: one_pinfreal_def del: Real_1)
+
+lemma SUP_mono:
+ assumes "\<And>n. f n \<le> g (N n)"
+ shows "(SUP n. f n) \<le> (SUP n. g n)"
+proof (safe intro!: SUPR_bound)
+ fix n note assms[of n]
+ also have "g (N n) \<le> (SUP n. g n)" by (auto intro!: le_SUPI)
+ finally show "f n \<le> (SUP n. g n)" .
+qed
+
+lemma isoton_Sup:
+ assumes "f \<up> u"
+ shows "f i \<le> u"
+ using le_SUPI[of i UNIV f] assms
+ unfolding isoton_def by auto
+
+lemma isoton_mono:
+ assumes iso: "x \<up> a" "y \<up> b" and *: "\<And>n. x n \<le> y (N n)"
+ shows "a \<le> b"
+proof -
+ from iso have "a = (SUP n. x n)" "b = (SUP n. y n)"
+ unfolding isoton_def by auto
+ with * show ?thesis by (auto intro!: SUP_mono)
+qed
+
+lemma pinfreal_le_mult_one_interval:
+ fixes x y :: pinfreal
+ assumes "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
+ shows "x \<le> y"
+proof (cases x, cases y)
+ assume "x = \<omega>"
+ with assms[of "1 / 2"]
+ show "x \<le> y" by simp
+next
+ fix r p assume *: "y = Real p" "x = Real r" and **: "0 \<le> r" "0 \<le> p"
+ have "r \<le> p"
+ proof (rule field_le_mult_one_interval)
+ fix z :: real assume "0 < z" and "z < 1"
+ with assms[of "Real z"]
+ show "z * r \<le> p" using ** * by (auto simp: zero_le_mult_iff)
+ qed
+ thus "x \<le> y" using ** * by simp
+qed simp
+
+lemma pinfreal_greater_0[intro]:
+ fixes a :: pinfreal
+ assumes "a \<noteq> 0"
+ shows "a > 0"
+using assms apply (cases a) by auto
+
+lemma pinfreal_mult_strict_right_mono:
+ assumes "a < b" and "0 < c" "c < \<omega>"
+ shows "a * c < b * c"
+ using assms
+ by (cases a, cases b, cases c)
+ (auto simp: zero_le_mult_iff pinfreal_less_\<omega>)
+
+lemma minus_pinfreal_eq2:
+ fixes x y z :: pinfreal
+ assumes "y \<le> x" and "y \<noteq> \<omega>" shows "z = x - y \<longleftrightarrow> z + y = x"
+ using assms
+ apply (subst eq_commute)
+ apply (subst minus_pinfreal_eq)
+ by (cases x, cases z, auto simp add: ac_simps not_less)
+
+lemma pinfreal_diff_eq_diff_imp_eq:
+ assumes "a \<noteq> \<omega>" "b \<le> a" "c \<le> a"
+ assumes "a - b = a - c"
+ shows "b = c"
+ using assms
+ by (cases a, cases b, cases c) (auto split: split_if_asm)
+
+lemma pinfreal_inverse_eq_0: "inverse x = 0 \<longleftrightarrow> x = \<omega>"
+ by (cases x) auto
+
+lemma pinfreal_mult_inverse:
+ "\<lbrakk> x \<noteq> \<omega> ; x \<noteq> 0 \<rbrakk> \<Longrightarrow> x * inverse x = 1"
+ by (cases x) auto
+
+lemma pinfreal_zero_less_diff_iff:
+ fixes a b :: pinfreal shows "0 < a - b \<longleftrightarrow> b < a"
+ apply (cases a, cases b)
+ apply (auto simp: pinfreal_noteq_omega_Ex pinfreal_less_\<omega>)
+ apply (cases b)
+ by auto
+
+lemma pinfreal_less_Real_Ex:
+ fixes a b :: pinfreal shows "x < Real r \<longleftrightarrow> (\<exists>p\<ge>0. p < r \<and> x = Real p)"
+ by (cases x) auto
+
+lemma open_Real: assumes "open S" shows "open (Real ` ({0..} \<inter> S))"
+ unfolding open_pinfreal_def apply(rule,rule,rule,rule assms) by auto
+
+lemma pinfreal_zero_le_diff:
+ fixes a b :: pinfreal shows "a - b = 0 \<longleftrightarrow> a \<le> b"
+ by (cases a, cases b, simp_all, cases b, auto)
+
+lemma lim_Real[simp]: assumes "\<forall>n. f n \<ge> 0" "m\<ge>0"
+ shows "(\<lambda>n. Real (f n)) ----> Real m \<longleftrightarrow> (\<lambda>n. f n) ----> m" (is "?l = ?r")
+proof assume ?l show ?r unfolding Lim_sequentially
+ proof safe fix e::real assume e:"e>0"
+ note open_ball[of m e] note open_Real[OF this]
+ note * = `?l`[unfolded tendsto_def,rule_format,OF this]
+ have "eventually (\<lambda>x. Real (f x) \<in> Real ` ({0..} \<inter> ball m e)) sequentially"
+ apply(rule *) unfolding image_iff using assms(2) e by auto
+ thus "\<exists>N. \<forall>n\<ge>N. dist (f n) m < e" unfolding eventually_sequentially
+ apply safe apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
+ proof- fix n x assume "Real (f n) = Real x" "0 \<le> x"
+ hence *:"f n = x" using assms(1) by auto
+ assume "x \<in> ball m e" thus "dist (f n) m < e" unfolding *
+ by (auto simp add:dist_commute)
+ qed qed
+next assume ?r show ?l unfolding tendsto_def eventually_sequentially
+ proof safe fix S assume S:"open S" "Real m \<in> S"
+ guess T y using S(1) apply-apply(erule pinfreal_openE) . note T=this
+ have "m\<in>real ` (S - {\<omega>})" unfolding image_iff
+ apply(rule_tac x="Real m" in bexI) using assms(2) S(2) by auto
+ hence "m \<in> T" unfolding T(2)[THEN sym] by auto
+ from `?r`[unfolded tendsto_def eventually_sequentially,rule_format,OF T(1) this]
+ guess N .. note N=this[rule_format]
+ show "\<exists>N. \<forall>n\<ge>N. Real (f n) \<in> S" apply(rule_tac x=N in exI)
+ proof safe fix n assume n:"N\<le>n"
+ have "f n \<in> real ` (S - {\<omega>})" using N[OF n] assms unfolding T(2)[THEN sym]
+ unfolding image_iff apply-apply(rule_tac x="Real (f n)" in bexI)
+ unfolding real_Real by auto
+ then guess x unfolding image_iff .. note x=this
+ show "Real (f n) \<in> S" unfolding x apply(subst Real_real) using x by auto
+ qed
+ qed
+qed
+
+lemma pinfreal_INFI:
+ fixes x :: pinfreal
+ assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> f i"
+ assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> f i) \<Longrightarrow> y \<le> x"
+ shows "(INF i:A. f i) = x"
+ unfolding INFI_def Inf_pinfreal_def
+ using assms by (auto intro!: Greatest_equality)
+
+lemma real_of_pinfreal_less:"x < y \<Longrightarrow> y\<noteq>\<omega> \<Longrightarrow> real x < real y"
+proof- case goal1
+ have *:"y = Real (real y)" "x = Real (real x)" using goal1 Real_real by auto
+ show ?case using goal1 apply- apply(subst(asm) *(1))apply(subst(asm) *(2))
+ unfolding pinfreal_less by auto
+qed
+
+lemma not_less_omega[simp]:"\<not> x < \<omega> \<longleftrightarrow> x = \<omega>"
+ by (metis antisym_conv3 pinfreal_less(3))
+
+lemma Real_real': assumes "x\<noteq>\<omega>" shows "Real (real x) = x"
+proof- have *:"(THE r. 0 \<le> r \<and> x = Real r) = real x"
+ apply(rule the_equality) using assms unfolding Real_real by auto
+ have "Real (THE r. 0 \<le> r \<and> x = Real r) = x" unfolding *
+ using assms unfolding Real_real by auto
+ thus ?thesis unfolding real_of_pinfreal_def of_pinfreal_def
+ unfolding pinfreal_case_def using assms by auto
+qed
+
+lemma Real_less_plus_one:"Real x < Real (max (x + 1) 1)"
+ unfolding pinfreal_less by auto
+
+lemma Lim_omega: "f ----> \<omega> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n \<ge> Real B)" (is "?l = ?r")
+proof assume ?r show ?l apply(rule topological_tendstoI)
+ unfolding eventually_sequentially
+ proof- fix S assume "open S" "\<omega> \<in> S"
+ from open_omega[OF this] guess B .. note B=this
+ from `?r`[rule_format,of "(max B 0)+1"] guess N .. note N=this
+ show "\<exists>N. \<forall>n\<ge>N. f n \<in> S" apply(rule_tac x=N in exI)
+ proof safe case goal1
+ have "Real B < Real ((max B 0) + 1)" by auto
+ also have "... \<le> f n" using goal1 N by auto
+ finally show ?case using B by fastsimp
+ qed
+ qed
+next assume ?l show ?r
+ proof fix B::real have "open {Real B<..}" "\<omega> \<in> {Real B<..}" by auto
+ from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
+ guess N .. note N=this
+ show "\<exists>N. \<forall>n\<ge>N. Real B \<le> f n" apply(rule_tac x=N in exI) using N by auto
+ qed
+qed
+
+lemma Lim_bounded_omgea: assumes lim:"f ----> l" and "\<And>n. f n \<le> Real B" shows "l \<noteq> \<omega>"
+proof(rule ccontr,unfold not_not) let ?B = "max (B + 1) 1" assume as:"l=\<omega>"
+ from lim[unfolded this Lim_omega,rule_format,of "?B"]
+ guess N .. note N=this[rule_format,OF le_refl]
+ hence "Real ?B \<le> Real B" using assms(2)[of N] by(rule order_trans)
+ hence "Real ?B < Real ?B" using Real_less_plus_one[of B] by(rule le_less_trans)
+ thus False by auto
+qed
+
+lemma incseq_le_pinfreal: assumes inc: "\<And>n m. n\<ge>m \<Longrightarrow> X n \<ge> X m"
+ and lim: "X ----> (L::pinfreal)" shows "X n \<le> L"
+proof(cases "L = \<omega>")
+ case False have "\<forall>n. X n \<noteq> \<omega>"
+ proof(rule ccontr,unfold not_all not_not,safe)
+ case goal1 hence "\<forall>n\<ge>x. X n = \<omega>" using inc[of x] by auto
+ hence "X ----> \<omega>" unfolding tendsto_def eventually_sequentially
+ apply safe apply(rule_tac x=x in exI) by auto
+ note Lim_unique[OF trivial_limit_sequentially this lim]
+ with False show False by auto
+ qed note * =this[rule_format]
+
+ have **:"\<forall>m n. m \<le> n \<longrightarrow> Real (real (X m)) \<le> Real (real (X n))"
+ unfolding Real_real using * inc by auto
+ have "real (X n) \<le> real L" apply-apply(rule incseq_le) defer
+ apply(subst lim_Real[THEN sym]) apply(rule,rule,rule)
+ unfolding Real_real'[OF *] Real_real'[OF False]
+ unfolding incseq_def using ** lim by auto
+ hence "Real (real (X n)) \<le> Real (real L)" by auto
+ thus ?thesis unfolding Real_real using * False by auto
+qed auto
+
+lemma SUP_Lim_pinfreal: assumes "\<And>n m. n\<ge>m \<Longrightarrow> f n \<ge> f m" "f ----> l"
+ shows "(SUP n. f n) = (l::pinfreal)" unfolding SUPR_def Sup_pinfreal_def
+proof (safe intro!: Least_equality)
+ fix n::nat show "f n \<le> l" apply(rule incseq_le_pinfreal)
+ using assms by auto
+next fix y assume y:"\<forall>x\<in>range f. x \<le> y" show "l \<le> y"
+ proof(rule ccontr,cases "y=\<omega>",unfold not_le)
+ case False assume as:"y < l"
+ have l:"l \<noteq> \<omega>" apply(rule Lim_bounded_omgea[OF assms(2), of "real y"])
+ using False y unfolding Real_real by auto
+
+ have yl:"real y < real l" using as apply-
+ apply(subst(asm) Real_real'[THEN sym,OF `y\<noteq>\<omega>`])
+ apply(subst(asm) Real_real'[THEN sym,OF `l\<noteq>\<omega>`])
+ unfolding pinfreal_less apply(subst(asm) if_P) by auto
+ hence "y + (y - l) * Real (1 / 2) < l" apply-
+ apply(subst Real_real'[THEN sym,OF `y\<noteq>\<omega>`]) apply(subst(2) Real_real'[THEN sym,OF `y\<noteq>\<omega>`])
+ apply(subst Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) apply(subst(2) Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) by auto
+ hence *:"l \<in> {y + (y - l) / 2<..}" by auto
+ have "open {y + (y-l)/2 <..}" by auto
+ note topological_tendstoD[OF assms(2) this *]
+ from this[unfolded eventually_sequentially] guess N .. note this[rule_format, of N]
+ hence "y + (y - l) * Real (1 / 2) < y" using y[rule_format,of "f N"] by auto
+ hence "Real (real y) + (Real (real y) - Real (real l)) * Real (1 / 2) < Real (real y)"
+ unfolding Real_real using `y\<noteq>\<omega>` `l\<noteq>\<omega>` by auto
+ thus False using yl by auto
+ qed auto
+qed
+
+lemma Real_max':"Real x = Real (max x 0)"
+proof(cases "x < 0") case True
+ hence *:"max x 0 = 0" by auto
+ show ?thesis unfolding * using True by auto
+qed auto
+
+lemma lim_pinfreal_increasing: assumes "\<forall>n m. n\<ge>m \<longrightarrow> f n \<ge> f m"
+ obtains l where "f ----> (l::pinfreal)"
+proof(cases "\<exists>B. \<forall>n. f n < Real B")
+ case False thus thesis apply- apply(rule that[of \<omega>]) unfolding Lim_omega not_ex not_all
+ apply safe apply(erule_tac x=B in allE,safe) apply(rule_tac x=x in exI,safe)
+ apply(rule order_trans[OF _ assms[rule_format]]) by auto
+next case True then guess B .. note B = this[rule_format]
+ hence *:"\<And>n. f n < \<omega>" apply-apply(rule less_le_trans,assumption) by auto
+ have *:"\<And>n. f n \<noteq> \<omega>" proof- case goal1 show ?case using *[of n] by auto qed
+ have B':"\<And>n. real (f n) \<le> max 0 B" proof- case goal1 thus ?case
+ using B[of n] apply-apply(subst(asm) Real_real'[THEN sym]) defer
+ apply(subst(asm)(2) Real_max') unfolding pinfreal_less apply(subst(asm) if_P) using *[of n] by auto
+ qed
+ have "\<exists>l. (\<lambda>n. real (f n)) ----> l" apply(rule Topology_Euclidean_Space.bounded_increasing_convergent)
+ proof safe show "bounded {real (f n) |n. True}"
+ unfolding bounded_def apply(rule_tac x=0 in exI,rule_tac x="max 0 B" in exI)
+ using B' unfolding dist_norm by auto
+ fix n::nat have "Real (real (f n)) \<le> Real (real (f (Suc n)))"
+ using assms[rule_format,of n "Suc n"] apply(subst Real_real)+
+ using *[of n] *[of "Suc n"] by fastsimp
+ thus "real (f n) \<le> real (f (Suc n))" by auto
+ qed then guess l .. note l=this
+ have "0 \<le> l" apply(rule LIMSEQ_le_const[OF l])
+ by(rule_tac x=0 in exI,auto)
+
+ thus ?thesis apply-apply(rule that[of "Real l"])
+ using l apply-apply(subst(asm) lim_Real[THEN sym]) prefer 3
+ unfolding Real_real using * by auto
+qed
+
+lemma setsum_neq_omega: assumes "finite s" "\<And>x. x \<in> s \<Longrightarrow> f x \<noteq> \<omega>"
+ shows "setsum f s \<noteq> \<omega>" using assms
+proof induct case (insert x s)
+ show ?case unfolding setsum.insert[OF insert(1-2)]
+ using insert by auto
+qed auto
+
+
+lemma real_Real': "0 \<le> x \<Longrightarrow> real (Real x) = x"
+ unfolding real_Real by auto
+
+lemma real_pinfreal_pos[intro]:
+ assumes "x \<noteq> 0" "x \<noteq> \<omega>"
+ shows "real x > 0"
+ apply(subst real_Real'[THEN sym,of 0]) defer
+ apply(rule real_of_pinfreal_less) using assms by auto
+
+lemma Lim_omega_gt: "f ----> \<omega> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n > Real B)" (is "?l = ?r")
+proof assume ?l thus ?r unfolding Lim_omega apply safe
+ apply(erule_tac x="max B 0 +1" in allE,safe)
+ apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
+ apply(rule_tac y="Real (max B 0 + 1)" in less_le_trans) by auto
+next assume ?r thus ?l unfolding Lim_omega apply safe
+ apply(erule_tac x=B in allE,safe) apply(rule_tac x=N in exI,safe) by auto
+qed
+
+lemma pinfreal_minus_le_cancel:
+ fixes a b c :: pinfreal
+ assumes "b \<le> a"
+ shows "c - a \<le> c - b"
+ using assms by (cases a, cases b, cases c, simp, simp, simp, cases b, cases c, simp_all)
+
+lemma pinfreal_minus_\<omega>[simp]: "x - \<omega> = 0" by (cases x) simp_all
+
+lemma pinfreal_minus_mono[intro]: "a - x \<le> (a::pinfreal)"
+proof- have "a - x \<le> a - 0"
+ apply(rule pinfreal_minus_le_cancel) by auto
+ thus ?thesis by auto
+qed
+
+lemma pinfreal_minus_eq_\<omega>[simp]: "x - y = \<omega> \<longleftrightarrow> (x = \<omega> \<and> y \<noteq> \<omega>)"
+ by (cases x, cases y) (auto, cases y, auto)
+
+lemma pinfreal_less_minus_iff:
+ fixes a b c :: pinfreal
+ shows "a < b - c \<longleftrightarrow> c + a < b"
+ by (cases c, cases a, cases b, auto)
+
+lemma pinfreal_minus_less_iff:
+ fixes a b c :: pinfreal shows "a - c < b \<longleftrightarrow> (0 < b \<and> (c \<noteq> \<omega> \<longrightarrow> a < b + c))"
+ by (cases c, cases a, cases b, auto)
+
+lemma pinfreal_le_minus_iff:
+ fixes a b c :: pinfreal
+ shows "a \<le> c - b \<longleftrightarrow> ((c \<le> b \<longrightarrow> a = 0) \<and> (b < c \<longrightarrow> a + b \<le> c))"
+ by (cases a, cases c, cases b, auto simp: pinfreal_noteq_omega_Ex)
+
+lemma pinfreal_minus_le_iff:
+ fixes a b c :: pinfreal
+ shows "a - c \<le> b \<longleftrightarrow> (c \<le> a \<longrightarrow> a \<le> b + c)"
+ by (cases a, cases c, cases b, auto simp: pinfreal_noteq_omega_Ex)
+
+lemmas pinfreal_minus_order = pinfreal_minus_le_iff pinfreal_minus_less_iff pinfreal_le_minus_iff pinfreal_less_minus_iff
+
+lemma pinfreal_minus_strict_mono:
+ assumes "a > 0" "x > 0" "a\<noteq>\<omega>"
+ shows "a - x < (a::pinfreal)"
+ using assms by(cases x, cases a, auto)
+
+lemma pinfreal_minus':
+ "Real r - Real p = (if 0 \<le> r \<and> p \<le> r then if 0 \<le> p then Real (r - p) else Real r else 0)"
+ by (auto simp: minus_pinfreal_eq not_less)
+
+lemma pinfreal_minus_plus:
+ "x \<le> (a::pinfreal) \<Longrightarrow> a - x + x = a"
+ by (cases a, cases x) auto
+
+lemma pinfreal_cancel_plus_minus: "b \<noteq> \<omega> \<Longrightarrow> a + b - b = a"
+ by (cases a, cases b) auto
+
+lemma pinfreal_minus_le_cancel_right:
+ fixes a b c :: pinfreal
+ assumes "a \<le> b" "c \<le> a"
+ shows "a - c \<le> b - c"
+ using assms by (cases a, cases b, cases c, auto, cases c, auto)
+
+lemma real_of_pinfreal_setsum':
+ assumes "\<forall>x \<in> S. f x \<noteq> \<omega>"
+ shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
+proof cases
+ assume "finite S"
+ from this assms show ?thesis
+ by induct (simp_all add: real_of_pinfreal_add setsum_\<omega>)
+qed simp
+
+lemma Lim_omega_pos: "f ----> \<omega> \<longleftrightarrow> (\<forall>B>0. \<exists>N. \<forall>n\<ge>N. f n \<ge> Real B)" (is "?l = ?r")
+ unfolding Lim_omega apply safe defer
+ apply(erule_tac x="max 1 B" in allE) apply safe defer
+ apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
+ apply(rule_tac y="Real (max 1 B)" in order_trans) by auto
+
+lemma (in complete_lattice) isotonD[dest]:
+ assumes "A \<up> X" shows "A i \<le> A (Suc i)" "(SUP i. A i) = X"
+ using assms unfolding isoton_def by auto
+
+lemma isotonD'[dest]:
+ assumes "(A::_=>_) \<up> X" shows "A i x \<le> A (Suc i) x" "(SUP i. A i) = X"
+ using assms unfolding isoton_def le_fun_def by auto
+
+lemma pinfreal_LimI_finite:
+ assumes "x \<noteq> \<omega>" "\<And>r. 0 < r \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. u n < x + r \<and> x < u n + r"
+ shows "u ----> x"
+proof (rule topological_tendstoI, unfold eventually_sequentially)
+ fix S assume "open S" "x \<in> S"
+ then obtain A where "open A" and A_eq: "Real ` (A \<inter> {0..}) = S - {\<omega>}" by (auto elim!: pinfreal_openE)
+ then have "x \<in> Real ` (A \<inter> {0..})" using `x \<in> S` `x \<noteq> \<omega>` by auto
+ then have "real x \<in> A" by auto
+ then obtain r where "0 < r" and dist: "\<And>y. dist y (real x) < r \<Longrightarrow> y \<in> A"
+ using `open A` unfolding open_real_def by auto
+ then obtain n where
+ upper: "\<And>N. n \<le> N \<Longrightarrow> u N < x + Real r" and
+ lower: "\<And>N. n \<le> N \<Longrightarrow> x < u N + Real r" using assms(2)[of "Real r"] by auto
+ show "\<exists>N. \<forall>n\<ge>N. u n \<in> S"
+ proof (safe intro!: exI[of _ n])
+ fix N assume "n \<le> N"
+ from upper[OF this] `x \<noteq> \<omega>` `0 < r`
+ have "u N \<noteq> \<omega>" by (force simp: pinfreal_noteq_omega_Ex)
+ with `x \<noteq> \<omega>` `0 < r` lower[OF `n \<le> N`] upper[OF `n \<le> N`]
+ have "dist (real (u N)) (real x) < r" "u N \<noteq> \<omega>"
+ by (auto simp: pinfreal_noteq_omega_Ex dist_real_def abs_diff_less_iff field_simps)
+ from dist[OF this(1)]
+ have "u N \<in> Real ` (A \<inter> {0..})" using `u N \<noteq> \<omega>`
+ by (auto intro!: image_eqI[of _ _ "real (u N)"] simp: pinfreal_noteq_omega_Ex Real_real)
+ thus "u N \<in> S" using A_eq by simp
+ qed
+qed
+
+lemma real_Real_max:"real (Real x) = max x 0"
+ unfolding real_Real by auto
+
+lemma (in complete_lattice) SUPR_upper:
+ "x \<in> A \<Longrightarrow> f x \<le> SUPR A f"
+ unfolding SUPR_def apply(rule Sup_upper) by auto
+
+lemma (in complete_lattice) SUPR_subset:
+ assumes "A \<subseteq> B" shows "SUPR A f \<le> SUPR B f"
+ apply(rule SUP_leI) apply(rule SUPR_upper) using assms by auto
+
+lemma (in complete_lattice) SUPR_mono:
+ assumes "\<forall>a\<in>A. \<exists>b\<in>B. f b \<ge> f a"
+ shows "SUPR A f \<le> SUPR B f"
+ unfolding SUPR_def apply(rule Sup_mono)
+ using assms by auto
+
+lemma Sup_lim:
+ assumes "\<forall>n. b n \<in> s" "b ----> (a::pinfreal)"
+ shows "a \<le> Sup s"
+proof(rule ccontr,unfold not_le)
+ assume as:"Sup s < a" hence om:"Sup s \<noteq> \<omega>" by auto
+ have s:"s \<noteq> {}" using assms by auto
+ { presume *:"\<forall>n. b n < a \<Longrightarrow> False"
+ show False apply(cases,rule *,assumption,unfold not_all not_less)
+ proof- case goal1 then guess n .. note n=this
+ thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]]
+ using as by auto
+ qed
+ } assume b:"\<forall>n. b n < a"
+ show False
+ proof(cases "a = \<omega>")
+ case False have *:"a - Sup s > 0"
+ using False as by(auto simp: pinfreal_zero_le_diff)
+ have "(a - Sup s) / 2 \<le> a / 2" unfolding divide_pinfreal_def
+ apply(rule mult_right_mono) by auto
+ also have "... = Real (real (a / 2))" apply(rule Real_real'[THEN sym])
+ using False by auto
+ also have "... < Real (real a)" unfolding pinfreal_less using as False
+ by(auto simp add: real_of_pinfreal_mult[THEN sym])
+ also have "... = a" apply(rule Real_real') using False by auto
+ finally have asup:"a > (a - Sup s) / 2" .
+ have "\<exists>n. a - b n < (a - Sup s) / 2"
+ proof(rule ccontr,unfold not_ex not_less)
+ case goal1
+ have "(a - Sup s) * Real (1 / 2) > 0"
+ using * by auto
+ hence "a - (a - Sup s) * Real (1 / 2) < a"
+ apply-apply(rule pinfreal_minus_strict_mono)
+ using False * by auto
+ hence *:"a \<in> {a - (a - Sup s) / 2<..}"using asup by auto
+ note topological_tendstoD[OF assms(2) open_pinfreal_greaterThan,OF *]
+ from this[unfolded eventually_sequentially] guess n ..
+ note n = this[rule_format,of n]
+ have "b n + (a - Sup s) / 2 \<le> a"
+ using add_right_mono[OF goal1[rule_format,of n],of "b n"]
+ unfolding pinfreal_minus_plus[OF less_imp_le[OF b[rule_format]]]
+ by(auto simp: add_commute)
+ hence "b n \<le> a - (a - Sup s) / 2" unfolding pinfreal_le_minus_iff
+ using asup by auto
+ hence "b n \<notin> {a - (a - Sup s) / 2<..}" by auto
+ thus False using n by auto
+ qed
+ then guess n .. note n = this
+ have "Sup s < a - (a - Sup s) / 2"
+ using False as om by (cases a) (auto simp: pinfreal_noteq_omega_Ex field_simps)
+ also have "... \<le> b n"
+ proof- note add_right_mono[OF less_imp_le[OF n],of "b n"]
+ note this[unfolded pinfreal_minus_plus[OF less_imp_le[OF b[rule_format]]]]
+ hence "a - (a - Sup s) / 2 \<le> (a - Sup s) / 2 + b n - (a - Sup s) / 2"
+ apply(rule pinfreal_minus_le_cancel_right) using asup by auto
+ also have "... = b n + (a - Sup s) / 2 - (a - Sup s) / 2"
+ by(auto simp add: add_commute)
+ also have "... = b n" apply(subst pinfreal_cancel_plus_minus)
+ proof(rule ccontr,unfold not_not) case goal1
+ show ?case using asup unfolding goal1 by auto
+ qed auto
+ finally show ?thesis .
+ qed
+ finally show False
+ using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]] by auto
+ next case True
+ from assms(2)[unfolded True Lim_omega_gt,rule_format,of "real (Sup s)"]
+ guess N .. note N = this[rule_format,of N]
+ thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of N]]
+ unfolding Real_real using om by auto
+ qed qed
+
+lemma less_SUP_iff:
+ fixes a :: pinfreal
+ shows "(a < (SUP i:A. f i)) \<longleftrightarrow> (\<exists>x\<in>A. a < f x)"
+ unfolding SUPR_def less_Sup_iff by auto
+
+lemma Sup_mono_lim:
+ assumes "\<forall>a\<in>A. \<exists>b. \<forall>n. b n \<in> B \<and> b ----> (a::pinfreal)"
+ shows "Sup A \<le> Sup B"
+ unfolding Sup_le_iff apply(rule) apply(drule assms[rule_format]) apply safe
+ apply(rule_tac b=b in Sup_lim) by auto
+
+lemma pinfreal_less_add:
+ assumes "x \<noteq> \<omega>" "a < b"
+ shows "x + a < x + b"
+ using assms by (cases a, cases b, cases x) auto
+
+lemma SUPR_lim:
+ assumes "\<forall>n. b n \<in> B" "(\<lambda>n. f (b n)) ----> (f a::pinfreal)"
+ shows "f a \<le> SUPR B f"
+ unfolding SUPR_def apply(rule Sup_lim[of "\<lambda>n. f (b n)"])
+ using assms by auto
+
+lemma SUP_\<omega>_imp:
+ assumes "(SUP i. f i) = \<omega>"
+ shows "\<exists>i. Real x < f i"
+proof (rule ccontr)
+ assume "\<not> ?thesis" hence "\<And>i. f i \<le> Real x" by (simp add: not_less)
+ hence "(SUP i. f i) \<le> Real x" unfolding SUP_le_iff by auto
+ with assms show False by auto
+qed
+
+lemma SUPR_mono_lim:
+ assumes "\<forall>a\<in>A. \<exists>b. \<forall>n. b n \<in> B \<and> (\<lambda>n. f (b n)) ----> (f a::pinfreal)"
+ shows "SUPR A f \<le> SUPR B f"
+ unfolding SUPR_def apply(rule Sup_mono_lim)
+ apply safe apply(drule assms[rule_format],safe)
+ apply(rule_tac x="\<lambda>n. f (b n)" in exI) by auto
+
+lemma real_0_imp_eq_0:
+ assumes "x \<noteq> \<omega>" "real x = 0"
+ shows "x = 0"
+using assms by (cases x) auto
+
+lemma SUPR_mono:
+ assumes "\<forall>a\<in>A. \<exists>b\<in>B. f b \<ge> f a"
+ shows "SUPR A f \<le> SUPR B f"
+ unfolding SUPR_def apply(rule Sup_mono)
+ using assms by auto
+
+lemma less_add_Real:
+ fixes x :: real
+ fixes a b :: pinfreal
+ assumes "x \<ge> 0" "a < b"
+ shows "a + Real x < b + Real x"
+using assms by (cases a, cases b) auto
+
+lemma le_add_Real:
+ fixes x :: real
+ fixes a b :: pinfreal
+ assumes "x \<ge> 0" "a \<le> b"
+ shows "a + Real x \<le> b + Real x"
+using assms by (cases a, cases b) auto
+
+lemma le_imp_less_pinfreal:
+ fixes x :: pinfreal
+ assumes "x > 0" "a + x \<le> b" "a \<noteq> \<omega>"
+ shows "a < b"
+using assms by (cases x, cases a, cases b) auto
+
+lemma pinfreal_INF_minus:
+ fixes f :: "nat \<Rightarrow> pinfreal"
+ assumes "c \<noteq> \<omega>"
+ shows "(INF i. c - f i) = c - (SUP i. f i)"
+proof (cases "SUP i. f i")
+ case infinite
+ from `c \<noteq> \<omega>` obtain x where [simp]: "c = Real x" by (cases c) auto
+ from SUP_\<omega>_imp[OF infinite] obtain i where "Real x < f i" by auto
+ have "(INF i. c - f i) \<le> c - f i"
+ by (auto intro!: complete_lattice_class.INF_leI)
+ also have "\<dots> = 0" using `Real x < f i` by (auto simp: minus_pinfreal_eq)
+ finally show ?thesis using infinite by auto
+next
+ case (preal r)
+ from `c \<noteq> \<omega>` obtain x where c: "c = Real x" by (cases c) auto
+
+ show ?thesis unfolding c
+ proof (rule pinfreal_INFI)
+ fix i have "f i \<le> (SUP i. f i)" by (rule le_SUPI) simp
+ thus "Real x - (SUP i. f i) \<le> Real x - f i" by (rule pinfreal_minus_le_cancel)
+ next
+ fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> y \<le> Real x - f i"
+ from this[of 0] obtain p where p: "y = Real p" "0 \<le> p"
+ by (cases "f 0", cases y, auto split: split_if_asm)
+ hence "\<And>i. Real p \<le> Real x - f i" using * by auto
+ hence *: "\<And>i. Real x \<le> f i \<Longrightarrow> Real p = 0"
+ "\<And>i. f i < Real x \<Longrightarrow> Real p + f i \<le> Real x"
+ unfolding pinfreal_le_minus_iff by auto
+ show "y \<le> Real x - (SUP i. f i)" unfolding p pinfreal_le_minus_iff
+ proof safe
+ assume x_less: "Real x \<le> (SUP i. f i)"
+ show "Real p = 0"
+ proof (rule ccontr)
+ assume "Real p \<noteq> 0"
+ hence "0 < Real p" by auto
+ from Sup_close[OF this, of "range f"]
+ obtain i where e: "(SUP i. f i) < f i + Real p"
+ using preal unfolding SUPR_def by auto
+ hence "Real x \<le> f i + Real p" using x_less by auto
+ show False
+ proof cases
+ assume "\<forall>i. f i < Real x"
+ hence "Real p + f i \<le> Real x" using * by auto
+ hence "f i + Real p \<le> (SUP i. f i)" using x_less by (auto simp: field_simps)
+ thus False using e by auto
+ next
+ assume "\<not> (\<forall>i. f i < Real x)"
+ then obtain i where "Real x \<le> f i" by (auto simp: not_less)
+ from *(1)[OF this] show False using `Real p \<noteq> 0` by auto
+ qed
+ qed
+ next
+ have "\<And>i. f i \<le> (SUP i. f i)" by (rule complete_lattice_class.le_SUPI) auto
+ also assume "(SUP i. f i) < Real x"
+ finally have "\<And>i. f i < Real x" by auto
+ hence *: "\<And>i. Real p + f i \<le> Real x" using * by auto
+ have "Real p \<le> Real x" using *[of 0] by (cases "f 0") (auto split: split_if_asm)
+
+ have SUP_eq: "(SUP i. f i) \<le> Real x - Real p"
+ proof (rule SUP_leI)
+ fix i show "f i \<le> Real x - Real p" unfolding pinfreal_le_minus_iff
+ proof safe
+ assume "Real x \<le> Real p"
+ with *[of i] show "f i = 0"
+ by (cases "f i") (auto split: split_if_asm)
+ next
+ assume "Real p < Real x"
+ show "f i + Real p \<le> Real x" using * by (auto simp: field_simps)
+ qed
+ qed
+
+ show "Real p + (SUP i. f i) \<le> Real x"
+ proof cases
+ assume "Real x \<le> Real p"
+ with `Real p \<le> Real x` have [simp]: "Real p = Real x" by (rule antisym)
+ { fix i have "f i = 0" using *[of i] by (cases "f i") (auto split: split_if_asm) }
+ hence "(SUP i. f i) \<le> 0" by (auto intro!: SUP_leI)
+ thus ?thesis by simp
+ next
+ assume "\<not> Real x \<le> Real p" hence "Real p < Real x" unfolding not_le .
+ with SUP_eq show ?thesis unfolding pinfreal_le_minus_iff by (auto simp: field_simps)
+ qed
+ qed
+ qed
+qed
+
+lemma pinfreal_SUP_minus:
+ fixes f :: "nat \<Rightarrow> pinfreal"
+ shows "(SUP i. c - f i) = c - (INF i. f i)"
+proof (rule pinfreal_SUPI)
+ fix i have "(INF i. f i) \<le> f i" by (rule INF_leI) simp
+ thus "c - f i \<le> c - (INF i. f i)" by (rule pinfreal_minus_le_cancel)
+next
+ fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c - f i \<le> y"
+ show "c - (INF i. f i) \<le> y"
+ proof (cases y)
+ case (preal p)
+
+ show ?thesis unfolding pinfreal_minus_le_iff preal
+ proof safe
+ assume INF_le_x: "(INF i. f i) \<le> c"
+ from * have *: "\<And>i. f i \<le> c \<Longrightarrow> c \<le> Real p + f i"
+ unfolding pinfreal_minus_le_iff preal by auto
+
+ have INF_eq: "c - Real p \<le> (INF i. f i)"
+ proof (rule le_INFI)
+ fix i show "c - Real p \<le> f i" unfolding pinfreal_minus_le_iff
+ proof safe
+ assume "Real p \<le> c"
+ show "c \<le> f i + Real p"
+ proof cases
+ assume "f i \<le> c" from *[OF this]
+ show ?thesis by (simp add: field_simps)
+ next
+ assume "\<not> f i \<le> c"
+ hence "c \<le> f i" by auto
+ also have "\<dots> \<le> f i + Real p" by auto
+ finally show ?thesis .
+ qed
+ qed
+ qed
+
+ show "c \<le> Real p + (INF i. f i)"
+ proof cases
+ assume "Real p \<le> c"
+ with INF_eq show ?thesis unfolding pinfreal_minus_le_iff by (auto simp: field_simps)
+ next
+ assume "\<not> Real p \<le> c"
+ hence "c \<le> Real p" by auto
+ also have "Real p \<le> Real p + (INF i. f i)" by auto
+ finally show ?thesis .
+ qed
+ qed
+ qed simp
+qed
+
+lemma pinfreal_le_minus_imp_0:
+ fixes a b :: pinfreal
+ shows "a \<le> a - b \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a \<noteq> \<omega> \<Longrightarrow> b = 0"
+ by (cases a, cases b, auto split: split_if_asm)
+
+lemma lim_INF_le_lim_SUP:
+ fixes f :: "nat \<Rightarrow> pinfreal"
+ shows "(SUP n. INF m. f (n + m)) \<le> (INF n. SUP m. f (n + m))"
+proof (rule complete_lattice_class.SUP_leI, rule complete_lattice_class.le_INFI)
+ fix i j show "(INF m. f (i + m)) \<le> (SUP m. f (j + m))"
+ proof (cases rule: le_cases)
+ assume "i \<le> j"
+ have "(INF m. f (i + m)) \<le> f (i + (j - i))" by (rule INF_leI) simp
+ also have "\<dots> = f (j + 0)" using `i \<le> j` by auto
+ also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
+ finally show ?thesis .
+ next
+ assume "j \<le> i"
+ have "(INF m. f (i + m)) \<le> f (i + 0)" by (rule INF_leI) simp
+ also have "\<dots> = f (j + (i - j))" using `j \<le> i` by auto
+ also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
+ finally show ?thesis .
+ qed
+qed
+
+lemma lim_INF_eq_lim_SUP:
+ fixes X :: "nat \<Rightarrow> real"
+ assumes "\<And>i. 0 \<le> X i" and "0 \<le> x"
+ and lim_INF: "(SUP n. INF m. Real (X (n + m))) = Real x" (is "(SUP n. ?INF n) = _")
+ and lim_SUP: "(INF n. SUP m. Real (X (n + m))) = Real x" (is "(INF n. ?SUP n) = _")
+ shows "X ----> x"
+proof (rule LIMSEQ_I)
+ fix r :: real assume "0 < r"
+ hence "0 \<le> r" by auto
+ from Sup_close[of "Real r" "range ?INF"]
+ obtain n where inf: "Real x < ?INF n + Real r"
+ unfolding SUPR_def lim_INF[unfolded SUPR_def] using `0 < r` by auto
+
+ from Inf_close[of "range ?SUP" "Real r"]
+ obtain n' where sup: "?SUP n' < Real x + Real r"
+ unfolding INFI_def lim_SUP[unfolded INFI_def] using `0 < r` by auto
+
+ show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
+ proof (safe intro!: exI[of _ "max n n'"])
+ fix m assume "max n n' \<le> m" hence "n \<le> m" "n' \<le> m" by auto
+
+ note inf
+ also have "?INF n + Real r \<le> Real (X (n + (m - n))) + Real r"
+ by (rule le_add_Real, auto simp: `0 \<le> r` intro: INF_leI)
+ finally have up: "x < X m + r"
+ using `0 \<le> X m` `0 \<le> x` `0 \<le> r` `n \<le> m` by auto
+
+ have "Real (X (n' + (m - n'))) \<le> ?SUP n'"
+ by (auto simp: `0 \<le> r` intro: le_SUPI)
+ also note sup
+ finally have down: "X m < x + r"
+ using `0 \<le> X m` `0 \<le> x` `0 \<le> r` `n' \<le> m` by auto
+
+ show "norm (X m - x) < r" using up down by auto
+ qed
+qed
+
+lemma Sup_countable_SUPR:
+ assumes "Sup A \<noteq> \<omega>" "A \<noteq> {}"
+ shows "\<exists> f::nat \<Rightarrow> pinfreal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
+proof -
+ have "\<And>n. 0 < 1 / (of_nat n :: pinfreal)" by auto
+ from Sup_close[OF this assms]
+ have "\<forall>n. \<exists>x. x \<in> A \<and> Sup A < x + 1 / of_nat n" by blast
+ from choice[OF this] obtain f where "range f \<subseteq> A" and
+ epsilon: "\<And>n. Sup A < f n + 1 / of_nat n" by blast
+ have "SUPR UNIV f = Sup A"
+ proof (rule pinfreal_SUPI)
+ fix i show "f i \<le> Sup A" using `range f \<subseteq> A`
+ by (auto intro!: complete_lattice_class.Sup_upper)
+ next
+ fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
+ show "Sup A \<le> y"
+ proof (rule pinfreal_le_epsilon)
+ fix e :: pinfreal assume "0 < e"
+ show "Sup A \<le> y + e"
+ proof (cases e)
+ case (preal r)
+ hence "0 < r" using `0 < e` by auto
+ then obtain n where *: "inverse (of_nat n) < r" "0 < n"
+ using ex_inverse_of_nat_less by auto
+ have "Sup A \<le> f n + 1 / of_nat n" using epsilon[of n] by auto
+ also have "1 / of_nat n \<le> e" using preal * by (auto simp: real_eq_of_nat)
+ with bound have "f n + 1 / of_nat n \<le> y + e" by (rule add_mono) simp
+ finally show "Sup A \<le> y + e" .
+ qed simp
+ qed
+ qed
+ with `range f \<subseteq> A` show ?thesis by (auto intro!: exI[of _ f])
+qed
+
+lemma SUPR_countable_SUPR:
+ assumes "SUPR A g \<noteq> \<omega>" "A \<noteq> {}"
+ shows "\<exists> f::nat \<Rightarrow> pinfreal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
+proof -
+ have "Sup (g`A) \<noteq> \<omega>" "g`A \<noteq> {}" using assms unfolding SUPR_def by auto
+ from Sup_countable_SUPR[OF this]
+ show ?thesis unfolding SUPR_def .
+qed
+
+lemma pinfreal_setsum_subtractf:
+ assumes "\<And>i. i\<in>A \<Longrightarrow> g i \<le> f i" and "\<And>i. i\<in>A \<Longrightarrow> f i \<noteq> \<omega>"
+ shows "(\<Sum>i\<in>A. f i - g i) = (\<Sum>i\<in>A. f i) - (\<Sum>i\<in>A. g i)"
+proof cases
+ assume "finite A" from this assms show ?thesis
+ proof induct
+ case (insert x A)
+ hence hyp: "(\<Sum>i\<in>A. f i - g i) = (\<Sum>i\<in>A. f i) - (\<Sum>i\<in>A. g i)"
+ by auto
+ { fix i assume *: "i \<in> insert x A"
+ hence "g i \<le> f i" using insert by simp
+ also have "f i < \<omega>" using * insert by (simp add: pinfreal_less_\<omega>)
+ finally have "g i \<noteq> \<omega>" by (simp add: pinfreal_less_\<omega>) }
+ hence "setsum g A \<noteq> \<omega>" "g x \<noteq> \<omega>" by (auto simp: setsum_\<omega>)
+ moreover have "setsum f A \<noteq> \<omega>" "f x \<noteq> \<omega>" using insert by (auto simp: setsum_\<omega>)
+ moreover have "g x \<le> f x" using insert by auto
+ moreover have "(\<Sum>i\<in>A. g i) \<le> (\<Sum>i\<in>A. f i)" using insert by (auto intro!: setsum_mono)
+ ultimately show ?case using `finite A` `x \<notin> A` hyp
+ by (auto simp: pinfreal_noteq_omega_Ex)
+ qed simp
+qed simp
+
+lemma real_of_pinfreal_diff:
+ "y \<le> x \<Longrightarrow> x \<noteq> \<omega> \<Longrightarrow> real x - real y = real (x - y)"
+ by (cases x, cases y) auto
+
+lemma psuminf_minus:
+ assumes ord: "\<And>i. g i \<le> f i" and fin: "psuminf g \<noteq> \<omega>" "psuminf f \<noteq> \<omega>"
+ shows "(\<Sum>\<^isub>\<infinity> i. f i - g i) = psuminf f - psuminf g"
+proof -
+ have [simp]: "\<And>i. f i \<noteq> \<omega>" using fin by (auto intro: psuminf_\<omega>)
+ from fin have "(\<lambda>x. real (f x)) sums real (\<Sum>\<^isub>\<infinity>x. f x)"
+ and "(\<lambda>x. real (g x)) sums real (\<Sum>\<^isub>\<infinity>x. g x)"
+ by (auto intro: psuminf_imp_suminf)
+ from sums_diff[OF this]
+ have "(\<lambda>n. real (f n - g n)) sums (real ((\<Sum>\<^isub>\<infinity>x. f x) - (\<Sum>\<^isub>\<infinity>x. g x)))" using fin ord
+ by (subst (asm) (1 2) real_of_pinfreal_diff) (auto simp: psuminf_\<omega> psuminf_le)
+ hence "(\<Sum>\<^isub>\<infinity> i. Real (real (f i - g i))) = Real (real ((\<Sum>\<^isub>\<infinity>x. f x) - (\<Sum>\<^isub>\<infinity>x. g x)))"
+ by (rule suminf_imp_psuminf) simp
+ thus ?thesis using fin by (simp add: Real_real psuminf_\<omega>)
+qed
+
+lemma INF_eq_LIMSEQ:
+ assumes "mono (\<lambda>i. - f i)" and "\<And>n. 0 \<le> f n" and "0 \<le> x"
+ shows "(INF n. Real (f n)) = Real x \<longleftrightarrow> f ----> x"
+proof
+ assume x: "(INF n. Real (f n)) = Real x"
+ { fix n
+ have "Real x \<le> Real (f n)" using x[symmetric] by (auto intro: INF_leI)
+ hence "x \<le> f n" using assms by simp
+ hence "\<bar>f n - x\<bar> = f n - x" by auto }
+ note abs_eq = this
+ show "f ----> x"
+ proof (rule LIMSEQ_I)
+ fix r :: real assume "0 < r"
+ show "\<exists>no. \<forall>n\<ge>no. norm (f n - x) < r"
+ proof (rule ccontr)
+ assume *: "\<not> ?thesis"
+ { fix N
+ from * obtain n where *: "N \<le> n" "r \<le> f n - x"
+ using abs_eq by (auto simp: not_less)
+ hence "x + r \<le> f n" by auto
+ also have "f n \<le> f N" using `mono (\<lambda>i. - f i)` * by (auto dest: monoD)
+ finally have "Real (x + r) \<le> Real (f N)" using `0 \<le> f N` by auto }
+ hence "Real x < Real (x + r)"
+ and "Real (x + r) \<le> (INF n. Real (f n))" using `0 < r` `0 \<le> x` by (auto intro: le_INFI)
+ hence "Real x < (INF n. Real (f n))" by (rule less_le_trans)
+ thus False using x by auto
+ qed
+ qed
+next
+ assume "f ----> x"
+ show "(INF n. Real (f n)) = Real x"
+ proof (rule pinfreal_INFI)
+ fix n
+ from decseq_le[OF _ `f ----> x`] assms
+ show "Real x \<le> Real (f n)" unfolding decseq_eq_incseq incseq_mono by auto
+ next
+ fix y assume *: "\<And>n. n\<in>UNIV \<Longrightarrow> y \<le> Real (f n)"
+ thus "y \<le> Real x"
+ proof (cases y)
+ case (preal r)
+ with * have "\<exists>N. \<forall>n\<ge>N. r \<le> f n" using assms by fastsimp
+ from LIMSEQ_le_const[OF `f ----> x` this]
+ show "y \<le> Real x" using `0 \<le> x` preal by auto
+ qed simp
+ qed
+qed
+
+lemma INFI_bound:
+ assumes "\<forall>N. x \<le> f N"
+ shows "x \<le> (INF n. f n)"
+ using assms by (simp add: INFI_def le_Inf_iff)
+
+lemma INF_mono:
+ assumes "\<And>n. f (N n) \<le> g n"
+ shows "(INF n. f n) \<le> (INF n. g n)"
+proof (safe intro!: INFI_bound)
+ fix n
+ have "(INF n. f n) \<le> f (N n)" by (auto intro!: INF_leI)
+ also note assms[of n]
+ finally show "(INF n. f n) \<le> g n" .
+qed
+
+lemma INFI_fun_expand: "(INF y:A. f y) = (\<lambda>x. INF y:A. f y x)"
+ unfolding INFI_def expand_fun_eq Inf_fun_def
+ by (auto intro!: arg_cong[where f=Inf])
+
+lemma LIMSEQ_imp_lim_INF:
+ assumes pos: "\<And>i. 0 \<le> X i" and lim: "X ----> x"
+ shows "(SUP n. INF m. Real (X (n + m))) = Real x"
+proof -
+ have "0 \<le> x" using assms by (auto intro!: LIMSEQ_le_const)
+
+ have "\<And>n. (INF m. Real (X (n + m))) \<le> Real (X (n + 0))" by (rule INF_leI) simp
+ also have "\<And>n. Real (X (n + 0)) < \<omega>" by simp
+ finally have "\<forall>n. \<exists>r\<ge>0. (INF m. Real (X (n + m))) = Real r"
+ by (auto simp: pinfreal_less_\<omega> pinfreal_noteq_omega_Ex)
+ from choice[OF this] obtain r where r: "\<And>n. (INF m. Real (X (n + m))) = Real (r n)" "\<And>n. 0 \<le> r n"
+ by auto
+
+ show ?thesis unfolding r
+ proof (subst SUP_eq_LIMSEQ)
+ show "mono r" unfolding mono_def
+ proof safe
+ fix x y :: nat assume "x \<le> y"
+ have "Real (r x) \<le> Real (r y)" unfolding r(1)[symmetric] using pos
+ proof (safe intro!: INF_mono)
+ fix m have "x + (m + y - x) = y + m"
+ using `x \<le> y` by auto
+ thus "Real (X (x + (m + y - x))) \<le> Real (X (y + m))" by simp
+ qed
+ thus "r x \<le> r y" using r by auto
+ qed
+ show "\<And>n. 0 \<le> r n" by fact
+ show "0 \<le> x" by fact
+ show "r ----> x"
+ proof (rule LIMSEQ_I)
+ fix e :: real assume "0 < e"
+ hence "0 < e/2" by auto
+ from LIMSEQ_D[OF lim this] obtain N where *: "\<And>n. N \<le> n \<Longrightarrow> \<bar>X n - x\<bar> < e/2"
+ by auto
+ show "\<exists>N. \<forall>n\<ge>N. norm (r n - x) < e"
+ proof (safe intro!: exI[of _ N])
+ fix n assume "N \<le> n"
+ show "norm (r n - x) < e"
+ proof cases
+ assume "r n < x"
+ have "x - r n \<le> e/2"
+ proof cases
+ assume e: "e/2 \<le> x"
+ have "Real (x - e/2) \<le> Real (r n)" unfolding r(1)[symmetric]
+ proof (rule le_INFI)
+ fix m show "Real (x - e / 2) \<le> Real (X (n + m))"
+ using *[of "n + m"] `N \<le> n`
+ using pos by (auto simp: field_simps abs_real_def split: split_if_asm)
+ qed
+ with e show ?thesis using pos `0 \<le> x` r(2) by auto
+ next
+ assume "\<not> e/2 \<le> x" hence "x - e/2 < 0" by auto
+ with `0 \<le> r n` show ?thesis by auto
+ qed
+ with `r n < x` show ?thesis by simp
+ next
+ assume e: "\<not> r n < x"
+ have "Real (r n) \<le> Real (X (n + 0))" unfolding r(1)[symmetric]
+ by (rule INF_leI) simp
+ hence "r n - x \<le> X n - x" using r pos by auto
+ also have "\<dots> < e/2" using *[OF `N \<le> n`] by (auto simp: field_simps abs_real_def split: split_if_asm)
+ finally have "r n - x < e" using `0 < e` by auto
+ with e show ?thesis by auto
+ qed
+ qed
+ qed
+ qed
+qed
+
+
+lemma real_of_pinfreal_strict_mono_iff:
+ "real a < real b \<longleftrightarrow> (b \<noteq> \<omega> \<and> ((a = \<omega> \<and> 0 < b) \<or> (a < b)))"
+proof (cases a)
+ case infinite thus ?thesis by (cases b) auto
+next
+ case preal thus ?thesis by (cases b) auto
+qed
+
+lemma real_of_pinfreal_mono_iff:
+ "real a \<le> real b \<longleftrightarrow> (a = \<omega> \<or> (b \<noteq> \<omega> \<and> a \<le> b) \<or> (b = \<omega> \<and> a = 0))"
+proof (cases a)
+ case infinite thus ?thesis by (cases b) auto
+next
+ case preal thus ?thesis by (cases b) auto
+qed
+
+lemma ex_pinfreal_inverse_of_nat_Suc_less:
+ fixes e :: pinfreal assumes "0 < e" shows "\<exists>n. inverse (of_nat (Suc n)) < e"
+proof (cases e)
+ case (preal r)
+ with `0 < e` ex_inverse_of_nat_Suc_less[of r]
+ obtain n where "inverse (of_nat (Suc n)) < r" by auto
+ with preal show ?thesis
+ by (auto simp: real_eq_of_nat[symmetric])
+qed auto
+
+lemma Lim_eq_Sup_mono:
+ fixes u :: "nat \<Rightarrow> pinfreal" assumes "mono u"
+ shows "u ----> (SUP i. u i)"
+proof -
+ from lim_pinfreal_increasing[of u] `mono u`
+ obtain l where l: "u ----> l" unfolding mono_def by auto
+ from SUP_Lim_pinfreal[OF _ this] `mono u`
+ have "(SUP i. u i) = l" unfolding mono_def by auto
+ with l show ?thesis by simp
+qed
+
+lemma isotone_Lim:
+ fixes x :: pinfreal assumes "u \<up> x"
+ shows "u ----> x" (is ?lim) and "mono u" (is ?mono)
+proof -
+ show ?mono using assms unfolding mono_iff_le_Suc isoton_def by auto
+ from Lim_eq_Sup_mono[OF this] `u \<up> x`
+ show ?lim unfolding isoton_def by simp
+qed
+
+lemma isoton_iff_Lim_mono:
+ fixes u :: "nat \<Rightarrow> pinfreal"
+ shows "u \<up> x \<longleftrightarrow> (mono u \<and> u ----> x)"
+proof safe
+ assume "mono u" and x: "u ----> x"
+ with SUP_Lim_pinfreal[OF _ x]
+ show "u \<up> x" unfolding isoton_def
+ using `mono u`[unfolded mono_def]
+ using `mono u`[unfolded mono_iff_le_Suc]
+ by auto
+qed (auto dest: isotone_Lim)
+
+lemma pinfreal_inverse_inverse[simp]:
+ fixes x :: pinfreal
+ shows "inverse (inverse x) = x"
+ by (cases x) auto
+
+lemma atLeastAtMost_omega_eq_atLeast:
+ shows "{a .. \<omega>} = {a ..}"
+by auto
+
+lemma atLeast0AtMost_eq_atMost: "{0 :: pinfreal .. a} = {.. a}" by auto
+
+lemma greaterThan_omega_Empty: "{\<omega> <..} = {}" by auto
+
+lemma lessThan_0_Empty: "{..< 0 :: pinfreal} = {}" by auto
+
+end
--- a/src/HOL/Probability/Probability_Space.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Probability/Probability_Space.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1,13 +1,52 @@
theory Probability_Space
-imports Lebesgue
+imports Lebesgue_Integration
begin
+lemma (in measure_space) measure_inter_full_set:
+ assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>"
+ assumes T: "\<mu> T = \<mu> (space M)"
+ shows "\<mu> (S \<inter> T) = \<mu> S"
+proof (rule antisym)
+ show " \<mu> (S \<inter> T) \<le> \<mu> S"
+ using assms by (auto intro!: measure_mono)
+
+ show "\<mu> S \<le> \<mu> (S \<inter> T)"
+ proof (rule ccontr)
+ assume contr: "\<not> ?thesis"
+ have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))"
+ unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"])
+ also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
+ using assms by (auto intro!: measure_subadditive)
+ also have "\<dots> < \<mu> (T - S) + \<mu> S"
+ by (rule pinfreal_less_add[OF not_\<omega>]) (insert contr, auto)
+ also have "\<dots> = \<mu> (T \<union> S)"
+ using assms by (subst measure_additive) auto
+ also have "\<dots> \<le> \<mu> (space M)"
+ using assms sets_into_space by (auto intro!: measure_mono)
+ finally show False ..
+ qed
+qed
+
+lemma (in finite_measure) finite_measure_inter_full_set:
+ assumes "S \<in> sets M" "T \<in> sets M"
+ assumes T: "\<mu> T = \<mu> (space M)"
+ shows "\<mu> (S \<inter> T) = \<mu> S"
+ using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms
+ by auto
+
locale prob_space = measure_space +
- assumes prob_space: "measure M (space M) = 1"
+ assumes measure_space_1: "\<mu> (space M) = 1"
+
+sublocale prob_space < finite_measure
+proof
+ from measure_space_1 show "\<mu> (space M) \<noteq> \<omega>" by simp
+qed
+
+context prob_space
begin
abbreviation "events \<equiv> sets M"
-abbreviation "prob \<equiv> measure M"
+abbreviation "prob \<equiv> \<lambda>A. real (\<mu> A)"
abbreviation "prob_preserving \<equiv> measure_preserving"
abbreviation "random_variable \<equiv> \<lambda> s X. X \<in> measurable M s"
abbreviation "expectation \<equiv> integral"
@@ -19,75 +58,50 @@
"indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)"
definition
- "distribution X = (\<lambda>s. prob ((X -` s) \<inter> (space M)))"
+ "distribution X = (\<lambda>s. \<mu> ((X -` s) \<inter> (space M)))"
abbreviation
"joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
-(*
-definition probably :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "\<forall>\<^sup>*" 10) where
- "probably P \<longleftrightarrow> { x. P x } \<in> events \<and> prob { x. P x } = 1"
-definition possibly :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "\<exists>\<^sup>*" 10) where
- "possibly P \<longleftrightarrow> { x. P x } \<in> events \<and> prob { x. P x } \<noteq> 0"
-*)
+lemma prob_space: "prob (space M) = 1"
+ unfolding measure_space_1 by simp
-definition
- "conditional_expectation X M' \<equiv> SOME f. f \<in> measurable M' borel_space \<and>
- (\<forall> g \<in> sets M'. measure_space.integral M' (\<lambda>x. f x * indicator_fn g x) =
- measure_space.integral M' (\<lambda>x. X x * indicator_fn g x))"
+lemma measure_le_1[simp, intro]:
+ assumes "A \<in> events" shows "\<mu> A \<le> 1"
+proof -
+ have "\<mu> A \<le> \<mu> (space M)"
+ using assms sets_into_space by(auto intro!: measure_mono)
+ also note measure_space_1
+ finally show ?thesis .
+qed
-definition
- "conditional_prob E M' \<equiv> conditional_expectation (indicator_fn E) M'"
-
-lemma positive': "positive M prob"
- unfolding positive_def using positive empty_measure by blast
+lemma measure_finite[simp, intro]:
+ assumes "A \<in> events" shows "\<mu> A \<noteq> \<omega>"
+ using measure_le_1[OF assms] by auto
lemma prob_compl:
- assumes "s \<in> events"
- shows "prob (space M - s) = 1 - prob s"
-using assms
-proof -
- have "prob ((space M - s) \<union> s) = prob (space M - s) + prob s"
- using assms additive[unfolded additive_def] by blast
- thus ?thesis by (simp add:Un_absorb2[OF sets_into_space[OF assms]] prob_space)
-qed
+ assumes "A \<in> events"
+ shows "prob (space M - A) = 1 - prob A"
+ using `A \<in> events`[THEN sets_into_space] `A \<in> events` measure_space_1
+ by (subst real_finite_measure_Diff) auto
lemma indep_space:
assumes "s \<in> events"
shows "indep (space M) s"
-using assms prob_space
-unfolding indep_def by auto
-
-
-lemma prob_space_increasing:
- "increasing M prob"
-by (rule additive_increasing[OF positive' additive])
+ using assms prob_space by (simp add: indep_def)
-lemma prob_subadditive:
- assumes "s \<in> events" "t \<in> events"
- shows "prob (s \<union> t) \<le> prob s + prob t"
-using assms
-proof -
- have "prob (s \<union> t) = prob ((s - t) \<union> t)" by simp
- also have "\<dots> = prob (s - t) + prob t"
- using additive[unfolded additive_def, rule_format, of "s-t" "t"]
- assms by blast
- also have "\<dots> \<le> prob s + prob t"
- using prob_space_increasing[unfolded increasing_def, rule_format] assms
- by auto
- finally show ?thesis by simp
-qed
+lemma prob_space_increasing: "increasing M prob"
+ by (auto intro!: real_measure_mono simp: increasing_def)
lemma prob_zero_union:
assumes "s \<in> events" "t \<in> events" "prob t = 0"
shows "prob (s \<union> t) = prob s"
-using assms
+using assms
proof -
have "prob (s \<union> t) \<le> prob s"
- using prob_subadditive[of s t] assms by auto
+ using real_finite_measure_subadditive[of s t] assms by auto
moreover have "prob (s \<union> t) \<ge> prob s"
- using prob_space_increasing[unfolded increasing_def, rule_format]
- assms by auto
+ using assms by (blast intro: real_measure_mono)
ultimately show ?thesis by simp
qed
@@ -95,18 +109,19 @@
assumes "s \<in> events" "t \<in> events"
assumes "prob (space M - s) = prob (space M - t)"
shows "prob s = prob t"
-using assms prob_compl by auto
+ using assms prob_compl by auto
lemma prob_one_inter:
assumes events:"s \<in> events" "t \<in> events"
assumes "prob t = 1"
shows "prob (s \<inter> t) = prob s"
-using assms
proof -
- have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)"
- using prob_compl[of "t"] prob_zero_union assms by auto
- then show "prob (s \<inter> t) = prob s"
- using prob_eq_compl[of "s \<inter> t"] events by (simp only: Diff_Int) auto
+ have "prob ((space M - s) \<union> (space M - t)) = prob (space M - s)"
+ using events assms prob_compl[of "t"] by (auto intro!: prob_zero_union)
+ also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
+ by blast
+ finally show "prob (s \<inter> t) = prob s"
+ using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s])
qed
lemma prob_eq_bigunion_image:
@@ -114,98 +129,24 @@
assumes "disjoint_family f" "disjoint_family g"
assumes "\<And> n :: nat. prob (f n) = prob (g n)"
shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))"
-using assms
-proof -
- have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))"
- using ca[unfolded countably_additive_def] assms by blast
- have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))"
- using ca[unfolded countably_additive_def] assms by blast
- show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
-qed
-
-lemma prob_countably_subadditive:
- assumes "range f \<subseteq> events"
- assumes "summable (prob \<circ> f)"
- shows "prob (\<Union>i. f i) \<le> (\<Sum> i. prob (f i))"
using assms
proof -
- def f' == "\<lambda> i. f i - (\<Union> j \<in> {0 ..< i}. f j)"
- have "(\<Union> i. f' i) \<subseteq> (\<Union> i. f i)" unfolding f'_def by auto
- moreover have "(\<Union> i. f' i) \<supseteq> (\<Union> i. f i)"
- proof (rule subsetI)
- fix x assume "x \<in> (\<Union> i. f i)"
- then obtain k where "x \<in> f k" by blast
- hence k: "k \<in> {m. x \<in> f m}" by simp
- have "\<exists> l. x \<in> f l \<and> (\<forall> l' < l. x \<notin> f l')"
- using wfE_min[of "{(x, y). x < y}" "k" "{m. x \<in> f m}",
- OF wf_less k] by auto
- thus "x \<in> (\<Union> i. f' i)" unfolding f'_def by auto
- qed
- ultimately have uf'f: "(\<Union> i. f' i) = (\<Union> i. f i)" by (rule equalityI)
-
- have df': "\<And> i j. i < j \<Longrightarrow> f' i \<inter> f' j = {}"
- unfolding f'_def by auto
- have "\<And> i j. i \<noteq> j \<Longrightarrow> f' i \<inter> f' j = {}"
- apply (drule iffD1[OF nat_neq_iff])
- using df' by auto
- hence df: "disjoint_family f'" unfolding disjoint_family_on_def by simp
-
- have rf': "\<And> i. f' i \<in> events"
- proof -
- fix i :: nat
- have "(\<Union> {f j | j. j \<in> {0 ..< i}}) = (\<Union> j \<in> {0 ..< i}. f j)" by blast
- hence "(\<Union> {f j | j. j \<in> {0 ..< i}}) \<in> events
- \<Longrightarrow> (\<Union> j \<in> {0 ..< i}. f j) \<in> events" by auto
- thus "f' i \<in> events"
- unfolding f'_def
- using assms finite_union[of "{f j | j. j \<in> {0 ..< i}}"]
- Diff[of "f i" "\<Union> j \<in> {0 ..< i}. f j"] by auto
- qed
- hence uf': "(\<Union> range f') \<in> events" by auto
-
- have "\<And> i. prob (f' i) \<le> prob (f i)"
- using prob_space_increasing[unfolded increasing_def, rule_format, OF rf']
- assms rf' unfolding f'_def by blast
-
- hence absinc: "\<And> i. \<bar> prob (f' i) \<bar> \<le> prob (f i)"
- using abs_of_nonneg positive'[unfolded positive_def]
- assms rf' by auto
-
- have "prob (\<Union> i. f i) = prob (\<Union> i. f' i)" using uf'f by simp
-
- also have "\<dots> = (\<Sum> i. prob (f' i))"
- using ca[unfolded countably_additive_def, rule_format]
- sums_unique rf' uf' df
- by auto
-
- also have "\<dots> \<le> (\<Sum> i. prob (f i))"
- using summable_le2[of "\<lambda> i. prob (f' i)" "\<lambda> i. prob (f i)",
- rule_format, OF absinc]
- assms[unfolded o_def] by auto
-
- finally show ?thesis by auto
+ have a: "(\<lambda> i. prob (f i)) sums (prob (\<Union> i. f i))"
+ by (rule real_finite_measure_UNION[OF assms(1,3)])
+ have b: "(\<lambda> i. prob (g i)) sums (prob (\<Union> i. g i))"
+ by (rule real_finite_measure_UNION[OF assms(2,4)])
+ show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
qed
lemma prob_countably_zero:
assumes "range c \<subseteq> events"
assumes "\<And> i. prob (c i) = 0"
- shows "(prob (\<Union> i :: nat. c i) = 0)"
- using assms
-proof -
- have leq0: "0 \<le> prob (\<Union> i. c i)"
- using assms positive'[unfolded positive_def, rule_format]
- by auto
-
- have "prob (\<Union> i. c i) \<le> (\<Sum> i. prob (c i))"
- using prob_countably_subadditive[of c, unfolded o_def]
- assms sums_zero sums_summable by auto
-
- also have "\<dots> = 0"
- using assms sums_zero
- sums_unique[of "\<lambda> i. prob (c i)" "0"] by auto
-
- finally show "prob (\<Union> i. c i) = 0"
- using leq0 by auto
+ shows "prob (\<Union> i :: nat. c i) = 0"
+proof (rule antisym)
+ show "prob (\<Union> i :: nat. c i) \<le> 0"
+ using real_finite_measurable_countably_subadditive[OF assms(1)]
+ by (simp add: assms(2) suminf_zero summable_zero)
+ show "0 \<le> prob (\<Union> i :: nat. c i)" by (rule real_pinfreal_nonneg)
qed
lemma indep_sym:
@@ -218,191 +159,192 @@
using assms unfolding indep_def by auto
lemma prob_equiprobable_finite_unions:
- assumes "s \<in> events"
- assumes "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
- assumes "finite s"
+ assumes "s \<in> events"
+ assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})"
- shows "prob s = of_nat (card s) * prob {SOME x. x \<in> s}"
-using assms
+ shows "prob s = real (card s) * prob {SOME x. x \<in> s}"
proof (cases "s = {}")
- case True thus ?thesis by simp
-next
- case False hence " \<exists> x. x \<in> s" by blast
+ case False hence "\<exists> x. x \<in> s" by blast
from someI_ex[OF this] assms
have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
have "prob s = (\<Sum> x \<in> s. prob {x})"
- using assms measure_real_sum_image by blast
+ using real_finite_measure_finite_singelton[OF s_finite] by simp
also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
- also have "\<dots> = of_nat (card s) * prob {(SOME x. x \<in> s)}"
- using setsum_constant assms by auto
+ also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
+ using setsum_constant assms by (simp add: real_eq_of_nat)
finally show ?thesis by simp
-qed
+qed simp
lemma prob_real_sum_image_fn:
assumes "e \<in> events"
assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events"
assumes "finite s"
- assumes "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
- assumes "space M \<subseteq> (\<Union> i \<in> s. f i)"
+ assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
+ assumes upper: "space M \<subseteq> (\<Union> i \<in> s. f i)"
shows "prob e = (\<Sum> x \<in> s. prob (e \<inter> f x))"
-using assms
proof -
- let ?S = "{0 ..< card s}"
- obtain g where "g ` ?S = s \<and> inj_on g ?S"
- using ex_bij_betw_nat_finite[unfolded bij_betw_def, of s] assms by auto
- moreover hence gs: "g ` ?S = s" by simp
- ultimately have ginj: "inj_on g ?S" by simp
- let ?f' = "\<lambda> i. e \<inter> f (g i)"
- have f': "?f' \<in> ?S \<rightarrow> events"
- using gs assms by blast
- hence "\<And> i j. \<lbrakk>i \<in> ?S ; j \<in> ?S ; i \<noteq> j\<rbrakk>
- \<Longrightarrow> ?f' i \<inter> ?f' j = {}" using assms ginj[unfolded inj_on_def] gs f' by blast
- hence df': "\<And> i j. \<lbrakk>i < card s ; j < card s ; i \<noteq> j\<rbrakk>
- \<Longrightarrow> ?f' i \<inter> ?f' j = {}" by simp
-
- have "e = e \<inter> space M" using assms sets_into_space by simp
- also hence "\<dots> = e \<inter> (\<Union> x \<in> s. f x)" using assms by blast
- also have "\<dots> = (\<Union> x \<in> g ` ?S. e \<inter> f x)" using gs by simp
- also have "\<dots> = (\<Union> i \<in> ?S. ?f' i)" by simp
- finally have "prob e = prob (\<Union> i \<in> ?S. ?f' i)" by simp
- also have "\<dots> = (\<Sum> i \<in> ?S. prob (?f' i))"
- apply (subst measure_finitely_additive'')
- using f' df' assms by (auto simp: disjoint_family_on_def)
- also have "\<dots> = (\<Sum> x \<in> g ` ?S. prob (e \<inter> f x))"
- using setsum_reindex[of g "?S" "\<lambda> x. prob (e \<inter> f x)"]
- ginj by simp
- also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))" using gs by simp
- finally show ?thesis by simp
+ have e: "e = (\<Union> i \<in> s. e \<inter> f i)"
+ using `e \<in> events` sets_into_space upper by blast
+ hence "prob e = prob (\<Union> i \<in> s. e \<inter> f i)" by simp
+ also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))"
+ proof (rule real_finite_measure_finite_Union)
+ show "finite s" by fact
+ show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact
+ show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
+ using disjoint by (auto simp: disjoint_family_on_def)
+ qed
+ finally show ?thesis .
qed
lemma distribution_prob_space:
- assumes "random_variable s X"
- shows "prob_space \<lparr>space = space s, sets = sets s, measure = distribution X\<rparr>"
-using assms
+ fixes S :: "('c, 'd) algebra_scheme"
+ assumes "sigma_algebra S" "random_variable S X"
+ shows "prob_space S (distribution X)"
proof -
- let ?N = "\<lparr>space = space s, sets = sets s, measure = distribution X\<rparr>"
- interpret s: sigma_algebra "s" using assms[unfolded measurable_def] by auto
- hence sigN: "sigma_algebra ?N" using s.sigma_algebra_extend by auto
-
- have pos: "\<And> e. e \<in> sets s \<Longrightarrow> distribution X e \<ge> 0"
- unfolding distribution_def
- using positive'[unfolded positive_def]
- assms[unfolded measurable_def] by auto
+ interpret S: sigma_algebra S by fact
+ show ?thesis
+ proof
+ show "distribution X {} = 0" unfolding distribution_def by simp
+ have "X -` space S \<inter> space M = space M"
+ using `random_variable S X` by (auto simp: measurable_def)
+ then show "distribution X (space S) = 1" using measure_space_1 by (simp add: distribution_def)
- have cas: "countably_additive ?N (distribution X)"
- proof -
- {
- fix f :: "nat \<Rightarrow> 'c \<Rightarrow> bool"
- let ?g = "\<lambda> n. X -` f n \<inter> space M"
- assume asm: "range f \<subseteq> sets s" "UNION UNIV f \<in> sets s" "disjoint_family f"
- hence "range ?g \<subseteq> events"
- using assms unfolding measurable_def by blast
- from ca[unfolded countably_additive_def,
- rule_format, of ?g, OF this] countable_UN[OF this] asm
- have "(\<lambda> n. prob (?g n)) sums prob (UNION UNIV ?g)"
- unfolding disjoint_family_on_def by blast
- moreover have "(X -` (\<Union> n. f n)) = (\<Union> n. X -` f n)" by blast
- ultimately have "(\<lambda> n. distribution X (f n)) sums distribution X (UNION UNIV f)"
- unfolding distribution_def by simp
- } thus ?thesis unfolding countably_additive_def by simp
+ show "countably_additive S (distribution X)"
+ proof (unfold countably_additive_def, safe)
+ fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets S" "disjoint_family A"
+ hence *: "\<And>i. X -` A i \<inter> space M \<in> sets M"
+ using `random_variable S X` by (auto simp: measurable_def)
+ moreover hence "\<And>i. \<mu> (X -` A i \<inter> space M) \<noteq> \<omega>"
+ using finite_measure by auto
+ moreover have "(\<Union>i. X -` A i \<inter> space M) \<in> sets M"
+ using * by blast
+ moreover hence "\<mu> (\<Union>i. X -` A i \<inter> space M) \<noteq> \<omega>"
+ using finite_measure by auto
+ moreover have **: "disjoint_family (\<lambda>i. X -` A i \<inter> space M)"
+ using `disjoint_family A` by (auto simp: disjoint_family_on_def)
+ ultimately show "(\<Sum>\<^isub>\<infinity> i. distribution X (A i)) = distribution X (\<Union>i. A i)"
+ using measure_countably_additive[OF _ **]
+ by (auto simp: distribution_def Real_real comp_def vimage_UN)
+ qed
qed
-
- have ds0: "distribution X {} = 0"
- unfolding distribution_def by simp
-
- have "X -` space s \<inter> space M = space M"
- using assms[unfolded measurable_def] by auto
- hence ds1: "distribution X (space s) = 1"
- unfolding measurable_def distribution_def using prob_space by simp
-
- from ds0 ds1 cas pos sigN
- show "prob_space ?N"
- unfolding prob_space_def prob_space_axioms_def
- measure_space_def measure_space_axioms_def by simp
qed
lemma distribution_lebesgue_thm1:
assumes "random_variable s X"
assumes "A \<in> sets s"
- shows "distribution X A = expectation (indicator_fn (X -` A \<inter> space M))"
+ shows "real (distribution X A) = expectation (indicator (X -` A \<inter> space M))"
unfolding distribution_def
using assms unfolding measurable_def
-using integral_indicator_fn by auto
+using integral_indicator by auto
lemma distribution_lebesgue_thm2:
- assumes "random_variable s X" "A \<in> sets s"
- shows "distribution X A = measure_space.integral \<lparr>space = space s, sets = sets s, measure = distribution X\<rparr> (indicator_fn A)"
- (is "_ = measure_space.integral ?M _")
+ assumes "sigma_algebra S" "random_variable S X" and "A \<in> sets S"
+ shows "distribution X A =
+ measure_space.positive_integral S (distribution X) (indicator A)"
+ (is "_ = measure_space.positive_integral _ ?D _")
proof -
- interpret S: prob_space ?M using assms(1) by (rule distribution_prob_space)
+ interpret S: prob_space S "distribution X" using assms(1,2) by (rule distribution_prob_space)
show ?thesis
- using S.integral_indicator_fn(1)
+ using S.positive_integral_indicator(1)
using assms unfolding distribution_def by auto
qed
lemma finite_expectation1:
- assumes "finite (space M)" "random_variable borel_space X"
+ assumes "finite (X`space M)" and rv: "random_variable borel_space X"
shows "expectation X = (\<Sum> r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))"
- using assms integral_finite measurable_def
- unfolding borel_measurable_def by auto
+proof (rule integral_on_finite(2)[OF assms(2,1)])
+ fix x have "X -` {x} \<inter> space M \<in> sets M"
+ using rv unfolding measurable_def by auto
+ thus "\<mu> (X -` {x} \<inter> space M) \<noteq> \<omega>" using finite_measure by simp
+qed
lemma finite_expectation:
- assumes "finite (space M) \<and> random_variable borel_space X"
- shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
-using assms unfolding distribution_def using finite_expectation1 by auto
+ assumes "finite (space M)" "random_variable borel_space X"
+ shows "expectation X = (\<Sum> r \<in> X ` (space M). r * real (distribution X {r}))"
+ using assms unfolding distribution_def using finite_expectation1 by auto
+
lemma prob_x_eq_1_imp_prob_y_eq_0:
assumes "{x} \<in> events"
- assumes "(prob {x} = 1)"
+ assumes "prob {x} = 1"
assumes "{y} \<in> events"
assumes "y \<noteq> x"
shows "prob {y} = 0"
using prob_one_inter[of "{y}" "{x}"] assms by auto
+lemma distribution_empty[simp]: "distribution X {} = 0"
+ unfolding distribution_def by simp
+
+lemma distribution_space[simp]: "distribution X (X ` space M) = 1"
+proof -
+ have "X -` X ` space M \<inter> space M = space M" by auto
+ thus ?thesis unfolding distribution_def by (simp add: measure_space_1)
+qed
+
+lemma distribution_one:
+ assumes "random_variable M X" and "A \<in> events"
+ shows "distribution X A \<le> 1"
+proof -
+ have "distribution X A \<le> \<mu> (space M)" unfolding distribution_def
+ using assms[unfolded measurable_def] by (auto intro!: measure_mono)
+ thus ?thesis by (simp add: measure_space_1)
+qed
+
+lemma distribution_finite:
+ assumes "random_variable M X" and "A \<in> events"
+ shows "distribution X A \<noteq> \<omega>"
+ using distribution_one[OF assms] by auto
+
lemma distribution_x_eq_1_imp_distribution_y_eq_0:
assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X"
- assumes "(distribution X {x} = 1)"
+ (is "random_variable ?S X")
+ assumes "distribution X {x} = 1"
assumes "y \<noteq> x"
shows "distribution X {y} = 0"
proof -
- let ?S = "\<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr>"
- let ?M = "\<lparr>space = X ` (space M), sets = Pow (X ` (space M)), measure = distribution X\<rparr>"
- interpret S: prob_space ?M
- using distribution_prob_space[OF X] by auto
- { assume "{x} \<notin> sets ?M"
- hence "x \<notin> X ` space M" by auto
+ have "sigma_algebra ?S" by (rule sigma_algebra_Pow)
+ from distribution_prob_space[OF this X]
+ interpret S: prob_space ?S "distribution X" by simp
+
+ have x: "{x} \<in> sets ?S"
+ proof (rule ccontr)
+ assume "{x} \<notin> sets ?S"
hence "X -` {x} \<inter> space M = {}" by auto
- hence "distribution X {x} = 0" unfolding distribution_def by auto
- hence "False" using assms by auto }
- hence x: "{x} \<in> sets ?M" by auto
- { assume "{y} \<notin> sets ?M"
- hence "y \<notin> X ` space M" by auto
+ thus "False" using assms unfolding distribution_def by auto
+ qed
+
+ have [simp]: "{y} \<inter> {x} = {}" "{x} - {y} = {x}" using `y \<noteq> x` by auto
+
+ show ?thesis
+ proof cases
+ assume "{y} \<in> sets ?S"
+ with `{x} \<in> sets ?S` assms show "distribution X {y} = 0"
+ using S.measure_inter_full_set[of "{y}" "{x}"]
+ by simp
+ next
+ assume "{y} \<notin> sets ?S"
hence "X -` {y} \<inter> space M = {}" by auto
- hence "distribution X {y} = 0" unfolding distribution_def by auto }
- moreover
- { assume "{y} \<in> sets ?M"
- hence "distribution X {y} = 0" using assms S.prob_x_eq_1_imp_prob_y_eq_0[OF x] by auto }
- ultimately show ?thesis by auto
+ thus "distribution X {y} = 0" unfolding distribution_def by auto
+ qed
qed
-
end
locale finite_prob_space = prob_space + finite_measure_space
lemma finite_prob_space_eq:
- "finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1"
+ "finite_prob_space M \<mu> \<longleftrightarrow> finite_measure_space M \<mu> \<and> \<mu> (space M) = 1"
unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def
by auto
lemma (in prob_space) not_empty: "space M \<noteq> {}"
using prob_space empty_measure by auto
-lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. measure M {x}) = 1"
- using prob_space sum_over_space by simp
+lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
+ using measure_space_1 sum_over_space by simp
lemma (in finite_prob_space) positive_distribution: "0 \<le> distribution X x"
- unfolding distribution_def using positive sets_eq_Pow by simp
+ unfolding distribution_def by simp
lemma (in finite_prob_space) joint_distribution_restriction_fst:
"joint_distribution X Y A \<le> distribution X (fst ` A)"
@@ -439,24 +381,27 @@
lemma (in finite_prob_space) finite_product_measure_space:
assumes "finite s1" "finite s2"
- shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = joint_distribution X Y\<rparr>"
- (is "finite_measure_space ?M")
+ shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2)\<rparr> (joint_distribution X Y)"
+ (is "finite_measure_space ?M ?D")
proof (rule finite_Pow_additivity_sufficient)
- show "positive ?M (measure ?M)"
- unfolding positive_def using positive'[unfolded positive_def] assms sets_eq_Pow
+ show "positive ?D"
+ unfolding positive_def using assms sets_eq_Pow
by (simp add: distribution_def)
- show "additive ?M (measure ?M)" unfolding additive_def
+ show "additive ?M ?D" unfolding additive_def
proof safe
fix x y
have A: "((\<lambda>x. (X x, Y x)) -` x) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto
have B: "((\<lambda>x. (X x, Y x)) -` y) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto
assume "x \<inter> y = {}"
+ hence "(\<lambda>x. (X x, Y x)) -` x \<inter> space M \<inter> ((\<lambda>x. (X x, Y x)) -` y \<inter> space M) = {}"
+ by auto
from additive[unfolded additive_def, rule_format, OF A B] this
- show "measure ?M (x \<union> y) = measure ?M x + measure ?M y"
+ finite_measure[OF A] finite_measure[OF B]
+ show "?D (x \<union> y) = ?D x + ?D y"
apply (simp add: distribution_def)
apply (subst Int_Un_distrib2)
- by auto
+ by (auto simp: real_of_pinfreal_add)
qed
show "finite (space ?M)"
@@ -464,23 +409,25 @@
show "sets ?M = Pow (space ?M)"
by simp
+
+ { fix x assume "x \<in> space ?M" thus "?D {x} \<noteq> \<omega>"
+ unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) }
qed
lemma (in finite_prob_space) finite_product_measure_space_of_images:
shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
- sets = Pow (X ` space M \<times> Y ` space M),
- measure = joint_distribution X Y\<rparr>"
- (is "finite_measure_space ?M")
+ sets = Pow (X ` space M \<times> Y ` space M) \<rparr>
+ (joint_distribution X Y)"
using finite_space by (auto intro!: finite_product_measure_space)
lemma (in finite_prob_space) finite_measure_space:
- shows "finite_measure_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = distribution X\<rparr>"
- (is "finite_measure_space ?S")
+ shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
+ (is "finite_measure_space ?S _")
proof (rule finite_Pow_additivity_sufficient, simp_all)
show "finite (X ` space M)" using finite_space by simp
- show "positive ?S (distribution X)" unfolding distribution_def
- unfolding positive_def using positive'[unfolded positive_def] sets_eq_Pow by auto
+ show "positive (distribution X)"
+ unfolding distribution_def positive_def using sets_eq_Pow by auto
show "additive ?S (distribution X)" unfolding additive_def distribution_def
proof (simp, safe)
@@ -488,36 +435,32 @@
have x: "(X -` x) \<inter> space M \<in> sets M"
and y: "(X -` y) \<inter> space M \<in> sets M" using sets_eq_Pow by auto
assume "x \<inter> y = {}"
+ hence "X -` x \<inter> space M \<inter> (X -` y \<inter> space M) = {}" by auto
from additive[unfolded additive_def, rule_format, OF x y] this
- have "prob (((X -` x) \<union> (X -` y)) \<inter> space M) =
- prob ((X -` x) \<inter> space M) + prob ((X -` y) \<inter> space M)"
- apply (subst Int_Un_distrib2)
- by auto
- thus "prob ((X -` x \<union> X -` y) \<inter> space M) = prob (X -` x \<inter> space M) + prob (X -` y \<inter> space M)"
+ finite_measure[OF x] finite_measure[OF y]
+ have "\<mu> (((X -` x) \<union> (X -` y)) \<inter> space M) =
+ \<mu> ((X -` x) \<inter> space M) + \<mu> ((X -` y) \<inter> space M)"
+ by (subst Int_Un_distrib2) auto
+ thus "\<mu> ((X -` x \<union> X -` y) \<inter> space M) = \<mu> (X -` x \<inter> space M) + \<mu> (X -` y \<inter> space M)"
by auto
qed
+
+ { fix x assume "x \<in> X ` space M" thus "distribution X {x} \<noteq> \<omega>"
+ unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) }
qed
lemma (in finite_prob_space) finite_prob_space_of_images:
- "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = distribution X\<rparr>"
- (is "finite_prob_space ?S")
-proof (simp add: finite_prob_space_eq, safe)
- show "finite_measure_space ?S" by (rule finite_measure_space)
- have "X -` X ` space M \<inter> space M = space M" by auto
- thus "distribution X (X`space M) = 1"
- by (simp add: distribution_def prob_space)
-qed
+ "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
+ by (simp add: finite_prob_space_eq finite_measure_space)
lemma (in finite_prob_space) finite_product_prob_space_of_images:
- "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M),
- measure = joint_distribution X Y\<rparr>"
- (is "finite_prob_space ?S")
-proof (simp add: finite_prob_space_eq, safe)
- show "finite_measure_space ?S" by (rule finite_product_measure_space_of_images)
-
+ "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr>
+ (joint_distribution X Y)"
+ (is "finite_prob_space ?S _")
+proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images)
have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
- by (simp add: distribution_def prob_space vimage_Times comp_def)
+ by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
qed
end
--- a/src/HOL/Probability/Product_Measure.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Probability/Product_Measure.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1,97 +1,172 @@
theory Product_Measure
-imports Lebesgue
+imports Lebesgue_Integration
begin
+definition prod_sets where
+ "prod_sets A B = {z. \<exists>x \<in> A. \<exists>y \<in> B. z = x \<times> y}"
+
definition
- "prod_measure M M' = (\<lambda>a. measure_space.integral M (\<lambda>s0. measure M' ((\<lambda>s1. (s0, s1)) -` a)))"
+ "prod_measure M \<mu> N \<nu> = (\<lambda>A. measure_space.positive_integral M \<mu> (\<lambda>s0. \<nu> ((\<lambda>s1. (s0, s1)) -` A)))"
definition
- "prod_measure_space M M' \<equiv>
- \<lparr> space = space M \<times> space M',
- sets = sets (sigma (space M \<times> space M') (prod_sets (sets M) (sets M'))),
- measure = prod_measure M M' \<rparr>"
+ "prod_measure_space M1 M2 = sigma (space M1 \<times> space M2) (prod_sets (sets M1) (sets M2))"
+
+lemma prod_setsI: "x \<in> A \<Longrightarrow> y \<in> B \<Longrightarrow> (x \<times> y) \<in> prod_sets A B"
+ by (auto simp add: prod_sets_def)
-lemma prod_measure_times:
- assumes "measure_space M" and "measure_space M'" and a: "a \<in> sets M"
- shows "prod_measure M M' (a \<times> a') = measure M a * measure M' a'"
-proof -
- interpret M: measure_space M by fact
- interpret M': measure_space M' by fact
+lemma sigma_prod_sets_finite:
+ assumes "finite A" and "finite B"
+ shows "sigma_sets (A \<times> B) (prod_sets (Pow A) (Pow B)) = Pow (A \<times> B)"
+proof safe
+ have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
- { fix \<omega>
- have "(\<lambda>\<omega>'. (\<omega>, \<omega>')) -` (a \<times> a') = (if \<omega> \<in> a then a' else {})"
- by auto
- hence "measure M' ((\<lambda>\<omega>'. (\<omega>, \<omega>')) -` (a \<times> a')) =
- measure M' a' * indicator_fn a \<omega>"
- unfolding indicator_fn_def by auto }
- note vimage_eq_indicator = this
-
- show ?thesis
- unfolding prod_measure_def vimage_eq_indicator
- M.integral_cmul_indicator(1)[OF `a \<in> sets M`]
- by simp
+ fix x assume subset: "x \<subseteq> A \<times> B"
+ hence "finite x" using fin by (rule finite_subset)
+ from this subset show "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))"
+ (is "x \<in> sigma_sets ?prod ?sets")
+ proof (induct x)
+ case empty show ?case by (rule sigma_sets.Empty)
+ next
+ case (insert a x)
+ hence "{a} \<in> sigma_sets ?prod ?sets" by (auto simp: prod_sets_def intro!: sigma_sets.Basic)
+ moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
+ ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
+ qed
+next
+ fix x a b
+ assume "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))" and "(a, b) \<in> x"
+ from sigma_sets_into_sp[OF _ this(1)] this(2)
+ show "a \<in> A" and "b \<in> B"
+ by (auto simp: prod_sets_def)
qed
-lemma finite_prod_measure_space:
- assumes "finite_measure_space M" and "finite_measure_space M'"
- shows "prod_measure_space M M' =
- \<lparr> space = space M \<times> space M',
- sets = Pow (space M \<times> space M'),
- measure = prod_measure M M' \<rparr>"
+lemma (in sigma_algebra) measurable_prod_sigma:
+ assumes sa1: "sigma_algebra a1" and sa2: "sigma_algebra a2"
+ assumes 1: "(fst o f) \<in> measurable M a1" and 2: "(snd o f) \<in> measurable M a2"
+ shows "f \<in> measurable M (sigma ((space a1) \<times> (space a2))
+ (prod_sets (sets a1) (sets a2)))"
proof -
- interpret M: finite_measure_space M by fact
- interpret M': finite_measure_space M' by fact
- show ?thesis using M.finite_space M'.finite_space
- by (simp add: sigma_prod_sets_finite M.sets_eq_Pow M'.sets_eq_Pow
- prod_measure_space_def)
+ from 1 have fn1: "fst \<circ> f \<in> space M \<rightarrow> space a1"
+ and q1: "\<forall>y\<in>sets a1. (fst \<circ> f) -` y \<inter> space M \<in> sets M"
+ by (auto simp add: measurable_def)
+ from 2 have fn2: "snd \<circ> f \<in> space M \<rightarrow> space a2"
+ and q2: "\<forall>y\<in>sets a2. (snd \<circ> f) -` y \<inter> space M \<in> sets M"
+ by (auto simp add: measurable_def)
+ show ?thesis
+ proof (rule measurable_sigma)
+ show "prod_sets (sets a1) (sets a2) \<subseteq> Pow (space a1 \<times> space a2)" using sa1 sa2
+ by (auto simp add: prod_sets_def sigma_algebra_iff dest: algebra.space_closed)
+ next
+ show "f \<in> space M \<rightarrow> space a1 \<times> space a2"
+ by (rule prod_final [OF fn1 fn2])
+ next
+ fix z
+ assume z: "z \<in> prod_sets (sets a1) (sets a2)"
+ thus "f -` z \<inter> space M \<in> sets M"
+ proof (auto simp add: prod_sets_def vimage_Times)
+ fix x y
+ assume x: "x \<in> sets a1" and y: "y \<in> sets a2"
+ have "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M =
+ ((fst \<circ> f) -` x \<inter> space M) \<inter> ((snd \<circ> f) -` y \<inter> space M)"
+ by blast
+ also have "... \<in> sets M" using x y q1 q2
+ by blast
+ finally show "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M \<in> sets M" .
+ qed
+ qed
qed
-lemma finite_measure_space_finite_prod_measure:
- assumes "finite_measure_space M" and "finite_measure_space M'"
- shows "finite_measure_space (prod_measure_space M M')"
-proof (rule finite_Pow_additivity_sufficient)
- interpret M: finite_measure_space M by fact
- interpret M': finite_measure_space M' by fact
+lemma (in sigma_finite_measure) prod_measure_times:
+ assumes "sigma_finite_measure N \<nu>"
+ and "A1 \<in> sets M" "A2 \<in> sets N"
+ shows "prod_measure M \<mu> N \<nu> (A1 \<times> A2) = \<mu> A1 * \<nu> A2"
+ oops
- from M.finite_space M'.finite_space
- show "finite (space (prod_measure_space M M'))" and
- "sets (prod_measure_space M M') = Pow (space (prod_measure_space M M'))"
- by (simp_all add: finite_prod_measure_space[OF assms])
+lemma (in sigma_finite_measure) sigma_finite_prod_measure_space:
+ assumes "sigma_finite_measure N \<nu>"
+ shows "sigma_finite_measure (prod_measure_space M N) (prod_measure M \<mu> N \<nu>)"
+ oops
+
+lemma (in finite_measure_space) simple_function_finite:
+ "simple_function f"
+ unfolding simple_function_def sets_eq_Pow using finite_space by auto
+
+lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M"
+ unfolding measurable_def sets_eq_Pow by auto
- show "additive (prod_measure_space M M') (measure (prod_measure_space M M'))"
- unfolding additive_def finite_prod_measure_space[OF assms]
- proof (simp, safe)
- fix x y assume subs: "x \<subseteq> space M \<times> space M'" "y \<subseteq> space M \<times> space M'"
- and disj_x_y: "x \<inter> y = {}"
- have "\<And>z. measure M' (Pair z -` x \<union> Pair z -` y) =
- measure M' (Pair z -` x) + measure M' (Pair z -` y)"
- using disj_x_y subs by (subst M'.measure_additive) (auto simp: M'.sets_eq_Pow)
- thus "prod_measure M M' (x \<union> y) = prod_measure M M' x + prod_measure M M' y"
- unfolding prod_measure_def M.integral_finite_singleton
- by (simp_all add: setsum_addf[symmetric] field_simps)
- qed
+lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
+ "positive_integral f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
+proof -
+ have *: "positive_integral f = positive_integral (\<lambda>x. \<Sum>y\<in>space M. f y * indicator {y} x)"
+ by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space])
+ show ?thesis unfolding * using borel_measurable_finite[of f]
+ by (simp add: positive_integral_setsum positive_integral_cmult_indicator sets_eq_Pow)
+qed
- show "positive (prod_measure_space M M') (measure (prod_measure_space M M'))"
- unfolding positive_def
- by (auto intro!: setsum_nonneg mult_nonneg_nonneg M'.positive M.positive
- simp add: M.integral_zero finite_prod_measure_space[OF assms]
- prod_measure_def M.integral_finite_singleton
- M.sets_eq_Pow M'.sets_eq_Pow)
+lemma (in finite_measure_space) finite_prod_measure_times:
+ assumes "finite_measure_space N \<nu>"
+ and "A1 \<in> sets M" "A2 \<in> sets N"
+ shows "prod_measure M \<mu> N \<nu> (A1 \<times> A2) = \<mu> A1 * \<nu> A2"
+proof -
+ interpret N: finite_measure_space N \<nu> by fact
+ have *: "\<And>x. \<nu> (Pair x -` (A1 \<times> A2)) * \<mu> {x} = (if x \<in> A1 then \<nu> A2 * \<mu> {x} else 0)"
+ by (auto simp: vimage_Times comp_def)
+ have "finite A1"
+ using `A1 \<in> sets M` finite_space by (auto simp: sets_eq_Pow intro: finite_subset)
+ then have "\<mu> A1 = (\<Sum>x\<in>A1. \<mu> {x})" using `A1 \<in> sets M`
+ by (auto intro!: measure_finite_singleton simp: sets_eq_Pow)
+ then show ?thesis using `A1 \<in> sets M`
+ unfolding prod_measure_def positive_integral_finite_eq_setsum *
+ by (auto simp add: sets_eq_Pow setsum_right_distrib[symmetric] mult_commute setsum_cases[OF finite_space])
qed
-lemma finite_measure_space_finite_prod_measure_alterantive:
- assumes M: "finite_measure_space M" and M': "finite_measure_space M'"
- shows "finite_measure_space \<lparr> space = space M \<times> space M', sets = Pow (space M \<times> space M'), measure = prod_measure M M' \<rparr>"
- (is "finite_measure_space ?M")
+lemma (in finite_measure_space) finite_prod_measure_space:
+ assumes "finite_measure_space N \<nu>"
+ shows "prod_measure_space M N = \<lparr> space = space M \<times> space N, sets = Pow (space M \<times> space N) \<rparr>"
proof -
- interpret M: finite_measure_space M by fact
- interpret M': finite_measure_space M' by fact
-
- show ?thesis
- using finite_measure_space_finite_prod_measure[OF assms]
- unfolding prod_measure_space_def M.sets_eq_Pow M'.sets_eq_Pow
- using M.finite_space M'.finite_space
- by (simp add: sigma_prod_sets_finite)
+ interpret N: finite_measure_space N \<nu> by fact
+ show ?thesis using finite_space N.finite_space
+ by (simp add: sigma_def prod_measure_space_def sigma_prod_sets_finite sets_eq_Pow N.sets_eq_Pow)
qed
+lemma (in finite_measure_space) finite_measure_space_finite_prod_measure:
+ assumes "finite_measure_space N \<nu>"
+ shows "finite_measure_space (prod_measure_space M N) (prod_measure M \<mu> N \<nu>)"
+ unfolding finite_prod_measure_space[OF assms]
+proof (rule finite_measure_spaceI)
+ interpret N: finite_measure_space N \<nu> by fact
+
+ let ?P = "\<lparr>space = space M \<times> space N, sets = Pow (space M \<times> space N)\<rparr>"
+ show "measure_space ?P (prod_measure M \<mu> N \<nu>)"
+ proof (rule sigma_algebra.finite_additivity_sufficient)
+ show "sigma_algebra ?P" by (rule sigma_algebra_Pow)
+ show "finite (space ?P)" using finite_space N.finite_space by auto
+ from finite_prod_measure_times[OF assms, of "{}" "{}"]
+ show "positive (prod_measure M \<mu> N \<nu>)"
+ unfolding positive_def by simp
+
+ show "additive ?P (prod_measure M \<mu> N \<nu>)"
+ unfolding additive_def
+ apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric]
+ intro!: positive_integral_cong)
+ apply (subst N.measure_additive[symmetric])
+ by (auto simp: N.sets_eq_Pow sets_eq_Pow)
+ qed
+ show "finite (space ?P)" using finite_space N.finite_space by auto
+ show "sets ?P = Pow (space ?P)" by simp
+
+ fix x assume "x \<in> space ?P"
+ with finite_prod_measure_times[OF assms, of "{fst x}" "{snd x}"]
+ finite_measure[of "{fst x}"] N.finite_measure[of "{snd x}"]
+ show "prod_measure M \<mu> N \<nu> {x} \<noteq> \<omega>"
+ by (auto simp add: sets_eq_Pow N.sets_eq_Pow elim!: SigmaE)
+qed
+
+lemma (in finite_measure_space) finite_measure_space_finite_prod_measure_alterantive:
+ assumes N: "finite_measure_space N \<nu>"
+ shows "finite_measure_space \<lparr> space = space M \<times> space N, sets = Pow (space M \<times> space N) \<rparr> (prod_measure M \<mu> N \<nu>)"
+ (is "finite_measure_space ?M ?m")
+ unfolding finite_prod_measure_space[OF N, symmetric]
+ using finite_measure_space_finite_prod_measure[OF N] .
+
end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Radon_Nikodym.thy Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,1059 @@
+theory Radon_Nikodym
+imports Lebesgue_Integration
+begin
+
+lemma (in measure_space) measure_finitely_subadditive:
+ assumes "finite I" "A ` I \<subseteq> sets M"
+ shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))"
+using assms proof induct
+ case (insert i I)
+ then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN)
+ then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)"
+ using insert by (simp add: measure_subadditive)
+ also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))"
+ using insert by (auto intro!: add_left_mono)
+ finally show ?case .
+qed simp
+
+lemma (in sigma_algebra) borel_measurable_restricted:
+ fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M"
+ shows "f \<in> borel_measurable (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<longleftrightarrow>
+ (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
+ (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
+proof -
+ interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
+ have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R"
+ by (auto intro!: measurable_cong)
+ show ?thesis unfolding *
+ unfolding in_borel_measurable_borel_space
+ proof (simp, safe)
+ fix S :: "pinfreal set" assume "S \<in> sets borel_space"
+ "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
+ then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
+ then have f: "?f -` S \<inter> A \<in> sets M"
+ using `A \<in> sets M` sets_into_space by fastsimp
+ show "?f -` S \<inter> space M \<in> sets M"
+ proof cases
+ assume "0 \<in> S"
+ then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)"
+ using `A \<in> sets M` sets_into_space by auto
+ then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff)
+ next
+ assume "0 \<notin> S"
+ then have "?f -` S \<inter> space M = ?f -` S \<inter> A"
+ using `A \<in> sets M` sets_into_space
+ by (auto simp: indicator_def split: split_if_asm)
+ then show ?thesis using f by auto
+ qed
+ next
+ fix S :: "pinfreal set" assume "S \<in> sets borel_space"
+ "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M"
+ then have f: "?f -` S \<inter> space M \<in> sets M" by auto
+ then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
+ using `A \<in> sets M` sets_into_space
+ apply (simp add: image_iff)
+ apply (rule bexI[OF _ f])
+ by auto
+ qed
+qed
+
+lemma (in sigma_algebra) simple_function_eq_borel_measurable:
+ fixes f :: "'a \<Rightarrow> pinfreal"
+ shows "simple_function f \<longleftrightarrow>
+ finite (f`space M) \<and> f \<in> borel_measurable M"
+ using simple_function_borel_measurable[of f]
+ borel_measurable_simple_function[of f]
+ by (fastsimp simp: simple_function_def)
+
+lemma (in measure_space) simple_function_restricted:
+ fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M"
+ shows "sigma_algebra.simple_function (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)"
+ (is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?f")
+proof -
+ interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
+ have "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
+ proof cases
+ assume "A = space M"
+ then have "f`A = ?f`space M" by (fastsimp simp: image_iff)
+ then show ?thesis by simp
+ next
+ assume "A \<noteq> space M"
+ then obtain x where x: "x \<in> space M" "x \<notin> A"
+ using sets_into_space `A \<in> sets M` by auto
+ have *: "?f`space M = f`A \<union> {0}"
+ proof (auto simp add: image_iff)
+ show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0"
+ using x by (auto intro!: bexI[of _ x])
+ next
+ fix x assume "x \<in> A"
+ then show "\<exists>y\<in>space M. f x = f y * indicator A y"
+ using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
+ next
+ fix x
+ assume "indicator A x \<noteq> (0::pinfreal)"
+ then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
+ moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
+ ultimately show "f x = 0" by auto
+ qed
+ then show ?thesis by auto
+ qed
+ then show ?thesis
+ unfolding simple_function_eq_borel_measurable
+ R.simple_function_eq_borel_measurable
+ unfolding borel_measurable_restricted[OF `A \<in> sets M`]
+ by auto
+qed
+
+lemma (in measure_space) simple_integral_restricted:
+ assumes "A \<in> sets M"
+ assumes sf: "simple_function (\<lambda>x. f x * indicator A x)"
+ shows "measure_space.simple_integral (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<mu> f = simple_integral (\<lambda>x. f x * indicator A x)"
+ (is "_ = simple_integral ?f")
+ unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]]
+ unfolding simple_integral_def
+proof (simp, safe intro!: setsum_mono_zero_cong_left)
+ from sf show "finite (?f ` space M)"
+ unfolding simple_function_def by auto
+next
+ fix x assume "x \<in> A"
+ then show "f x \<in> ?f ` space M"
+ using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x])
+next
+ fix x assume "x \<in> space M" "?f x \<notin> f`A"
+ then have "x \<notin> A" by (auto simp: image_iff)
+ then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp
+next
+ fix x assume "x \<in> A"
+ then have "f x \<noteq> 0 \<Longrightarrow>
+ f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M"
+ using `A \<in> sets M` sets_into_space
+ by (auto simp: indicator_def split: split_if_asm)
+ then show "f x * \<mu> (f -` {f x} \<inter> A) =
+ f x * \<mu> (?f -` {f x} \<inter> space M)"
+ unfolding pinfreal_mult_cancel_left by auto
+qed
+
+lemma (in measure_space) positive_integral_restricted:
+ assumes "A \<in> sets M"
+ shows "measure_space.positive_integral (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)"
+ (is "measure_space.positive_integral ?R \<mu> f = positive_integral ?f")
+proof -
+ have msR: "measure_space ?R \<mu>" by (rule restricted_measure_space[OF `A \<in> sets M`])
+ then interpret R: measure_space ?R \<mu> .
+ have saR: "sigma_algebra ?R" by fact
+ have *: "R.positive_integral f = R.positive_integral ?f"
+ by (auto intro!: R.positive_integral_cong)
+ show ?thesis
+ unfolding * R.positive_integral_def positive_integral_def
+ unfolding simple_function_restricted[OF `A \<in> sets M`]
+ apply (simp add: SUPR_def)
+ apply (rule arg_cong[where f=Sup])
+ proof (auto simp: image_iff simple_integral_restricted[OF `A \<in> sets M`])
+ fix g assume "simple_function (\<lambda>x. g x * indicator A x)"
+ "g \<le> f" "\<forall>x\<in>A. \<omega> \<noteq> g x"
+ then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> (\<forall>y\<in>space M. \<omega> \<noteq> x y) \<and>
+ simple_integral (\<lambda>x. g x * indicator A x) = simple_integral x"
+ apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
+ by (auto simp: indicator_def le_fun_def)
+ next
+ fix g assume g: "simple_function g" "g \<le> (\<lambda>x. f x * indicator A x)"
+ "\<forall>x\<in>space M. \<omega> \<noteq> g x"
+ then have *: "(\<lambda>x. g x * indicator A x) = g"
+ "\<And>x. g x * indicator A x = g x"
+ "\<And>x. g x \<le> f x"
+ by (auto simp: le_fun_def expand_fun_eq indicator_def split: split_if_asm)
+ from g show "\<exists>x. simple_function (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and> (\<forall>xa\<in>A. \<omega> \<noteq> x xa) \<and>
+ simple_integral g = simple_integral (\<lambda>xa. x xa * indicator A xa)"
+ using `A \<in> sets M`[THEN sets_into_space]
+ apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
+ by (fastsimp simp: le_fun_def *)
+ qed
+qed
+
+lemma (in sigma_algebra) borel_measurable_psuminf:
+ assumes "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
+ using assms unfolding psuminf_def
+ by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand])
+
+lemma (in sigma_finite_measure) disjoint_sigma_finite:
+ "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and>
+ (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A"
+proof -
+ obtain A :: "nat \<Rightarrow> 'a set" where
+ range: "range A \<subseteq> sets M" and
+ space: "(\<Union>i. A i) = space M" and
+ measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+ using sigma_finite by auto
+
+ note range' = range_disjointed_sets[OF range] range
+
+ { fix i
+ have "\<mu> (disjointed A i) \<le> \<mu> (A i)"
+ using range' disjointed_subset[of A i] by (auto intro!: measure_mono)
+ then have "\<mu> (disjointed A i) \<noteq> \<omega>"
+ using measure[of i] by auto }
+ with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
+ show ?thesis by (auto intro!: exI[of _ "disjointed A"])
+qed
+
+lemma (in sigma_finite_measure) Ex_finite_integrable_function:
+ shows "\<exists>h\<in>borel_measurable M. positive_integral h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)"
+proof -
+ obtain A :: "nat \<Rightarrow> 'a set" where
+ range: "range A \<subseteq> sets M" and
+ space: "(\<Union>i. A i) = space M" and
+ measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" and
+ disjoint: "disjoint_family A"
+ using disjoint_sigma_finite by auto
+
+ let "?B i" = "2^Suc i * \<mu> (A i)"
+ have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
+ proof
+ fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
+ proof cases
+ assume "\<mu> (A i) = 0"
+ then show ?thesis by (auto intro!: exI[of _ 1])
+ next
+ assume not_0: "\<mu> (A i) \<noteq> 0"
+ then have "?B i \<noteq> \<omega>" using measure[of i] by auto
+ then have "inverse (?B i) \<noteq> 0" unfolding pinfreal_inverse_eq_0 by simp
+ then show ?thesis using measure[of i] not_0
+ by (auto intro!: exI[of _ "inverse (?B i) / 2"]
+ simp: pinfreal_noteq_omega_Ex zero_le_mult_iff zero_less_mult_iff mult_le_0_iff power_le_zero_eq)
+ qed
+ qed
+ from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
+ "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
+
+ let "?h x" = "\<Sum>\<^isub>\<infinity> i. n i * indicator (A i) x"
+ show ?thesis
+ proof (safe intro!: bexI[of _ ?h] del: notI)
+ have "positive_integral ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))"
+ apply (subst positive_integral_psuminf)
+ using range apply (fastsimp intro!: borel_measurable_pinfreal_times borel_measurable_const borel_measurable_indicator)
+ apply (subst positive_integral_cmult_indicator)
+ using range by auto
+ also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. Real ((1 / 2)^Suc i))"
+ proof (rule psuminf_le)
+ fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)"
+ using measure[of N] n[of N]
+ by (cases "n N") (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff mult_le_0_iff mult_less_0_iff power_less_zero_eq power_le_zero_eq inverse_eq_divide less_divide_eq power_divide split: split_if_asm)
+ qed
+ also have "\<dots> = Real 1"
+ by (rule suminf_imp_psuminf, rule power_half_series, auto)
+ finally show "positive_integral ?h \<noteq> \<omega>" by auto
+ next
+ fix x assume "x \<in> space M"
+ then obtain i where "x \<in> A i" using space[symmetric] by auto
+ from psuminf_cmult_indicator[OF disjoint, OF this]
+ have "?h x = n i" by simp
+ then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto
+ next
+ show "?h \<in> borel_measurable M" using range
+ by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times borel_measurable_indicator)
+ qed
+qed
+
+definition (in measure_space)
+ "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pinfreal))"
+
+lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
+ fixes e :: real assumes "0 < e"
+ assumes "finite_measure M s"
+ shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (s (space M)) \<le>
+ real (\<mu> A) - real (s A) \<and>
+ (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < real (\<mu> B) - real (s B))"
+proof -
+ let "?d A" = "real (\<mu> A) - real (s A)"
+ interpret M': finite_measure M s by fact
+
+ let "?A A" = "if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)
+ then {}
+ else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)"
+ def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}"
+
+ have A_simps[simp]:
+ "A 0 = {}"
+ "\<And>n. A (Suc n) = (A n \<union> ?A (A n))" unfolding A_def by simp_all
+
+ { fix A assume "A \<in> sets M"
+ have "?A A \<in> sets M"
+ by (auto intro!: someI2[of _ _ "\<lambda>A. A \<in> sets M"] simp: not_less) }
+ note A'_in_sets = this
+
+ { fix n have "A n \<in> sets M"
+ proof (induct n)
+ case (Suc n) thus "A (Suc n) \<in> sets M"
+ using A'_in_sets[of "A n"] by (auto split: split_if_asm)
+ qed (simp add: A_def) }
+ note A_in_sets = this
+ hence "range A \<subseteq> sets M" by auto
+
+ { fix n B
+ assume Ex: "\<exists>B. B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> -e"
+ hence False: "\<not> (\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B)" by (auto simp: not_less)
+ have "?d (A (Suc n)) \<le> ?d (A n) - e" unfolding A_simps if_not_P[OF False]
+ proof (rule someI2_ex[OF Ex])
+ fix B assume "B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"
+ hence "A n \<inter> B = {}" "B \<in> sets M" and dB: "?d B \<le> -e" by auto
+ hence "?d (A n \<union> B) = ?d (A n) + ?d B"
+ using `A n \<in> sets M` real_finite_measure_Union M'.real_finite_measure_Union by simp
+ also have "\<dots> \<le> ?d (A n) - e" using dB by simp
+ finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .
+ qed }
+ note dA_epsilon = this
+
+ { fix n have "?d (A (Suc n)) \<le> ?d (A n)"
+ proof (cases "\<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e")
+ case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp
+ next
+ case False
+ hence "\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B" by (auto simp: not_le)
+ thus ?thesis by simp
+ qed }
+ note dA_mono = this
+
+ show ?thesis
+ proof (cases "\<exists>n. \<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B")
+ case True then obtain n where B: "\<And>B. \<lbrakk> B \<in> sets M; B \<subseteq> space M - A n\<rbrakk> \<Longrightarrow> -e < ?d B" by blast
+ show ?thesis
+ proof (safe intro!: bexI[of _ "space M - A n"])
+ fix B assume "B \<in> sets M" "B \<subseteq> space M - A n"
+ from B[OF this] show "-e < ?d B" .
+ next
+ show "space M - A n \<in> sets M" by (rule compl_sets) fact
+ next
+ show "?d (space M) \<le> ?d (space M - A n)"
+ proof (induct n)
+ fix n assume "?d (space M) \<le> ?d (space M - A n)"
+ also have "\<dots> \<le> ?d (space M - A (Suc n))"
+ using A_in_sets sets_into_space dA_mono[of n]
+ real_finite_measure_Diff[of "space M"]
+ real_finite_measure_Diff[of "space M"]
+ M'.real_finite_measure_Diff[of "space M"]
+ M'.real_finite_measure_Diff[of "space M"]
+ by (simp del: A_simps)
+ finally show "?d (space M) \<le> ?d (space M - A (Suc n))" .
+ qed simp
+ qed
+ next
+ case False hence B: "\<And>n. \<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"
+ by (auto simp add: not_less)
+ { fix n have "?d (A n) \<le> - real n * e"
+ proof (induct n)
+ case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps)
+ qed simp } note dA_less = this
+ have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
+ proof (rule incseq_SucI)
+ fix n show "- ?d (A n) \<le> - ?d (A (Suc n))" using dA_mono[of n] by auto
+ qed
+ from real_finite_continuity_from_below[of A] `range A \<subseteq> sets M`
+ M'.real_finite_continuity_from_below[of A]
+ have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"
+ by (auto intro!: LIMSEQ_diff)
+ obtain n :: nat where "- ?d (\<Union>i. A i) / e < real n" using reals_Archimedean2 by auto
+ moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less]
+ have "real n \<le> - ?d (\<Union>i. A i) / e" using `0<e` by (simp add: field_simps)
+ ultimately show ?thesis by auto
+ qed
+qed
+
+lemma (in finite_measure) Radon_Nikodym_aux:
+ assumes "finite_measure M s"
+ shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (s (space M)) \<le>
+ real (\<mu> A) - real (s A) \<and>
+ (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> real (\<mu> B) - real (s B))"
+proof -
+ let "?d A" = "real (\<mu> A) - real (s A)"
+ let "?P A B n" = "A \<in> sets M \<and> A \<subseteq> B \<and> ?d B \<le> ?d A \<and> (\<forall>C\<in>sets M. C \<subseteq> A \<longrightarrow> - 1 / real (Suc n) < ?d C)"
+
+ interpret M': finite_measure M s by fact
+
+ let "?r S" = "M\<lparr> space := S, sets := (\<lambda>C. S \<inter> C)`sets M\<rparr>"
+
+ { fix S n
+ assume S: "S \<in> sets M"
+ hence **: "\<And>X. X \<in> op \<inter> S ` sets M \<longleftrightarrow> X \<in> sets M \<and> X \<subseteq> S" by auto
+ from M'.restricted_finite_measure[of S] restricted_finite_measure[of S] S
+ have "finite_measure (?r S) \<mu>" "0 < 1 / real (Suc n)"
+ "finite_measure (?r S) s" by auto
+ from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X ..
+ hence "?P X S n"
+ proof (simp add: **, safe)
+ fix C assume C: "C \<in> sets M" "C \<subseteq> X" "X \<subseteq> S" and
+ *: "\<forall>B\<in>sets M. S \<inter> B \<subseteq> X \<longrightarrow> - (1 / real (Suc n)) < ?d (S \<inter> B)"
+ hence "C \<subseteq> S" "C \<subseteq> X" "S \<inter> C = C" by auto
+ with *[THEN bspec, OF `C \<in> sets M`]
+ show "- (1 / real (Suc n)) < ?d C" by auto
+ qed
+ hence "\<exists>A. ?P A S n" by auto }
+ note Ex_P = this
+
+ def A \<equiv> "nat_rec (space M) (\<lambda>n A. SOME B. ?P B A n)"
+ have A_Suc: "\<And>n. A (Suc n) = (SOME B. ?P B (A n) n)" by (simp add: A_def)
+ have A_0[simp]: "A 0 = space M" unfolding A_def by simp
+
+ { fix i have "A i \<in> sets M" unfolding A_def
+ proof (induct i)
+ case (Suc i)
+ from Ex_P[OF this, of i] show ?case unfolding nat_rec_Suc
+ by (rule someI2_ex) simp
+ qed simp }
+ note A_in_sets = this
+
+ { fix n have "?P (A (Suc n)) (A n) n"
+ using Ex_P[OF A_in_sets] unfolding A_Suc
+ by (rule someI2_ex) simp }
+ note P_A = this
+
+ have "range A \<subseteq> sets M" using A_in_sets by auto
+
+ have A_mono: "\<And>i. A (Suc i) \<subseteq> A i" using P_A by simp
+ have mono_dA: "mono (\<lambda>i. ?d (A i))" using P_A by (simp add: mono_iff_le_Suc)
+ have epsilon: "\<And>C i. \<lbrakk> C \<in> sets M; C \<subseteq> A (Suc i) \<rbrakk> \<Longrightarrow> - 1 / real (Suc i) < ?d C"
+ using P_A by auto
+
+ show ?thesis
+ proof (safe intro!: bexI[of _ "\<Inter>i. A i"])
+ show "(\<Inter>i. A i) \<in> sets M" using A_in_sets by auto
+ from `range A \<subseteq> sets M` A_mono
+ real_finite_continuity_from_above[of A]
+ M'.real_finite_continuity_from_above[of A]
+ have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (auto intro!: LIMSEQ_diff)
+ thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _]
+ by (rule_tac LIMSEQ_le_const) (auto intro!: exI)
+ next
+ fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)"
+ show "0 \<le> ?d B"
+ proof (rule ccontr)
+ assume "\<not> 0 \<le> ?d B"
+ hence "0 < - ?d B" by auto
+ from ex_inverse_of_nat_Suc_less[OF this]
+ obtain n where *: "?d B < - 1 / real (Suc n)"
+ by (auto simp: real_eq_of_nat inverse_eq_divide field_simps)
+ have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat_rec_Suc)
+ from epsilon[OF B(1) this] *
+ show False by auto
+ qed
+ qed
+qed
+
+lemma (in finite_measure) Radon_Nikodym_finite_measure:
+ assumes "finite_measure M \<nu>"
+ assumes "absolutely_continuous \<nu>"
+ shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+proof -
+ interpret M': finite_measure M \<nu> using assms(1) .
+
+ def G \<equiv> "{g \<in> borel_measurable M. \<forall>A\<in>sets M. positive_integral (\<lambda>x. g x * indicator A x) \<le> \<nu> A}"
+ have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
+ hence "G \<noteq> {}" by auto
+
+ { fix f g assume f: "f \<in> G" and g: "g \<in> G"
+ have "(\<lambda>x. max (g x) (f x)) \<in> G" (is "?max \<in> G") unfolding G_def
+ proof safe
+ show "?max \<in> borel_measurable M" using f g unfolding G_def by auto
+
+ let ?A = "{x \<in> space M. f x \<le> g x}"
+ have "?A \<in> sets M" using f g unfolding G_def by auto
+
+ fix A assume "A \<in> sets M"
+ hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto
+ have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"
+ using sets_into_space[OF `A \<in> sets M`] by auto
+
+ have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =
+ g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x"
+ by (auto simp: indicator_def max_def)
+ hence "positive_integral (\<lambda>x. max (g x) (f x) * indicator A x) =
+ positive_integral (\<lambda>x. g x * indicator (?A \<inter> A) x) +
+ positive_integral (\<lambda>x. f x * indicator ((space M - ?A) \<inter> A) x)"
+ using f g sets unfolding G_def
+ by (auto cong: positive_integral_cong intro!: positive_integral_add borel_measurable_indicator)
+ also have "\<dots> \<le> \<nu> (?A \<inter> A) + \<nu> ((space M - ?A) \<inter> A)"
+ using f g sets unfolding G_def by (auto intro!: add_mono)
+ also have "\<dots> = \<nu> A"
+ using M'.measure_additive[OF sets] union by auto
+ finally show "positive_integral (\<lambda>x. max (g x) (f x) * indicator A x) \<le> \<nu> A" .
+ qed }
+ note max_in_G = this
+
+ { fix f g assume "f \<up> g" and f: "\<And>i. f i \<in> G"
+ have "g \<in> G" unfolding G_def
+ proof safe
+ from `f \<up> g` have [simp]: "g = (SUP i. f i)" unfolding isoton_def by simp
+ have f_borel: "\<And>i. f i \<in> borel_measurable M" using f unfolding G_def by simp
+ thus "g \<in> borel_measurable M" by (auto intro!: borel_measurable_SUP)
+
+ fix A assume "A \<in> sets M"
+ hence "\<And>i. (\<lambda>x. f i x * indicator A x) \<in> borel_measurable M"
+ using f_borel by (auto intro!: borel_measurable_indicator)
+ from positive_integral_isoton[OF isoton_indicator[OF `f \<up> g`] this]
+ have SUP: "positive_integral (\<lambda>x. g x * indicator A x) =
+ (SUP i. positive_integral (\<lambda>x. f i x * indicator A x))"
+ unfolding isoton_def by simp
+ show "positive_integral (\<lambda>x. g x * indicator A x) \<le> \<nu> A" unfolding SUP
+ using f `A \<in> sets M` unfolding G_def by (auto intro!: SUP_leI)
+ qed }
+ note SUP_in_G = this
+
+ let ?y = "SUP g : G. positive_integral g"
+ have "?y \<le> \<nu> (space M)" unfolding G_def
+ proof (safe intro!: SUP_leI)
+ fix g assume "\<forall>A\<in>sets M. positive_integral (\<lambda>x. g x * indicator A x) \<le> \<nu> A"
+ from this[THEN bspec, OF top] show "positive_integral g \<le> \<nu> (space M)"
+ by (simp cong: positive_integral_cong)
+ qed
+ hence "?y \<noteq> \<omega>" using M'.finite_measure_of_space by auto
+ from SUPR_countable_SUPR[OF this `G \<noteq> {}`] guess ys .. note ys = this
+ hence "\<forall>n. \<exists>g. g\<in>G \<and> positive_integral g = ys n"
+ proof safe
+ fix n assume "range ys \<subseteq> positive_integral ` G"
+ hence "ys n \<in> positive_integral ` G" by auto
+ thus "\<exists>g. g\<in>G \<and> positive_integral g = ys n" by auto
+ qed
+ from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. positive_integral (gs n) = ys n" by auto
+ hence y_eq: "?y = (SUP i. positive_integral (gs i))" using ys by auto
+ let "?g i x" = "Max ((\<lambda>n. gs n x) ` {..i})"
+ def f \<equiv> "SUP i. ?g i"
+ have gs_not_empty: "\<And>i. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
+ { fix i have "?g i \<in> G"
+ proof (induct i)
+ case 0 thus ?case by simp fact
+ next
+ case (Suc i)
+ with Suc gs_not_empty `gs (Suc i) \<in> G` show ?case
+ by (auto simp add: atMost_Suc intro!: max_in_G)
+ qed }
+ note g_in_G = this
+ have "\<And>x. \<forall>i. ?g i x \<le> ?g (Suc i) x"
+ using gs_not_empty by (simp add: atMost_Suc)
+ hence isoton_g: "?g \<up> f" by (simp add: isoton_def le_fun_def f_def)
+ from SUP_in_G[OF this g_in_G] have "f \<in> G" .
+ hence [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto
+
+ have "(\<lambda>i. positive_integral (?g i)) \<up> positive_integral f"
+ using isoton_g g_in_G by (auto intro!: positive_integral_isoton simp: G_def f_def)
+ hence "positive_integral f = (SUP i. positive_integral (?g i))"
+ unfolding isoton_def by simp
+ also have "\<dots> = ?y"
+ proof (rule antisym)
+ show "(SUP i. positive_integral (?g i)) \<le> ?y"
+ using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_def)
+ show "?y \<le> (SUP i. positive_integral (?g i))" unfolding y_eq
+ by (auto intro!: SUP_mono positive_integral_mono Max_ge)
+ qed
+ finally have int_f_eq_y: "positive_integral f = ?y" .
+
+ let "?t A" = "\<nu> A - positive_integral (\<lambda>x. f x * indicator A x)"
+
+ have "finite_measure M ?t"
+ proof
+ show "?t {} = 0" by simp
+ show "?t (space M) \<noteq> \<omega>" using M'.finite_measure by simp
+ show "countably_additive M ?t" unfolding countably_additive_def
+ proof safe
+ fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A"
+ have "(\<Sum>\<^isub>\<infinity> n. positive_integral (\<lambda>x. f x * indicator (A n) x))
+ = positive_integral (\<lambda>x. (\<Sum>\<^isub>\<infinity>n. f x * indicator (A n) x))"
+ using `range A \<subseteq> sets M`
+ by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator)
+ also have "\<dots> = positive_integral (\<lambda>x. f x * indicator (\<Union>n. A n) x)"
+ apply (rule positive_integral_cong)
+ apply (subst psuminf_cmult_right)
+ unfolding psuminf_indicator[OF `disjoint_family A`] ..
+ finally have "(\<Sum>\<^isub>\<infinity> n. positive_integral (\<lambda>x. f x * indicator (A n) x))
+ = positive_integral (\<lambda>x. f x * indicator (\<Union>n. A n) x)" .
+ moreover have "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
+ using M'.measure_countably_additive A by (simp add: comp_def)
+ moreover have "\<And>i. positive_integral (\<lambda>x. f x * indicator (A i) x) \<le> \<nu> (A i)"
+ using A `f \<in> G` unfolding G_def by auto
+ moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<omega>" using M'.finite_measure A by (simp add: countable_UN)
+ moreover {
+ have "positive_integral (\<lambda>x. f x * indicator (\<Union>i. A i) x) \<le> \<nu> (\<Union>i. A i)"
+ using A `f \<in> G` unfolding G_def by (auto simp: countable_UN)
+ also have "\<nu> (\<Union>i. A i) < \<omega>" using v_fin by (simp add: pinfreal_less_\<omega>)
+ finally have "positive_integral (\<lambda>x. f x * indicator (\<Union>i. A i) x) \<noteq> \<omega>"
+ by (simp add: pinfreal_less_\<omega>) }
+ ultimately
+ show "(\<Sum>\<^isub>\<infinity> n. ?t (A n)) = ?t (\<Union>i. A i)"
+ apply (subst psuminf_minus) by simp_all
+ qed
+ qed
+ then interpret M: finite_measure M ?t .
+
+ have ac: "absolutely_continuous ?t" using `absolutely_continuous \<nu>` unfolding absolutely_continuous_def by auto
+
+ have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0"
+ proof (rule ccontr)
+ assume "\<not> ?thesis"
+ then obtain A where A: "A \<in> sets M" and pos: "0 < ?t A"
+ by (auto simp: not_le)
+ note pos
+ also have "?t A \<le> ?t (space M)"
+ using M.measure_mono[of A "space M"] A sets_into_space by simp
+ finally have pos_t: "0 < ?t (space M)" by simp
+ moreover
+ hence pos_M: "0 < \<mu> (space M)"
+ using ac top unfolding absolutely_continuous_def by auto
+ moreover
+ have "positive_integral (\<lambda>x. f x * indicator (space M) x) \<le> \<nu> (space M)"
+ using `f \<in> G` unfolding G_def by auto
+ hence "positive_integral (\<lambda>x. f x * indicator (space M) x) \<noteq> \<omega>"
+ using M'.finite_measure_of_space by auto
+ moreover
+ def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
+ ultimately have b: "b \<noteq> 0" "b \<noteq> \<omega>"
+ using M'.finite_measure_of_space
+ by (auto simp: pinfreal_inverse_eq_0 finite_measure_of_space)
+
+ have "finite_measure M (\<lambda>A. b * \<mu> A)" (is "finite_measure M ?b")
+ proof
+ show "?b {} = 0" by simp
+ show "?b (space M) \<noteq> \<omega>" using finite_measure_of_space b by auto
+ show "countably_additive M ?b"
+ unfolding countably_additive_def psuminf_cmult_right
+ using measure_countably_additive by auto
+ qed
+
+ from M.Radon_Nikodym_aux[OF this]
+ obtain A0 where "A0 \<in> sets M" and
+ space_less_A0: "real (?t (space M)) - real (b * \<mu> (space M)) \<le> real (?t A0) - real (b * \<mu> A0)" and
+ *: "\<And>B. \<lbrakk> B \<in> sets M ; B \<subseteq> A0 \<rbrakk> \<Longrightarrow> 0 \<le> real (?t B) - real (b * \<mu> B)" by auto
+ { fix B assume "B \<in> sets M" "B \<subseteq> A0"
+ with *[OF this] have "b * \<mu> B \<le> ?t B"
+ using M'.finite_measure b finite_measure
+ by (cases "b * \<mu> B", cases "?t B") (auto simp: field_simps) }
+ note bM_le_t = this
+
+ let "?f0 x" = "f x + b * indicator A0 x"
+
+ { fix A assume A: "A \<in> sets M"
+ hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
+ have "positive_integral (\<lambda>x. ?f0 x * indicator A x) =
+ positive_integral (\<lambda>x. f x * indicator A x + b * indicator (A \<inter> A0) x)"
+ by (auto intro!: positive_integral_cong simp: field_simps indicator_inter_arith)
+ hence "positive_integral (\<lambda>x. ?f0 x * indicator A x) =
+ positive_integral (\<lambda>x. f x * indicator A x) + b * \<mu> (A \<inter> A0)"
+ using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A
+ by (simp add: borel_measurable_indicator positive_integral_add positive_integral_cmult_indicator) }
+ note f0_eq = this
+
+ { fix A assume A: "A \<in> sets M"
+ hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
+ have f_le_v: "positive_integral (\<lambda>x. f x * indicator A x) \<le> \<nu> A"
+ using `f \<in> G` A unfolding G_def by auto
+ note f0_eq[OF A]
+ also have "positive_integral (\<lambda>x. f x * indicator A x) + b * \<mu> (A \<inter> A0) \<le>
+ positive_integral (\<lambda>x. f x * indicator A x) + ?t (A \<inter> A0)"
+ using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M`
+ by (auto intro!: add_left_mono)
+ also have "\<dots> \<le> positive_integral (\<lambda>x. f x * indicator A x) + ?t A"
+ using M.measure_mono[simplified, OF _ `A \<inter> A0 \<in> sets M` `A \<in> sets M`]
+ by (auto intro!: add_left_mono)
+ also have "\<dots> \<le> \<nu> A"
+ using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`]
+ by (cases "positive_integral (\<lambda>x. f x * indicator A x)", cases "\<nu> A", auto)
+ finally have "positive_integral (\<lambda>x. ?f0 x * indicator A x) \<le> \<nu> A" . }
+ hence "?f0 \<in> G" using `A0 \<in> sets M` unfolding G_def
+ by (auto intro!: borel_measurable_indicator borel_measurable_pinfreal_add borel_measurable_pinfreal_times)
+
+ have real: "?t (space M) \<noteq> \<omega>" "?t A0 \<noteq> \<omega>"
+ "b * \<mu> (space M) \<noteq> \<omega>" "b * \<mu> A0 \<noteq> \<omega>"
+ using `A0 \<in> sets M` b
+ finite_measure[of A0] M.finite_measure[of A0]
+ finite_measure_of_space M.finite_measure_of_space
+ by auto
+
+ have int_f_finite: "positive_integral f \<noteq> \<omega>"
+ using M'.finite_measure_of_space pos_t unfolding pinfreal_zero_less_diff_iff
+ by (auto cong: positive_integral_cong)
+
+ have "?t (space M) > b * \<mu> (space M)" unfolding b_def
+ apply (simp add: field_simps)
+ apply (subst mult_assoc[symmetric])
+ apply (subst pinfreal_mult_inverse)
+ using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
+ using pinfreal_mult_strict_right_mono[of "Real (1/2)" 1 "?t (space M)"]
+ by simp_all
+ hence "0 < ?t (space M) - b * \<mu> (space M)"
+ by (simp add: pinfreal_zero_less_diff_iff)
+ also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
+ using space_less_A0 pos_M pos_t b real[unfolded pinfreal_noteq_omega_Ex] by auto
+ finally have "b * \<mu> A0 < ?t A0" unfolding pinfreal_zero_less_diff_iff .
+ hence "0 < ?t A0" by auto
+ hence "0 < \<mu> A0" using ac unfolding absolutely_continuous_def
+ using `A0 \<in> sets M` by auto
+ hence "0 < b * \<mu> A0" using b by auto
+
+ from int_f_finite this
+ have "?y + 0 < positive_integral f + b * \<mu> A0" unfolding int_f_eq_y
+ by (rule pinfreal_less_add)
+ also have "\<dots> = positive_integral ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
+ by (simp cong: positive_integral_cong)
+ finally have "?y < positive_integral ?f0" by simp
+
+ moreover from `?f0 \<in> G` have "positive_integral ?f0 \<le> ?y" by (auto intro!: le_SUPI)
+ ultimately show False by auto
+ qed
+
+ show ?thesis
+ proof (safe intro!: bexI[of _ f])
+ fix A assume "A\<in>sets M"
+ show "\<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+ proof (rule antisym)
+ show "positive_integral (\<lambda>x. f x * indicator A x) \<le> \<nu> A"
+ using `f \<in> G` `A \<in> sets M` unfolding G_def by auto
+ show "\<nu> A \<le> positive_integral (\<lambda>x. f x * indicator A x)"
+ using upper_bound[THEN bspec, OF `A \<in> sets M`]
+ by (simp add: pinfreal_zero_le_diff)
+ qed
+ qed simp
+qed
+
+lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
+ assumes "measure_space M \<nu>"
+ assumes "absolutely_continuous \<nu>"
+ shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+proof -
+ interpret v: measure_space M \<nu> by fact
+ let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<omega>}"
+ let ?a = "SUP Q:?Q. \<mu> Q"
+
+ have "{} \<in> ?Q" using v.empty_measure by auto
+ then have Q_not_empty: "?Q \<noteq> {}" by blast
+
+ have "?a \<le> \<mu> (space M)" using sets_into_space
+ by (auto intro!: SUP_leI measure_mono top)
+ then have "?a \<noteq> \<omega>" using finite_measure_of_space
+ by auto
+ from SUPR_countable_SUPR[OF this Q_not_empty]
+ obtain Q'' where "range Q'' \<subseteq> \<mu> ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
+ by auto
+ then have "\<forall>i. \<exists>Q'. Q'' i = \<mu> Q' \<and> Q' \<in> ?Q" by auto
+ from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = \<mu> (Q' i)" "\<And>i. Q' i \<in> ?Q"
+ by auto
+ then have a_Lim: "?a = (SUP i::nat. \<mu> (Q' i))" using a by simp
+ let "?O n" = "\<Union>i\<le>n. Q' i"
+ have Union: "(SUP i. \<mu> (?O i)) = \<mu> (\<Union>i. ?O i)"
+ proof (rule continuity_from_below[of ?O])
+ show "range ?O \<subseteq> sets M" using Q' by (auto intro!: finite_UN)
+ show "\<And>i. ?O i \<subseteq> ?O (Suc i)" by fastsimp
+ qed
+
+ have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto
+
+ have O_sets: "\<And>i. ?O i \<in> sets M"
+ using Q' by (auto intro!: finite_UN Un)
+ then have O_in_G: "\<And>i. ?O i \<in> ?Q"
+ proof (safe del: notI)
+ fix i have "Q' ` {..i} \<subseteq> sets M"
+ using Q' by (auto intro: finite_UN)
+ with v.measure_finitely_subadditive[of "{.. i}" Q']
+ have "\<nu> (?O i) \<le> (\<Sum>i\<le>i. \<nu> (Q' i))" by auto
+ also have "\<dots> < \<omega>" unfolding setsum_\<omega> pinfreal_less_\<omega> using Q' by auto
+ finally show "\<nu> (?O i) \<noteq> \<omega>" unfolding pinfreal_less_\<omega> by auto
+ qed auto
+ have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastsimp
+
+ have a_eq: "?a = \<mu> (\<Union>i. ?O i)" unfolding Union[symmetric]
+ proof (rule antisym)
+ show "?a \<le> (SUP i. \<mu> (?O i))" unfolding a_Lim
+ using Q' by (auto intro!: SUP_mono measure_mono finite_UN)
+ show "(SUP i. \<mu> (?O i)) \<le> ?a" unfolding SUPR_def
+ proof (safe intro!: Sup_mono, unfold bex_simps)
+ fix i
+ have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
+ then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<omega>) \<and>
+ \<mu> (\<Union>Q' ` {..i}) \<le> \<mu> x"
+ using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
+ qed
+ qed
+
+ let "?O_0" = "(\<Union>i. ?O i)"
+ have "?O_0 \<in> sets M" using Q' by auto
+
+ { fix A assume *: "A \<in> ?Q" "A \<subseteq> space M - ?O_0"
+ then have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)"
+ using Q' by (auto intro!: measure_additive countable_UN)
+ also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))"
+ proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
+ show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
+ using * O_sets by auto
+ qed fastsimp
+ also have "\<dots> \<le> ?a"
+ proof (safe intro!: SUPR_bound)
+ fix i have "?O i \<union> A \<in> ?Q"
+ proof (safe del: notI)
+ show "?O i \<union> A \<in> sets M" using O_sets * by auto
+ from O_in_G[of i]
+ moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A"
+ using v.measure_subadditive[of "?O i" A] * O_sets by auto
+ ultimately show "\<nu> (?O i \<union> A) \<noteq> \<omega>"
+ using * by auto
+ qed
+ then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule le_SUPI)
+ qed
+ finally have "\<mu> A = 0" unfolding a_eq using finite_measure[OF `?O_0 \<in> sets M`]
+ by (cases "\<mu> A") (auto simp: pinfreal_noteq_omega_Ex) }
+ note stetic = this
+
+ def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> ?O 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n"
+
+ { fix i have "Q i \<in> sets M" unfolding Q_def using Q'[of 0] by (cases i) (auto intro: O_sets) }
+ note Q_sets = this
+
+ { fix i have "\<nu> (Q i) \<noteq> \<omega>"
+ proof (cases i)
+ case 0 then show ?thesis
+ unfolding Q_def using Q'[of 0] by simp
+ next
+ case (Suc n)
+ then show ?thesis unfolding Q_def
+ using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q` O_mono
+ using v.measure_Diff[of "?O n" "?O (Suc n)"] by auto
+ qed }
+ note Q_omega = this
+
+ { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
+ proof (induct j)
+ case 0 then show ?case by (simp add: Q_def)
+ next
+ case (Suc j)
+ have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastsimp
+ have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
+ then have "(\<Union>i\<le>Suc j. Q' i) = (\<Union>i\<le>j. Q' i) \<union> Q (Suc j)"
+ by (simp add: UN_Un[symmetric] Q_def del: UN_Un)
+ then show ?case using Suc by (auto simp add: eq atMost_Suc)
+ qed }
+ then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp
+ then have O_0_eq_Q: "?O_0 = (\<Union>j. Q j)" by fastsimp
+
+ have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
+ \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f x * indicator (Q i \<inter> A) x))"
+ proof
+ fix i
+ have indicator_eq: "\<And>f x A. (f x :: pinfreal) * indicator (Q i \<inter> A) x * indicator (Q i) x
+ = (f x * indicator (Q i) x) * indicator A x"
+ unfolding indicator_def by auto
+
+ have fm: "finite_measure (M\<lparr>space := Q i, sets := op \<inter> (Q i) ` sets M\<rparr>) \<mu>"
+ (is "finite_measure ?R \<mu>") by (rule restricted_finite_measure[OF Q_sets[of i]])
+ then interpret R: finite_measure ?R .
+ have fmv: "finite_measure ?R \<nu>"
+ unfolding finite_measure_def finite_measure_axioms_def
+ proof
+ show "measure_space ?R \<nu>"
+ using v.restricted_measure_space Q_sets[of i] by auto
+ show "\<nu> (space ?R) \<noteq> \<omega>"
+ using Q_omega by simp
+ qed
+ have "R.absolutely_continuous \<nu>"
+ using `absolutely_continuous \<nu>` `Q i \<in> sets M`
+ by (auto simp: R.absolutely_continuous_def absolutely_continuous_def)
+ from finite_measure.Radon_Nikodym_finite_measure[OF fm fmv this]
+ obtain f where f: "(\<lambda>x. f x * indicator (Q i) x) \<in> borel_measurable M"
+ and f_int: "\<And>A. A\<in>sets M \<Longrightarrow> \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. (f x * indicator (Q i) x) * indicator A x)"
+ unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`]
+ positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq)
+ then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
+ \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f x * indicator (Q i \<inter> A) x))"
+ by (fastsimp intro!: exI[of _ "\<lambda>x. f x * indicator (Q i) x"] positive_integral_cong
+ simp: indicator_def)
+ qed
+ from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M"
+ and f: "\<And>A i. A \<in> sets M \<Longrightarrow>
+ \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f i x * indicator (Q i \<inter> A) x)"
+ by auto
+ let "?f x" =
+ "(\<Sum>\<^isub>\<infinity> i. f i x * indicator (Q i) x) + \<omega> * indicator (space M - ?O_0) x"
+ show ?thesis
+ proof (safe intro!: bexI[of _ ?f])
+ show "?f \<in> borel_measurable M"
+ by (safe intro!: borel_measurable_psuminf borel_measurable_pinfreal_times
+ borel_measurable_pinfreal_add borel_measurable_indicator
+ borel_measurable_const borel Q_sets O_sets Diff countable_UN)
+ fix A assume "A \<in> sets M"
+ let ?C = "(space M - (\<Union>i. Q i)) \<inter> A"
+ have *:
+ "\<And>x i. indicator A x * (f i x * indicator (Q i) x) =
+ f i x * indicator (Q i \<inter> A) x"
+ "\<And>x i. (indicator A x * indicator (space M - (\<Union>i. UNION {..i} Q')) x :: pinfreal) =
+ indicator ?C x" unfolding O_0_eq_Q by (auto simp: indicator_def)
+ have "positive_integral (\<lambda>x. ?f x * indicator A x) =
+ (\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> ?C"
+ unfolding f[OF `A \<in> sets M`]
+ apply (simp del: pinfreal_times(2) add: field_simps)
+ apply (subst positive_integral_add)
+ apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
+ borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
+ unfolding psuminf_cmult_right[symmetric]
+ apply (subst positive_integral_psuminf)
+ apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
+ borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
+ apply (subst positive_integral_cmult)
+ apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
+ borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
+ unfolding *
+ apply (subst positive_integral_indicator)
+ apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const Int
+ borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
+ by simp
+ moreover have "(\<Sum>\<^isub>\<infinity>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)"
+ proof (rule v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified])
+ show "range (\<lambda>i. Q i \<inter> A) \<subseteq> sets M"
+ using Q_sets `A \<in> sets M` by auto
+ show "disjoint_family (\<lambda>i. Q i \<inter> A)"
+ by (fastsimp simp: disjoint_family_on_def Q_def
+ split: nat.split_asm)
+ qed
+ moreover have "\<omega> * \<mu> ?C = \<nu> ?C"
+ proof cases
+ assume null: "\<mu> ?C = 0"
+ hence "?C \<in> null_sets" using Q_sets `A \<in> sets M` by auto
+ with `absolutely_continuous \<nu>` and null
+ show ?thesis by (simp add: absolutely_continuous_def)
+ next
+ assume not_null: "\<mu> ?C \<noteq> 0"
+ have "\<nu> ?C = \<omega>"
+ proof (rule ccontr)
+ assume "\<nu> ?C \<noteq> \<omega>"
+ then have "?C \<in> ?Q"
+ using Q_sets `A \<in> sets M` by auto
+ from stetic[OF this] not_null
+ show False unfolding O_0_eq_Q by auto
+ qed
+ then show ?thesis using not_null by simp
+ qed
+ moreover have "?C \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
+ using Q_sets `A \<in> sets M` by (auto intro!: countable_UN)
+ moreover have "((\<Union>i. Q i) \<inter> A) \<union> ?C = A" "((\<Union>i. Q i) \<inter> A) \<inter> ?C = {}"
+ using `A \<in> sets M` sets_into_space by auto
+ ultimately show "\<nu> A = positive_integral (\<lambda>x. ?f x * indicator A x)"
+ using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" ?C] by auto
+ qed
+qed
+
+lemma (in measure_space) positive_integral_translated_density:
+ assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ shows "measure_space.positive_integral M (\<lambda>A. positive_integral (\<lambda>x. f x * indicator A x)) g =
+ positive_integral (\<lambda>x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _")
+proof -
+ from measure_space_density[OF assms(1)]
+ interpret T: measure_space M ?T .
+
+ from borel_measurable_implies_simple_function_sequence[OF assms(2)]
+ obtain G where G: "\<And>i. simple_function (G i)" "G \<up> g" by blast
+ note G_borel = borel_measurable_simple_function[OF this(1)]
+
+ from T.positive_integral_isoton[OF `G \<up> g` G_borel]
+ have *: "(\<lambda>i. T.positive_integral (G i)) \<up> T.positive_integral g" .
+
+ { fix i
+ have [simp]: "finite (G i ` space M)"
+ using G(1) unfolding simple_function_def by auto
+ have "T.positive_integral (G i) = T.simple_integral (G i)"
+ using G T.positive_integral_eq_simple_integral by simp
+ also have "\<dots> = positive_integral (\<lambda>x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x))"
+ apply (simp add: T.simple_integral_def)
+ apply (subst positive_integral_cmult[symmetric])
+ using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
+ apply (subst positive_integral_setsum[symmetric])
+ using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
+ by (simp add: setsum_right_distrib field_simps)
+ also have "\<dots> = positive_integral (\<lambda>x. f x * G i x)"
+ by (auto intro!: positive_integral_cong
+ simp: indicator_def if_distrib setsum_cases)
+ finally have "T.positive_integral (G i) = positive_integral (\<lambda>x. f x * G i x)" . }
+ with * have eq_Tg: "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> T.positive_integral g" by simp
+
+ from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)"
+ unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right)
+ then have "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)"
+ using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times)
+ with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)"
+ unfolding isoton_def by simp
+qed
+
+lemma (in sigma_finite_measure) Radon_Nikodym:
+ assumes "measure_space M \<nu>"
+ assumes "absolutely_continuous \<nu>"
+ shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+proof -
+ from Ex_finite_integrable_function
+ obtain h where finite: "positive_integral h \<noteq> \<omega>" and
+ borel: "h \<in> borel_measurable M" and
+ pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
+ "\<And>x. x \<in> space M \<Longrightarrow> h x < \<omega>" by auto
+ let "?T A" = "positive_integral (\<lambda>x. h x * indicator A x)"
+ from measure_space_density[OF borel] finite
+ interpret T: finite_measure M ?T
+ unfolding finite_measure_def finite_measure_axioms_def
+ by (simp cong: positive_integral_cong)
+ have "\<And>N. N \<in> sets M \<Longrightarrow> {x \<in> space M. h x \<noteq> 0 \<and> indicator N x \<noteq> (0::pinfreal)} = N"
+ using sets_into_space pos by (force simp: indicator_def)
+ then have "T.absolutely_continuous \<nu>" using assms(2) borel
+ unfolding T.absolutely_continuous_def absolutely_continuous_def
+ by (fastsimp simp: borel_measurable_indicator positive_integral_0_iff)
+ from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this]
+ obtain f where f_borel: "f \<in> borel_measurable M" and
+ fT: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = T.positive_integral (\<lambda>x. f x * indicator A x)" by auto
+ show ?thesis
+ proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"])
+ show "(\<lambda>x. h x * f x) \<in> borel_measurable M"
+ using borel f_borel by (auto intro: borel_measurable_pinfreal_times)
+ fix A assume "A \<in> sets M"
+ then have "(\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
+ using f_borel by (auto intro: borel_measurable_pinfreal_times borel_measurable_indicator)
+ from positive_integral_translated_density[OF borel this]
+ show "\<nu> A = positive_integral (\<lambda>x. h x * f x * indicator A x)"
+ unfolding fT[OF `A \<in> sets M`] by (simp add: field_simps)
+ qed
+qed
+
+section "Radon Nikodym derivative"
+
+definition (in sigma_finite_measure)
+ "RN_deriv \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and>
+ (\<forall>A \<in> sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x))"
+
+lemma (in sigma_finite_measure) RN_deriv:
+ assumes "measure_space M \<nu>"
+ assumes "absolutely_continuous \<nu>"
+ shows "RN_deriv \<nu> \<in> borel_measurable M" (is ?borel)
+ and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = positive_integral (\<lambda>x. RN_deriv \<nu> x * indicator A x)"
+ (is "\<And>A. _ \<Longrightarrow> ?int A")
+proof -
+ note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]
+ thus ?borel unfolding RN_deriv_def by (rule someI2_ex) auto
+ fix A assume "A \<in> sets M"
+ from Ex show "?int A" unfolding RN_deriv_def
+ by (rule someI2_ex) (simp add: `A \<in> sets M`)
+qed
+
+lemma (in sigma_finite_measure) RN_deriv_singleton:
+ assumes "measure_space M \<nu>"
+ and ac: "absolutely_continuous \<nu>"
+ and "{x} \<in> sets M"
+ shows "\<nu> {x} = RN_deriv \<nu> x * \<mu> {x}"
+proof -
+ note deriv = RN_deriv[OF assms(1, 2)]
+ from deriv(2)[OF `{x} \<in> sets M`]
+ have "\<nu> {x} = positive_integral (\<lambda>w. RN_deriv \<nu> x * indicator {x} w)"
+ by (auto simp: indicator_def intro!: positive_integral_cong)
+ thus ?thesis using positive_integral_cmult_indicator[OF `{x} \<in> sets M`]
+ by auto
+qed
+
+theorem (in finite_measure_space) RN_deriv_finite_measure:
+ assumes "measure_space M \<nu>"
+ and ac: "absolutely_continuous \<nu>"
+ and "x \<in> space M"
+ shows "\<nu> {x} = RN_deriv \<nu> x * \<mu> {x}"
+proof -
+ have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
+ from RN_deriv_singleton[OF assms(1,2) this] show ?thesis .
+qed
+
+end
--- a/src/HOL/Probability/SeriesPlus.thy Mon Aug 23 16:47:57 2010 +0200
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,125 +0,0 @@
-theory SeriesPlus
- imports Complex_Main Nat_Bijection
-begin
-
-text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
-
-lemma choice2: "(!!x. \<exists>y z. Q x y z) ==> \<exists>f g. \<forall>x. Q x (f x) (g x)"
- by metis
-
-lemma range_subsetD: "range f \<subseteq> B \<Longrightarrow> f i \<in> B"
- by blast
-
-
-lemma suminf_2dimen:
- fixes f:: "nat * nat \<Rightarrow> real"
- assumes f_nneg: "!!m n. 0 \<le> f(m,n)"
- and fsums: "!!m. (\<lambda>n. f (m,n)) sums (g m)"
- and sumg: "summable g"
- shows "(f o prod_decode) sums suminf g"
-proof (simp add: sums_def)
- {
- fix x
- have "0 \<le> f x"
- by (cases x) (simp add: f_nneg)
- } note [iff] = this
- have g_nneg: "!!m. 0 \<le> g m"
- proof -
- fix m
- have "0 \<le> suminf (\<lambda>n. f (m,n))"
- by (rule suminf_0_le, simp add: f_nneg, metis fsums sums_iff)
- thus "0 \<le> g m" using fsums [of m]
- by (auto simp add: sums_iff)
- qed
- show "(\<lambda>n. \<Sum>x = 0..<n. f (prod_decode x)) ----> suminf g"
- proof (rule increasing_LIMSEQ, simp add: f_nneg)
- fix n
- def i \<equiv> "Max (Domain (prod_decode ` {0..<n}))"
- def j \<equiv> "Max (Range (prod_decode ` {0..<n}))"
- have ij: "prod_decode ` {0..<n} \<subseteq> ({0..i} \<times> {0..j})"
- by (force simp add: i_def j_def
- intro: finite_imageI Max_ge finite_Domain finite_Range)
- have "(\<Sum>x = 0..<n. f (prod_decode x)) = setsum f (prod_decode ` {0..<n})"
- using setsum_reindex [of prod_decode "{0..<n}" f]
- by (simp add: o_def)
- (metis inj_prod_decode inj_prod_decode)
- also have "... \<le> setsum f ({0..i} \<times> {0..j})"
- by (rule setsum_mono2) (auto simp add: ij)
- also have "... = setsum (\<lambda>m. setsum (\<lambda>n. f (m,n)) {0..j}) {0..<Suc i}"
- by (metis atLeast0AtMost atLeast0LessThan lessThan_Suc_atMost
- setsum_cartesian_product split_eta)
- also have "... \<le> setsum g {0..<Suc i}"
- proof (rule setsum_mono, simp add: less_Suc_eq_le)
- fix m
- assume m: "m \<le> i"
- thus " (\<Sum>n = 0..j. f (m, n)) \<le> g m" using fsums [of m]
- by (auto simp add: sums_iff)
- (metis atLeastLessThanSuc_atLeastAtMost f_nneg series_pos_le f_nneg)
- qed
- finally have "(\<Sum>x = 0..<n. f (prod_decode x)) \<le> setsum g {0..<Suc i}" .
- also have "... \<le> suminf g"
- by (rule series_pos_le [OF sumg]) (simp add: g_nneg)
- finally show "(\<Sum>x = 0..<n. f (prod_decode x)) \<le> suminf g" .
- next
- fix e :: real
- assume e: "0 < e"
- with summable_sums [OF sumg]
- obtain N where "\<bar>setsum g {0..<N} - suminf g\<bar> < e/2" and nz: "N>0"
- by (simp add: sums_def LIMSEQ_iff_nz dist_real_def)
- (metis e half_gt_zero le_refl that)
- hence gless: "suminf g < setsum g {0..<N} + e/2" using series_pos_le [OF sumg]
- by (simp add: g_nneg)
- { fix m
- assume m: "m<N"
- hence enneg: "0 < e / (2 * real N)" using e
- by (simp add: zero_less_divide_iff)
- hence "\<exists>j. \<bar>(\<Sum>n = 0..<j. f (m, n)) - g m\<bar> < e/(2 * real N)"
- using fsums [of m] m
- by (force simp add: sums_def LIMSEQ_def dist_real_def)
- hence "\<exists>j. g m < setsum (\<lambda>n. f (m,n)) {0..<j} + e/(2 * real N)"
- using fsums [of m]
- by (auto simp add: sums_iff)
- (metis abs_diff_less_iff add_less_cancel_right eq_diff_eq')
- }
- hence "\<exists>jj. \<forall>m. m<N \<longrightarrow> g m < (\<Sum>n = 0..<jj m. f (m, n)) + e/(2 * real N)"
- by (force intro: choice)
- then obtain jj where jj:
- "!!m. m<N \<Longrightarrow> g m < (\<Sum>n = 0..<jj m. f (m, n)) + e/(2 * real N)"
- by auto
- def IJ \<equiv> "SIGMA i : {0..<N}. {0..<jj i}"
- have "setsum g {0..<N} <
- (\<Sum>m = 0..<N. (\<Sum>n = 0..<jj m. f (m, n)) + e/(2 * real N))"
- by (auto simp add: nz jj intro: setsum_strict_mono)
- also have "... = (\<Sum>m = 0..<N. \<Sum>n = 0..<jj m. f (m, n)) + e/2" using nz
- by (auto simp add: setsum_addf real_of_nat_def)
- also have "... = setsum f IJ + e/2"
- by (simp add: IJ_def setsum_Sigma)
- finally have "setsum g {0..<N} < setsum f IJ + e/2" .
- hence glessf: "suminf g < setsum f IJ + e" using gless
- by auto
- have finite_IJ: "finite IJ"
- by (simp add: IJ_def)
- def NIJ \<equiv> "Max (prod_decode -` IJ)"
- have IJsb: "IJ \<subseteq> prod_decode ` {0..NIJ}"
- proof (auto simp add: NIJ_def)
- fix i j
- assume ij:"(i,j) \<in> IJ"
- hence "(i,j) = prod_decode (prod_encode (i,j))"
- by simp
- thus "(i,j) \<in> prod_decode ` {0..Max (prod_decode -` IJ)}"
- by (rule image_eqI)
- (simp add: ij finite_vimageI [OF finite_IJ inj_prod_decode])
- qed
- have "setsum f IJ \<le> setsum f (prod_decode `{0..NIJ})"
- by (rule setsum_mono2) (auto simp add: IJsb)
- also have "... = (\<Sum>k = 0..NIJ. f (prod_decode k))"
- by (simp add: setsum_reindex inj_prod_decode)
- also have "... = (\<Sum>k = 0..<Suc NIJ. f (prod_decode k))"
- by (metis atLeast0AtMost atLeast0LessThan lessThan_Suc_atMost)
- finally have "setsum f IJ \<le> (\<Sum>k = 0..<Suc NIJ. f (prod_decode k))" .
- thus "\<exists>n. suminf g \<le> (\<Sum>x = 0..<n. f (prod_decode x)) + e" using glessf
- by (metis add_right_mono local.glessf not_leE order_le_less_trans less_asym)
- qed
-qed
-
-end
--- a/src/HOL/Probability/Sigma_Algebra.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Probability/Sigma_Algebra.thy Tue Aug 24 08:22:17 2010 +0200
@@ -1,12 +1,12 @@
(* Title: Sigma_Algebra.thy
Author: Stefan Richter, Markus Wenzel, TU Muenchen
- Plus material from the Hurd/Coble measure theory development,
+ Plus material from the Hurd/Coble measure theory development,
translated by Lawrence Paulson.
*)
header {* Sigma Algebras *}
-theory Sigma_Algebra imports Complex_Main begin
+theory Sigma_Algebra imports Main Countable FuncSet begin
text {* Sigma algebras are an elementary concept in measure
theory. To measure --- that is to integrate --- functions, we first have
@@ -18,8 +18,8 @@
subsection {* Algebras *}
-record 'a algebra =
- space :: "'a set"
+record 'a algebra =
+ space :: "'a set"
sets :: "'a set set"
locale algebra =
@@ -35,20 +35,20 @@
lemma (in algebra) sets_into_space: "x \<in> sets M \<Longrightarrow> x \<subseteq> space M"
by (metis PowD contra_subsetD space_closed)
-lemma (in algebra) Int [intro]:
+lemma (in algebra) Int [intro]:
assumes a: "a \<in> sets M" and b: "b \<in> sets M" shows "a \<inter> b \<in> sets M"
proof -
- have "((space M - a) \<union> (space M - b)) \<in> sets M"
+ have "((space M - a) \<union> (space M - b)) \<in> sets M"
by (metis a b compl_sets Un)
moreover
- have "a \<inter> b = space M - ((space M - a) \<union> (space M - b))"
+ have "a \<inter> b = space M - ((space M - a) \<union> (space M - b))"
using space_closed a b
by blast
ultimately show ?thesis
by (metis compl_sets)
qed
-lemma (in algebra) Diff [intro]:
+lemma (in algebra) Diff [intro]:
assumes a: "a \<in> sets M" and b: "b \<in> sets M" shows "a - b \<in> sets M"
proof -
have "(a \<inter> (space M - b)) \<in> sets M"
@@ -60,74 +60,143 @@
by metis
qed
-lemma (in algebra) finite_union [intro]:
+lemma (in algebra) finite_union [intro]:
"finite X \<Longrightarrow> X \<subseteq> sets M \<Longrightarrow> Union X \<in> sets M"
- by (induct set: finite) (auto simp add: Un)
+ by (induct set: finite) (auto simp add: Un)
+lemma algebra_iff_Int:
+ "algebra M \<longleftrightarrow>
+ sets M \<subseteq> Pow (space M) & {} \<in> sets M &
+ (\<forall>a \<in> sets M. space M - a \<in> sets M) &
+ (\<forall>a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)"
+proof (auto simp add: algebra.Int, auto simp add: algebra_def)
+ fix a b
+ assume ab: "sets M \<subseteq> Pow (space M)"
+ "\<forall>a\<in>sets M. space M - a \<in> sets M"
+ "\<forall>a\<in>sets M. \<forall>b\<in>sets M. a \<inter> b \<in> sets M"
+ "a \<in> sets M" "b \<in> sets M"
+ hence "a \<union> b = space M - ((space M - a) \<inter> (space M - b))"
+ by blast
+ also have "... \<in> sets M"
+ by (metis ab)
+ finally show "a \<union> b \<in> sets M" .
+qed
+
+lemma (in algebra) insert_in_sets:
+ assumes "{x} \<in> sets M" "A \<in> sets M" shows "insert x A \<in> sets M"
+proof -
+ have "{x} \<union> A \<in> sets M" using assms by (rule Un)
+ thus ?thesis by auto
+qed
+
+lemma (in algebra) Int_space_eq1 [simp]: "x \<in> sets M \<Longrightarrow> space M \<inter> x = x"
+ by (metis Int_absorb1 sets_into_space)
+
+lemma (in algebra) Int_space_eq2 [simp]: "x \<in> sets M \<Longrightarrow> x \<inter> space M = x"
+ by (metis Int_absorb2 sets_into_space)
+
+lemma (in algebra) restricted_algebra:
+ assumes "S \<in> sets M"
+ shows "algebra (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>)"
+ (is "algebra ?r")
+ using assms by unfold_locales auto
subsection {* Sigma Algebras *}
locale sigma_algebra = algebra +
- assumes countable_UN [intro]:
+ assumes countable_nat_UN [intro]:
"!!A. range A \<subseteq> sets M \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
+lemma countable_UN_eq:
+ fixes A :: "'i::countable \<Rightarrow> 'a set"
+ shows "(range A \<subseteq> sets M \<longrightarrow> (\<Union>i. A i) \<in> sets M) \<longleftrightarrow>
+ (range (A \<circ> from_nat) \<subseteq> sets M \<longrightarrow> (\<Union>i. (A \<circ> from_nat) i) \<in> sets M)"
+proof -
+ let ?A' = "A \<circ> from_nat"
+ have *: "(\<Union>i. ?A' i) = (\<Union>i. A i)" (is "?l = ?r")
+ proof safe
+ fix x i assume "x \<in> A i" thus "x \<in> ?l"
+ by (auto intro!: exI[of _ "to_nat i"])
+ next
+ fix x i assume "x \<in> ?A' i" thus "x \<in> ?r"
+ by (auto intro!: exI[of _ "from_nat i"])
+ qed
+ have **: "range ?A' = range A"
+ using surj_range[OF surj_from_nat]
+ by (auto simp: image_compose intro!: imageI)
+ show ?thesis unfolding * ** ..
+qed
+
+lemma (in sigma_algebra) countable_UN[intro]:
+ fixes A :: "'i::countable \<Rightarrow> 'a set"
+ assumes "A`X \<subseteq> sets M"
+ shows "(\<Union>x\<in>X. A x) \<in> sets M"
+proof -
+ let "?A i" = "if i \<in> X then A i else {}"
+ from assms have "range ?A \<subseteq> sets M" by auto
+ with countable_nat_UN[of "?A \<circ> from_nat"] countable_UN_eq[of ?A M]
+ have "(\<Union>x. ?A x) \<in> sets M" by auto
+ moreover have "(\<Union>x. ?A x) = (\<Union>x\<in>X. A x)" by (auto split: split_if_asm)
+ ultimately show ?thesis by simp
+qed
+
+lemma (in sigma_algebra) finite_UN:
+ assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
+ shows "(\<Union>i\<in>I. A i) \<in> sets M"
+ using assms by induct auto
+
lemma (in sigma_algebra) countable_INT [intro]:
- assumes a: "range a \<subseteq> sets M"
- shows "(\<Inter>i::nat. a i) \<in> sets M"
+ fixes A :: "'i::countable \<Rightarrow> 'a set"
+ assumes A: "A`X \<subseteq> sets M" "X \<noteq> {}"
+ shows "(\<Inter>i\<in>X. A i) \<in> sets M"
proof -
- from a have "\<forall>i. a i \<in> sets M" by fast
- hence "space M - (\<Union>i. space M - a i) \<in> sets M" by blast
+ from A have "\<forall>i\<in>X. A i \<in> sets M" by fast
+ hence "space M - (\<Union>i\<in>X. space M - A i) \<in> sets M" by blast
moreover
- have "(\<Inter>i. a i) = space M - (\<Union>i. space M - a i)" using space_closed a
+ have "(\<Inter>i\<in>X. A i) = space M - (\<Union>i\<in>X. space M - A i)" using space_closed A
by blast
ultimately show ?thesis by metis
qed
-lemma (in sigma_algebra) gen_countable_UN:
- fixes f :: "nat \<Rightarrow> 'c"
- shows "I = range f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> (\<Union>x\<in>I. A x) \<in> sets M"
-by auto
-
-lemma (in sigma_algebra) gen_countable_INT:
- fixes f :: "nat \<Rightarrow> 'c"
- shows "I = range f \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> (\<Inter>x\<in>I. A x) \<in> sets M"
-by auto
+lemma (in sigma_algebra) finite_INT:
+ assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
+ shows "(\<Inter>i\<in>I. A i) \<in> sets M"
+ using assms by (induct rule: finite_ne_induct) auto
lemma algebra_Pow:
- "algebra (| space = sp, sets = Pow sp |)"
- by (auto simp add: algebra_def)
+ "algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
+ by (auto simp add: algebra_def)
lemma sigma_algebra_Pow:
- "sigma_algebra (| space = sp, sets = Pow sp |)"
- by (auto simp add: sigma_algebra_def sigma_algebra_axioms_def algebra_Pow)
+ "sigma_algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
+ by (auto simp add: sigma_algebra_def sigma_algebra_axioms_def algebra_Pow)
+
+lemma sigma_algebra_iff:
+ "sigma_algebra M \<longleftrightarrow>
+ algebra M \<and> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
+ by (simp add: sigma_algebra_def sigma_algebra_axioms_def)
subsection {* Binary Unions *}
definition binary :: "'a \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
where "binary a b = (\<lambda>\<^isup>x. b)(0 := a)"
-lemma range_binary_eq: "range(binary a b) = {a,b}"
- by (auto simp add: binary_def)
-
-lemma Un_range_binary: "a \<union> b = (\<Union>i::nat. binary a b i)"
- by (simp add: UNION_eq_Union_image range_binary_eq)
+lemma range_binary_eq: "range(binary a b) = {a,b}"
+ by (auto simp add: binary_def)
-lemma Int_range_binary: "a \<inter> b = (\<Inter>i::nat. binary a b i)"
- by (simp add: INTER_eq_Inter_image range_binary_eq)
+lemma Un_range_binary: "a \<union> b = (\<Union>i::nat. binary a b i)"
+ by (simp add: UNION_eq_Union_image range_binary_eq)
-lemma sigma_algebra_iff:
- "sigma_algebra M \<longleftrightarrow>
- algebra M & (\<forall>A. range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
- by (simp add: sigma_algebra_def sigma_algebra_axioms_def)
+lemma Int_range_binary: "a \<inter> b = (\<Inter>i::nat. binary a b i)"
+ by (simp add: INTER_eq_Inter_image range_binary_eq)
lemma sigma_algebra_iff2:
"sigma_algebra M \<longleftrightarrow>
- sets M \<subseteq> Pow (space M) &
- {} \<in> sets M & (\<forall>s \<in> sets M. space M - s \<in> sets M) &
+ sets M \<subseteq> Pow (space M) \<and>
+ {} \<in> sets M \<and> (\<forall>s \<in> sets M. space M - s \<in> sets M) \<and>
(\<forall>A. range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
- by (force simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def
- algebra_def Un_range_binary)
-
+ by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def
+ algebra_def Un_range_binary)
subsection {* Initial Sigma Algebra *}
@@ -148,19 +217,21 @@
sigma where
"sigma sp A = (| space = sp, sets = sigma_sets sp A |)"
+lemma sets_sigma: "sets (sigma A B) = sigma_sets A B"
+ unfolding sigma_def by simp
lemma space_sigma [simp]: "space (sigma X B) = X"
- by (simp add: sigma_def)
+ by (simp add: sigma_def)
lemma sigma_sets_top: "sp \<in> sigma_sets sp A"
by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty)
lemma sigma_sets_into_sp: "A \<subseteq> Pow sp \<Longrightarrow> x \<in> sigma_sets sp A \<Longrightarrow> x \<subseteq> sp"
- by (erule sigma_sets.induct, auto)
+ by (erule sigma_sets.induct, auto)
-lemma sigma_sets_Un:
+lemma sigma_sets_Un:
"a \<in> sigma_sets sp A \<Longrightarrow> b \<in> sigma_sets sp A \<Longrightarrow> a \<union> b \<in> sigma_sets sp A"
-apply (simp add: Un_range_binary range_binary_eq)
+apply (simp add: Un_range_binary range_binary_eq)
apply (rule Union, simp add: binary_def COMBK_def fun_upd_apply)
done
@@ -169,21 +240,21 @@
shows "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Inter>i. a i) \<in> sigma_sets sp A"
proof -
assume ai: "\<And>i::nat. a i \<in> sigma_sets sp A"
- hence "\<And>i::nat. sp-(a i) \<in> sigma_sets sp A"
+ hence "\<And>i::nat. sp-(a i) \<in> sigma_sets sp A"
by (rule sigma_sets.Compl)
- hence "(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"
+ hence "(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"
by (rule sigma_sets.Union)
- hence "sp-(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"
+ hence "sp-(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"
by (rule sigma_sets.Compl)
- also have "sp-(\<Union>i. sp-(a i)) = sp Int (\<Inter>i. a i)"
+ also have "sp-(\<Union>i. sp-(a i)) = sp Int (\<Inter>i. a i)"
by auto
also have "... = (\<Inter>i. a i)" using ai
- by (blast dest: sigma_sets_into_sp [OF Asb])
- finally show ?thesis .
+ by (blast dest: sigma_sets_into_sp [OF Asb])
+ finally show ?thesis .
qed
lemma sigma_sets_INTER:
- assumes Asb: "A \<subseteq> Pow sp"
+ assumes Asb: "A \<subseteq> Pow sp"
and ai: "\<And>i::nat. i \<in> S \<Longrightarrow> a i \<in> sigma_sets sp A" and non: "S \<noteq> {}"
shows "(\<Inter>i\<in>S. a i) \<in> sigma_sets sp A"
proof -
@@ -197,13 +268,13 @@
qed
lemma (in sigma_algebra) sigma_sets_subset:
- assumes a: "a \<subseteq> sets M"
+ assumes a: "a \<subseteq> sets M"
shows "sigma_sets (space M) a \<subseteq> sets M"
proof
fix x
assume "x \<in> sigma_sets (space M) a"
from this show "x \<in> sets M"
- by (induct rule: sigma_sets.induct, auto) (metis a subsetD)
+ by (induct rule: sigma_sets.induct, auto) (metis a subsetD)
qed
lemma (in sigma_algebra) sigma_sets_eq:
@@ -219,19 +290,612 @@
lemma sigma_algebra_sigma_sets:
"a \<subseteq> Pow (space M) \<Longrightarrow> sets M = sigma_sets (space M) a \<Longrightarrow> sigma_algebra M"
apply (auto simp add: sigma_algebra_def sigma_algebra_axioms_def
- algebra_def sigma_sets.Empty sigma_sets.Compl sigma_sets_Un)
+ algebra_def sigma_sets.Empty sigma_sets.Compl sigma_sets_Un)
apply (blast dest: sigma_sets_into_sp)
apply (rule sigma_sets.Union, fast)
done
lemma sigma_algebra_sigma:
"a \<subseteq> Pow X \<Longrightarrow> sigma_algebra (sigma X a)"
- apply (rule sigma_algebra_sigma_sets)
- apply (auto simp add: sigma_def)
+ apply (rule sigma_algebra_sigma_sets)
+ apply (auto simp add: sigma_def)
done
lemma (in sigma_algebra) sigma_subset:
"a \<subseteq> sets M ==> sets (sigma (space M) a) \<subseteq> (sets M)"
by (simp add: sigma_def sigma_sets_subset)
+lemma (in sigma_algebra) restriction_in_sets:
+ fixes A :: "nat \<Rightarrow> 'a set"
+ assumes "S \<in> sets M"
+ and *: "range A \<subseteq> (\<lambda>A. S \<inter> A) ` sets M" (is "_ \<subseteq> ?r")
+ shows "range A \<subseteq> sets M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` sets M"
+proof -
+ { fix i have "A i \<in> ?r" using * by auto
+ hence "\<exists>B. A i = B \<inter> S \<and> B \<in> sets M" by auto
+ hence "A i \<subseteq> S" "A i \<in> sets M" using `S \<in> sets M` by auto }
+ thus "range A \<subseteq> sets M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` sets M"
+ by (auto intro!: image_eqI[of _ _ "(\<Union>i. A i)"])
+qed
+
+lemma (in sigma_algebra) restricted_sigma_algebra:
+ assumes "S \<in> sets M"
+ shows "sigma_algebra (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>)"
+ (is "sigma_algebra ?r")
+ unfolding sigma_algebra_def sigma_algebra_axioms_def
+proof safe
+ show "algebra ?r" using restricted_algebra[OF assms] .
+next
+ fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets ?r"
+ from restriction_in_sets[OF assms this[simplified]]
+ show "(\<Union>i. A i) \<in> sets ?r" by simp
+qed
+
+section {* Measurable functions *}
+
+definition
+ "measurable A B = {f \<in> space A -> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
+
+lemma (in sigma_algebra) measurable_sigma:
+ assumes B: "B \<subseteq> Pow X"
+ and f: "f \<in> space M -> X"
+ and ba: "\<And>y. y \<in> B \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
+ shows "f \<in> measurable M (sigma X B)"
+proof -
+ have "sigma_sets X B \<subseteq> {y . (f -` y) \<inter> space M \<in> sets M & y \<subseteq> X}"
+ proof clarify
+ fix x
+ assume "x \<in> sigma_sets X B"
+ thus "f -` x \<inter> space M \<in> sets M \<and> x \<subseteq> X"
+ proof induct
+ case (Basic a)
+ thus ?case
+ by (auto simp add: ba) (metis B subsetD PowD)
+ next
+ case Empty
+ thus ?case
+ by auto
+ next
+ case (Compl a)
+ have [simp]: "f -` X \<inter> space M = space M"
+ by (auto simp add: funcset_mem [OF f])
+ thus ?case
+ by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl)
+ next
+ case (Union a)
+ thus ?case
+ by (simp add: vimage_UN, simp only: UN_extend_simps(4))
+ (blast intro: countable_UN)
+ qed
+ qed
+ thus ?thesis
+ by (simp add: measurable_def sigma_algebra_axioms sigma_algebra_sigma B f)
+ (auto simp add: sigma_def)
+qed
+
+lemma measurable_cong:
+ assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
+ shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
+ unfolding measurable_def using assms
+ by (simp cong: vimage_inter_cong Pi_cong)
+
+lemma measurable_space:
+ "f \<in> measurable M A \<Longrightarrow> x \<in> space M \<Longrightarrow> f x \<in> space A"
+ unfolding measurable_def by auto
+
+lemma measurable_sets:
+ "f \<in> measurable M A \<Longrightarrow> S \<in> sets A \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
+ unfolding measurable_def by auto
+
+lemma (in sigma_algebra) measurable_subset:
+ "(\<And>S. S \<in> sets A \<Longrightarrow> S \<subseteq> space A) \<Longrightarrow> measurable M A \<subseteq> measurable M (sigma (space A) (sets A))"
+ by (auto intro: measurable_sigma measurable_sets measurable_space)
+
+lemma measurable_eqI:
+ "\<lbrakk> space m1 = space m1' ; space m2 = space m2' ;
+ sets m1 = sets m1' ; sets m2 = sets m2' \<rbrakk>
+ \<Longrightarrow> measurable m1 m2 = measurable m1' m2'"
+ by (simp add: measurable_def sigma_algebra_iff2)
+
+lemma (in sigma_algebra) measurable_const[intro, simp]:
+ assumes "c \<in> space M'"
+ shows "(\<lambda>x. c) \<in> measurable M M'"
+ using assms by (auto simp add: measurable_def)
+
+lemma (in sigma_algebra) measurable_If:
+ assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
+ assumes P: "{x\<in>space M. P x} \<in> sets M"
+ shows "(\<lambda>x. if P x then f x else g x) \<in> measurable M M'"
+ unfolding measurable_def
+proof safe
+ fix x assume "x \<in> space M"
+ thus "(if P x then f x else g x) \<in> space M'"
+ using measure unfolding measurable_def by auto
+next
+ fix A assume "A \<in> sets M'"
+ hence *: "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M =
+ ((f -` A \<inter> space M) \<inter> {x\<in>space M. P x}) \<union>
+ ((g -` A \<inter> space M) \<inter> (space M - {x\<in>space M. P x}))"
+ using measure unfolding measurable_def by (auto split: split_if_asm)
+ show "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M \<in> sets M"
+ using `A \<in> sets M'` measure P unfolding * measurable_def
+ by (auto intro!: Un)
+qed
+
+lemma (in sigma_algebra) measurable_If_set:
+ assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
+ assumes P: "A \<in> sets M"
+ shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
+proof (rule measurable_If[OF measure])
+ have "{x \<in> space M. x \<in> A} = A" using `A \<in> sets M` sets_into_space by auto
+ thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<in> sets M` by auto
+qed
+
+lemma (in algebra) measurable_ident[intro, simp]: "id \<in> measurable M M"
+ by (auto simp add: measurable_def)
+
+lemma measurable_comp[intro]:
+ fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
+ shows "f \<in> measurable a b \<Longrightarrow> g \<in> measurable b c \<Longrightarrow> (g o f) \<in> measurable a c"
+ apply (auto simp add: measurable_def vimage_compose)
+ apply (subgoal_tac "f -` g -` y \<inter> space a = f -` (g -` y \<inter> space b) \<inter> space a")
+ apply force+
+ done
+
+lemma measurable_strong:
+ fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
+ assumes f: "f \<in> measurable a b" and g: "g \<in> (space b -> space c)"
+ and a: "sigma_algebra a" and b: "sigma_algebra b" and c: "sigma_algebra c"
+ and t: "f ` (space a) \<subseteq> t"
+ and cb: "\<And>s. s \<in> sets c \<Longrightarrow> (g -` s) \<inter> t \<in> sets b"
+ shows "(g o f) \<in> measurable a c"
+proof -
+ have fab: "f \<in> (space a -> space b)"
+ and ba: "\<And>y. y \<in> sets b \<Longrightarrow> (f -` y) \<inter> (space a) \<in> sets a" using f
+ by (auto simp add: measurable_def)
+ have eq: "f -` g -` y \<inter> space a = f -` (g -` y \<inter> t) \<inter> space a" using t
+ by force
+ show ?thesis
+ apply (auto simp add: measurable_def vimage_compose a c)
+ apply (metis funcset_mem fab g)
+ apply (subst eq, metis ba cb)
+ done
+qed
+
+lemma measurable_mono1:
+ "a \<subseteq> b \<Longrightarrow> sigma_algebra \<lparr>space = X, sets = b\<rparr>
+ \<Longrightarrow> measurable \<lparr>space = X, sets = a\<rparr> c \<subseteq> measurable \<lparr>space = X, sets = b\<rparr> c"
+ by (auto simp add: measurable_def)
+
+lemma measurable_up_sigma:
+ "measurable A M \<subseteq> measurable (sigma (space A) (sets A)) M"
+ unfolding measurable_def
+ by (auto simp: sigma_def intro: sigma_sets.Basic)
+
+lemma (in sigma_algebra) measurable_range_reduce:
+ "\<lbrakk> f \<in> measurable M \<lparr>space = s, sets = Pow s\<rparr> ; s \<noteq> {} \<rbrakk>
+ \<Longrightarrow> f \<in> measurable M \<lparr>space = s \<inter> (f ` space M), sets = Pow (s \<inter> (f ` space M))\<rparr>"
+ by (simp add: measurable_def sigma_algebra_Pow del: Pow_Int_eq) blast
+
+lemma (in sigma_algebra) measurable_Pow_to_Pow:
+ "(sets M = Pow (space M)) \<Longrightarrow> f \<in> measurable M \<lparr>space = UNIV, sets = Pow UNIV\<rparr>"
+ by (auto simp add: measurable_def sigma_algebra_def sigma_algebra_axioms_def algebra_def)
+
+lemma (in sigma_algebra) measurable_Pow_to_Pow_image:
+ "sets M = Pow (space M)
+ \<Longrightarrow> f \<in> measurable M \<lparr>space = f ` space M, sets = Pow (f ` space M)\<rparr>"
+ by (simp add: measurable_def sigma_algebra_Pow) intro_locales
+
+lemma (in sigma_algebra) sigma_algebra_preimages:
+ fixes f :: "'x \<Rightarrow> 'a"
+ assumes "f \<in> A \<rightarrow> space M"
+ shows "sigma_algebra \<lparr> space = A, sets = (\<lambda>M. f -` M \<inter> A) ` sets M \<rparr>"
+ (is "sigma_algebra \<lparr> space = _, sets = ?F ` sets M \<rparr>")
+proof (simp add: sigma_algebra_iff2, safe)
+ show "{} \<in> ?F ` sets M" by blast
+next
+ fix S assume "S \<in> sets M"
+ moreover have "A - ?F S = ?F (space M - S)"
+ using assms by auto
+ ultimately show "A - ?F S \<in> ?F ` sets M"
+ by blast
+next
+ fix S :: "nat \<Rightarrow> 'x set" assume *: "range S \<subseteq> ?F ` sets M"
+ have "\<forall>i. \<exists>b. b \<in> sets M \<and> S i = ?F b"
+ proof safe
+ fix i
+ have "S i \<in> ?F ` sets M" using * by auto
+ then show "\<exists>b. b \<in> sets M \<and> S i = ?F b" by auto
+ qed
+ from choice[OF this] obtain b where b: "range b \<subseteq> sets M" "\<And>i. S i = ?F (b i)"
+ by auto
+ then have "(\<Union>i. S i) = ?F (\<Union>i. b i)" by auto
+ then show "(\<Union>i. S i) \<in> ?F ` sets M" using b(1) by blast
+qed
+
+section "Disjoint families"
+
+definition
+ disjoint_family_on where
+ "disjoint_family_on A S \<longleftrightarrow> (\<forall>m\<in>S. \<forall>n\<in>S. m \<noteq> n \<longrightarrow> A m \<inter> A n = {})"
+
+abbreviation
+ "disjoint_family A \<equiv> disjoint_family_on A UNIV"
+
+lemma range_subsetD: "range f \<subseteq> B \<Longrightarrow> f i \<in> B"
+ by blast
+
+lemma Int_Diff_disjoint: "A \<inter> B \<inter> (A - B) = {}"
+ by blast
+
+lemma Int_Diff_Un: "A \<inter> B \<union> (A - B) = A"
+ by blast
+
+lemma disjoint_family_subset:
+ "disjoint_family A \<Longrightarrow> (!!x. B x \<subseteq> A x) \<Longrightarrow> disjoint_family B"
+ by (force simp add: disjoint_family_on_def)
+
+lemma disjoint_family_on_mono:
+ "A \<subseteq> B \<Longrightarrow> disjoint_family_on f B \<Longrightarrow> disjoint_family_on f A"
+ unfolding disjoint_family_on_def by auto
+
+lemma disjoint_family_Suc:
+ assumes Suc: "!!n. A n \<subseteq> A (Suc n)"
+ shows "disjoint_family (\<lambda>i. A (Suc i) - A i)"
+proof -
+ {
+ fix m
+ have "!!n. A n \<subseteq> A (m+n)"
+ proof (induct m)
+ case 0 show ?case by simp
+ next
+ case (Suc m) thus ?case
+ by (metis Suc_eq_plus1 assms nat_add_commute nat_add_left_commute subset_trans)
+ qed
+ }
+ hence "!!m n. m < n \<Longrightarrow> A m \<subseteq> A n"
+ by (metis add_commute le_add_diff_inverse nat_less_le)
+ thus ?thesis
+ by (auto simp add: disjoint_family_on_def)
+ (metis insert_absorb insert_subset le_SucE le_antisym not_leE)
+qed
+
+definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set "
+ where "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
+
+lemma finite_UN_disjointed_eq: "(\<Union>i\<in>{0..<n}. disjointed A i) = (\<Union>i\<in>{0..<n}. A i)"
+proof (induct n)
+ case 0 show ?case by simp
+next
+ case (Suc n)
+ thus ?case by (simp add: atLeastLessThanSuc disjointed_def)
+qed
+
+lemma UN_disjointed_eq: "(\<Union>i. disjointed A i) = (\<Union>i. A i)"
+ apply (rule UN_finite2_eq [where k=0])
+ apply (simp add: finite_UN_disjointed_eq)
+ done
+
+lemma less_disjoint_disjointed: "m<n \<Longrightarrow> disjointed A m \<inter> disjointed A n = {}"
+ by (auto simp add: disjointed_def)
+
+lemma disjoint_family_disjointed: "disjoint_family (disjointed A)"
+ by (simp add: disjoint_family_on_def)
+ (metis neq_iff Int_commute less_disjoint_disjointed)
+
+lemma disjointed_subset: "disjointed A n \<subseteq> A n"
+ by (auto simp add: disjointed_def)
+
+lemma (in algebra) UNION_in_sets:
+ fixes A:: "nat \<Rightarrow> 'a set"
+ assumes A: "range A \<subseteq> sets M "
+ shows "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
+proof (induct n)
+ case 0 show ?case by simp
+next
+ case (Suc n)
+ thus ?case
+ by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff)
+qed
+
+lemma (in algebra) range_disjointed_sets:
+ assumes A: "range A \<subseteq> sets M "
+ shows "range (disjointed A) \<subseteq> sets M"
+proof (auto simp add: disjointed_def)
+ fix n
+ show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> sets M" using UNION_in_sets
+ by (metis A Diff UNIV_I image_subset_iff)
+qed
+
+lemma sigma_algebra_disjoint_iff:
+ "sigma_algebra M \<longleftrightarrow>
+ algebra M &
+ (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow>
+ (\<Union>i::nat. A i) \<in> sets M)"
+proof (auto simp add: sigma_algebra_iff)
+ fix A :: "nat \<Rightarrow> 'a set"
+ assume M: "algebra M"
+ and A: "range A \<subseteq> sets M"
+ and UnA: "\<forall>A. range A \<subseteq> sets M \<longrightarrow>
+ disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
+ hence "range (disjointed A) \<subseteq> sets M \<longrightarrow>
+ disjoint_family (disjointed A) \<longrightarrow>
+ (\<Union>i. disjointed A i) \<in> sets M" by blast
+ hence "(\<Union>i. disjointed A i) \<in> sets M"
+ by (simp add: algebra.range_disjointed_sets M A disjoint_family_disjointed)
+ thus "(\<Union>i::nat. A i) \<in> sets M" by (simp add: UN_disjointed_eq)
+qed
+
+subsection {* A Two-Element Series *}
+
+definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "
+ where "binaryset A B = (\<lambda>\<^isup>x. {})(0 := A, Suc 0 := B)"
+
+lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}"
+ apply (simp add: binaryset_def)
+ apply (rule set_ext)
+ apply (auto simp add: image_iff)
+ done
+
+lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"
+ by (simp add: UNION_eq_Union_image range_binaryset_eq)
+
+section {* Closed CDI *}
+
+definition
+ closed_cdi where
+ "closed_cdi M \<longleftrightarrow>
+ sets M \<subseteq> Pow (space M) &
+ (\<forall>s \<in> sets M. space M - s \<in> sets M) &
+ (\<forall>A. (range A \<subseteq> sets M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
+ (\<Union>i. A i) \<in> sets M) &
+ (\<forall>A. (range A \<subseteq> sets M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
+
+
+inductive_set
+ smallest_ccdi_sets :: "('a, 'b) algebra_scheme \<Rightarrow> 'a set set"
+ for M
+ where
+ Basic [intro]:
+ "a \<in> sets M \<Longrightarrow> a \<in> smallest_ccdi_sets M"
+ | Compl [intro]:
+ "a \<in> smallest_ccdi_sets M \<Longrightarrow> space M - a \<in> smallest_ccdi_sets M"
+ | Inc:
+ "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
+ \<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets M"
+ | Disj:
+ "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> disjoint_family A
+ \<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets M"
+ monos Pow_mono
+
+
+definition
+ smallest_closed_cdi where
+ "smallest_closed_cdi M = (|space = space M, sets = smallest_ccdi_sets M|)"
+
+lemma space_smallest_closed_cdi [simp]:
+ "space (smallest_closed_cdi M) = space M"
+ by (simp add: smallest_closed_cdi_def)
+
+lemma (in algebra) smallest_closed_cdi1: "sets M \<subseteq> sets (smallest_closed_cdi M)"
+ by (auto simp add: smallest_closed_cdi_def)
+
+lemma (in algebra) smallest_ccdi_sets:
+ "smallest_ccdi_sets M \<subseteq> Pow (space M)"
+ apply (rule subsetI)
+ apply (erule smallest_ccdi_sets.induct)
+ apply (auto intro: range_subsetD dest: sets_into_space)
+ done
+
+lemma (in algebra) smallest_closed_cdi2: "closed_cdi (smallest_closed_cdi M)"
+ apply (auto simp add: closed_cdi_def smallest_closed_cdi_def smallest_ccdi_sets)
+ apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) +
+ done
+
+lemma (in algebra) smallest_closed_cdi3:
+ "sets (smallest_closed_cdi M) \<subseteq> Pow (space M)"
+ by (simp add: smallest_closed_cdi_def smallest_ccdi_sets)
+
+lemma closed_cdi_subset: "closed_cdi M \<Longrightarrow> sets M \<subseteq> Pow (space M)"
+ by (simp add: closed_cdi_def)
+
+lemma closed_cdi_Compl: "closed_cdi M \<Longrightarrow> s \<in> sets M \<Longrightarrow> space M - s \<in> sets M"
+ by (simp add: closed_cdi_def)
+
+lemma closed_cdi_Inc:
+ "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow>
+ (\<Union>i. A i) \<in> sets M"
+ by (simp add: closed_cdi_def)
+
+lemma closed_cdi_Disj:
+ "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
+ by (simp add: closed_cdi_def)
+
+lemma closed_cdi_Un:
+ assumes cdi: "closed_cdi M" and empty: "{} \<in> sets M"
+ and A: "A \<in> sets M" and B: "B \<in> sets M"
+ and disj: "A \<inter> B = {}"
+ shows "A \<union> B \<in> sets M"
+proof -
+ have ra: "range (binaryset A B) \<subseteq> sets M"
+ by (simp add: range_binaryset_eq empty A B)
+ have di: "disjoint_family (binaryset A B)" using disj
+ by (simp add: disjoint_family_on_def binaryset_def Int_commute)
+ from closed_cdi_Disj [OF cdi ra di]
+ show ?thesis
+ by (simp add: UN_binaryset_eq)
+qed
+
+lemma (in algebra) smallest_ccdi_sets_Un:
+ assumes A: "A \<in> smallest_ccdi_sets M" and B: "B \<in> smallest_ccdi_sets M"
+ and disj: "A \<inter> B = {}"
+ shows "A \<union> B \<in> smallest_ccdi_sets M"
+proof -
+ have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets M)"
+ by (simp add: range_binaryset_eq A B smallest_ccdi_sets.Basic)
+ have di: "disjoint_family (binaryset A B)" using disj
+ by (simp add: disjoint_family_on_def binaryset_def Int_commute)
+ from Disj [OF ra di]
+ show ?thesis
+ by (simp add: UN_binaryset_eq)
+qed
+
+lemma (in algebra) smallest_ccdi_sets_Int1:
+ assumes a: "a \<in> sets M"
+ shows "b \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M"
+proof (induct rule: smallest_ccdi_sets.induct)
+ case (Basic x)
+ thus ?case
+ by (metis a Int smallest_ccdi_sets.Basic)
+next
+ case (Compl x)
+ have "a \<inter> (space M - x) = space M - ((space M - a) \<union> (a \<inter> x))"
+ by blast
+ also have "... \<in> smallest_ccdi_sets M"
+ by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2
+ Diff_disjoint Int_Diff Int_empty_right Un_commute
+ smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl
+ smallest_ccdi_sets_Un)
+ finally show ?case .
+next
+ case (Inc A)
+ have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
+ by blast
+ have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Inc
+ by blast
+ moreover have "(\<lambda>i. a \<inter> A i) 0 = {}"
+ by (simp add: Inc)
+ moreover have "!!n. (\<lambda>i. a \<inter> A i) n \<subseteq> (\<lambda>i. a \<inter> A i) (Suc n)" using Inc
+ by blast
+ ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M"
+ by (rule smallest_ccdi_sets.Inc)
+ show ?case
+ by (metis 1 2)
+next
+ case (Disj A)
+ have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
+ by blast
+ have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Disj
+ by blast
+ moreover have "disjoint_family (\<lambda>i. a \<inter> A i)" using Disj
+ by (auto simp add: disjoint_family_on_def)
+ ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M"
+ by (rule smallest_ccdi_sets.Disj)
+ show ?case
+ by (metis 1 2)
+qed
+
+
+lemma (in algebra) smallest_ccdi_sets_Int:
+ assumes b: "b \<in> smallest_ccdi_sets M"
+ shows "a \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M"
+proof (induct rule: smallest_ccdi_sets.induct)
+ case (Basic x)
+ thus ?case
+ by (metis b smallest_ccdi_sets_Int1)
+next
+ case (Compl x)
+ have "(space M - x) \<inter> b = space M - (x \<inter> b \<union> (space M - b))"
+ by blast
+ also have "... \<in> smallest_ccdi_sets M"
+ by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b
+ smallest_ccdi_sets.Compl smallest_ccdi_sets_Un)
+ finally show ?case .
+next
+ case (Inc A)
+ have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
+ by blast
+ have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Inc
+ by blast
+ moreover have "(\<lambda>i. A i \<inter> b) 0 = {}"
+ by (simp add: Inc)
+ moreover have "!!n. (\<lambda>i. A i \<inter> b) n \<subseteq> (\<lambda>i. A i \<inter> b) (Suc n)" using Inc
+ by blast
+ ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M"
+ by (rule smallest_ccdi_sets.Inc)
+ show ?case
+ by (metis 1 2)
+next
+ case (Disj A)
+ have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
+ by blast
+ have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Disj
+ by blast
+ moreover have "disjoint_family (\<lambda>i. A i \<inter> b)" using Disj
+ by (auto simp add: disjoint_family_on_def)
+ ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M"
+ by (rule smallest_ccdi_sets.Disj)
+ show ?case
+ by (metis 1 2)
+qed
+
+lemma (in algebra) sets_smallest_closed_cdi_Int:
+ "a \<in> sets (smallest_closed_cdi M) \<Longrightarrow> b \<in> sets (smallest_closed_cdi M)
+ \<Longrightarrow> a \<inter> b \<in> sets (smallest_closed_cdi M)"
+ by (simp add: smallest_ccdi_sets_Int smallest_closed_cdi_def)
+
+lemma (in algebra) sigma_property_disjoint_lemma:
+ assumes sbC: "sets M \<subseteq> C"
+ and ccdi: "closed_cdi (|space = space M, sets = C|)"
+ shows "sigma_sets (space M) (sets M) \<subseteq> C"
+proof -
+ have "smallest_ccdi_sets M \<in> {B . sets M \<subseteq> B \<and> sigma_algebra (|space = space M, sets = B|)}"
+ apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int
+ smallest_ccdi_sets_Int)
+ apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets)
+ apply (blast intro: smallest_ccdi_sets.Disj)
+ done
+ hence "sigma_sets (space M) (sets M) \<subseteq> smallest_ccdi_sets M"
+ by clarsimp
+ (drule sigma_algebra.sigma_sets_subset [where a="sets M"], auto)
+ also have "... \<subseteq> C"
+ proof
+ fix x
+ assume x: "x \<in> smallest_ccdi_sets M"
+ thus "x \<in> C"
+ proof (induct rule: smallest_ccdi_sets.induct)
+ case (Basic x)
+ thus ?case
+ by (metis Basic subsetD sbC)
+ next
+ case (Compl x)
+ thus ?case
+ by (blast intro: closed_cdi_Compl [OF ccdi, simplified])
+ next
+ case (Inc A)
+ thus ?case
+ by (auto intro: closed_cdi_Inc [OF ccdi, simplified])
+ next
+ case (Disj A)
+ thus ?case
+ by (auto intro: closed_cdi_Disj [OF ccdi, simplified])
+ qed
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in algebra) sigma_property_disjoint:
+ assumes sbC: "sets M \<subseteq> C"
+ and compl: "!!s. s \<in> C \<inter> sigma_sets (space M) (sets M) \<Longrightarrow> space M - s \<in> C"
+ and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M)
+ \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n))
+ \<Longrightarrow> (\<Union>i. A i) \<in> C"
+ and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M)
+ \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> C"
+ shows "sigma_sets (space M) (sets M) \<subseteq> C"
+proof -
+ have "sigma_sets (space M) (sets M) \<subseteq> C \<inter> sigma_sets (space M) (sets M)"
+ proof (rule sigma_property_disjoint_lemma)
+ show "sets M \<subseteq> C \<inter> sigma_sets (space M) (sets M)"
+ by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic)
+ next
+ show "closed_cdi \<lparr>space = space M, sets = C \<inter> sigma_sets (space M) (sets M)\<rparr>"
+ by (simp add: closed_cdi_def compl inc disj)
+ (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed
+ IntE sigma_sets.Compl range_subsetD sigma_sets.Union)
+ qed
+ thus ?thesis
+ by blast
+qed
+
end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/document/root.tex Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,31 @@
+
+% HOL/Multivariate_Analysis/document/root.tex
+
+\documentclass[11pt,a4paper]{article}
+\usepackage{graphicx,isabelle,isabellesym,latexsym}
+\usepackage[only,bigsqcap]{stmaryrd}
+\usepackage[latin1]{inputenc}
+\usepackage{pdfsetup}
+
+\urlstyle{rm}
+\isabellestyle{it}
+\pagestyle{myheadings}
+
+\begin{document}
+
+\title{Multivariate Analysis}
+\maketitle
+
+\tableofcontents
+
+\begin{center}
+ \includegraphics[scale=0.45]{session_graph}
+\end{center}
+
+\renewcommand{\isamarkupheader}[1]%
+{\section{\isabellecontext: #1}\markright{THEORY~``\isabellecontext''}}
+
+\parindent 0pt\parskip 0.5ex
+\input{session}
+
+\end{document}
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Aug 23 16:47:57 2010 +0200
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Tue Aug 24 08:22:17 2010 +0200
@@ -3,7 +3,7 @@
begin
lemma finite_information_spaceI:
- "\<lbrakk> finite_measure_space M ; measure M (space M) = 1 ; 1 < b \<rbrakk> \<Longrightarrow> finite_information_space M b"
+ "\<lbrakk> finite_measure_space M \<mu> ; \<mu> (space M) = 1 ; 1 < b \<rbrakk> \<Longrightarrow> finite_information_space M \<mu> b"
unfolding finite_information_space_def finite_measure_space_def finite_measure_space_axioms_def
finite_prob_space_def prob_space_def prob_space_axioms_def finite_information_space_axioms_def
by auto
@@ -13,32 +13,26 @@
assumes finite[simp]: "finite S"
and not_empty[simp]: "S \<noteq> {}"
-definition (in finite_space) "M = \<lparr> space = S, sets = Pow S, measure = (\<lambda>s. real (card s) / real (card S)) \<rparr>"
+definition (in finite_space) "M = \<lparr> space = S, sets = Pow S \<rparr>"
+definition (in finite_space) \<mu>_def[simp]: "\<mu> A = (of_nat (card A) / of_nat (card S) :: pinfreal)"
lemma (in finite_space)
shows space_M[simp]: "space M = S"
and sets_M[simp]: "sets M = Pow S"
- and measure_M[simp]: "measure M s = real (card s) / real (card S)"
by (simp_all add: M_def)
-sublocale finite_space \<subseteq> finite_information_space M 2
+sublocale finite_space \<subseteq> finite_information_space M \<mu> 2
proof (rule finite_information_spaceI)
let ?measure = "\<lambda>s::'a set. real (card s) / real (card S)"
- show "finite_measure_space M"
+ show "finite_measure_space M \<mu>"
proof (rule finite_Pow_additivity_sufficient, simp_all)
- show "positive M (measure M)"
- by (simp add: positive_def le_divide_eq)
+ show "positive \<mu>" by (simp add: positive_def)
- show "additive M (measure M)"
- proof (simp add: additive_def, safe)
- fix x y assume "x \<subseteq> S" and "y \<subseteq> S" and "x \<inter> y = {}"
- with this(1,2)[THEN finite_subset]
- have "card (x \<union> y) = card x + card y"
- by (simp add: card_Un_disjoint)
- thus "?measure (x \<union> y) = ?measure x + ?measure y"
- by (cases "card S = 0") (simp_all add: field_simps)
- qed
+ show "additive M \<mu>"
+ by (simp add: additive_def inverse_eq_divide field_simps Real_real
+ divide_le_0_iff zero_le_divide_iff
+ card_Un_disjoint finite_subset[OF _ finite])
qed
qed simp_all
@@ -508,7 +502,7 @@
let ?dI = "distribution inversion"
{ have "\<H>(inversion|payer) =
- - (\<Sum>x\<in>inversion`dc_crypto. (\<Sum>z\<in>Some ` {0..<n}. ?dIP {(x, z)} * log 2 (?dIP {(x, z)} / ?dP {z})))"
+ - (\<Sum>x\<in>inversion`dc_crypto. (\<Sum>z\<in>Some ` {0..<n}. real (?dIP {(x, z)}) * log 2 (real (?dIP {(x, z)}) / real (?dP {z}))))"
unfolding conditional_entropy_eq
by (simp add: image_payer_dc_crypto setsum_Sigma)
also have "... =
@@ -522,18 +516,20 @@
moreover from x z obtain i where "z = Some i" and "i < n" by auto
moreover from x have "length x = n" by (auto simp: inversion_def_raw dc_crypto)
ultimately
- have "?dIP {(x, z)} = 2 / (real n * 2^n)" using x
- by (simp add: distribution_def card_dc_crypto card_payer_and_inversion)
+ have "real (?dIP {(x, z)}) = 2 / (real n * 2^n)" using x
+ by (simp add: distribution_def card_dc_crypto card_payer_and_inversion
+ inverse_eq_divide mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
moreover
from z have "payer -` {z} \<inter> dc_crypto = {z} \<times> {xs. length xs = n}"
by (auto simp: dc_crypto payer_def)
hence "card (payer -` {z} \<inter> dc_crypto) = 2^n"
using card_list_length[where A="UNIV::bool set"]
by (simp add: card_cartesian_product_singleton)
- hence "?dP {z} = 1 / real n"
- by (simp add: distribution_def card_dc_crypto)
+ hence "real (?dP {z}) = 1 / real n"
+ by (simp add: distribution_def card_dc_crypto field_simps inverse_eq_divide
+ mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
ultimately
- show "?dIP {(x,z)} * log 2 (?dIP {(x,z)} / ?dP {z}) =
+ show "real (?dIP {(x,z)}) * log 2 (real (?dIP {(x,z)}) / real (?dP {z})) =
2 / (real n * 2^n) * (1 - real n)"
by (simp add: field_simps log_divide log_nat_power[of 2])
qed
@@ -542,7 +538,7 @@
by (simp add: card_image_inversion card_image[OF inj_Some] field_simps real_eq_of_nat[symmetric])
finally have "\<H>(inversion|payer) = real n - 1" . }
moreover
- { have "\<H>(inversion) = - (\<Sum>x \<in> inversion`dc_crypto. ?dI {x} * log 2 (?dI {x}))"
+ { have "\<H>(inversion) = - (\<Sum>x \<in> inversion`dc_crypto. real (?dI {x}) * log 2 (real (?dI {x})))"
unfolding entropy_eq by simp
also have "... = - (\<Sum>x \<in> inversion`dc_crypto. 2 * (1 - real n) / 2^n)"
unfolding neg_equal_iff_equal
@@ -551,9 +547,10 @@
hence "length x = n" by (auto simp: inversion_def_raw dc_crypto)
moreover have "inversion -` {x} \<inter> dc_crypto = {dc \<in> dc_crypto. inversion dc = x}" by auto
ultimately have "?dI {x} = 2 / 2^n" using `0 < n`
- by (simp add: distribution_def card_inversion[OF x_inv] card_dc_crypto)
- thus "?dI {x} * log 2 (?dI {x}) = 2 * (1 - real n) / 2^n"
- by (simp add: log_divide log_nat_power)
+ by (simp add: distribution_def card_inversion[OF x_inv] card_dc_crypto
+ mult_le_0_iff zero_le_mult_iff power_le_zero_eq)
+ thus "real (?dI {x}) * log 2 (real (?dI {x})) = 2 * (1 - real n) / 2^n"
+ by (simp add: log_divide log_nat_power power_le_zero_eq inverse_eq_divide)
qed
also have "... = real n - 1"
by (simp add: card_image_inversion real_of_nat_def[symmetric] field_simps)
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/Concurrent/simple_thread.scala Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,37 @@
+/* Title: Pure/Concurrent/simple_thread.scala
+ Author: Makarius
+
+Simplified thread operations.
+*/
+
+package isabelle
+
+
+import java.lang.Thread
+
+import scala.actors.Actor
+
+
+object Simple_Thread
+{
+ /* plain thread */
+
+ def fork(name: String, daemon: Boolean = false)(body: => Unit): Thread =
+ {
+ val thread = new Thread(name) { override def run = body }
+ thread.setDaemon(daemon)
+ thread.start
+ thread
+ }
+
+
+ /* thread as actor */
+
+ def actor(name: String, daemon: Boolean = false)(body: => Unit): Actor =
+ {
+ val actor = Future.promise[Actor]
+ val thread = fork(name, daemon) { actor.fulfill(Actor.self); body }
+ actor.join
+ }
+}
+
--- a/src/Pure/General/table.ML Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/General/table.ML Tue Aug 24 08:22:17 2010 +0200
@@ -392,6 +392,16 @@
fun merge_list eq = join (fn _ => Library.merge eq);
+(* ML pretty-printing -- requires Poly/ML 5.3.0 or later *)
+
+val _ =
+ PolyML.addPrettyPrinter (fn depth => fn pretty => fn tab =>
+ ml_pretty
+ (ML_Pretty.enum "," "{" "}"
+ (ML_Pretty.pair (pretty_ml o PolyML.prettyRepresentation) (pretty_ml o pretty))
+ (dest tab, depth)));
+
+
(*final declarations of this structure!*)
fun map f = map_table (K f);
val map' = map_table;
--- a/src/Pure/IsaMakefile Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/IsaMakefile Tue Aug 24 08:22:17 2010 +0200
@@ -32,6 +32,7 @@
ML-Systems/polyml-5.2.ML \
ML-Systems/polyml.ML \
ML-Systems/polyml_common.ML \
+ ML-Systems/pp_dummy.ML \
ML-Systems/pp_polyml.ML \
ML-Systems/proper_int.ML \
ML-Systems/single_assignment.ML \
--- a/src/Pure/ML-Systems/ml_pretty.ML Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/ML-Systems/ml_pretty.ML Tue Aug 24 08:22:17 2010 +0200
@@ -1,7 +1,7 @@
(* Title: Pure/ML-Systems/ml_pretty.ML
Author: Makarius
-Raw datatype for ML pretty printing.
+Minimal support for raw ML pretty printing -- for boot-strapping only.
*)
structure ML_Pretty =
@@ -12,5 +12,20 @@
String of string * int |
Break of bool * int;
+fun block prts = Block (("", ""), prts, 2);
+fun str s = String (s, size s);
+fun brk wd = Break (false, wd);
+
+fun pair pretty1 pretty2 ((x, y), depth: int) =
+ block [str "(", pretty1 (x, depth), str ",", brk 1, pretty2 (y, depth - 1), str ")"];
+
+fun enum sep lpar rpar pretty (args, depth) =
+ let
+ fun elems _ [] = []
+ | elems 0 _ = [str "..."]
+ | elems d [x] = [pretty (x, d)]
+ | elems d (x :: xs) = pretty (x, d) :: str sep :: brk 1 :: elems (d - 1) xs;
+ in block (str lpar :: (elems (Int.max (depth, 0)) args @ [str rpar])) end;
+
end;
--- a/src/Pure/ML-Systems/polyml-5.2.1.ML Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/ML-Systems/polyml-5.2.1.ML Tue Aug 24 08:22:17 2010 +0200
@@ -25,4 +25,5 @@
use "ML-Systems/compiler_polyml-5.2.ML";
use "ML-Systems/pp_polyml.ML";
+use "ML-Systems/pp_dummy.ML";
--- a/src/Pure/ML-Systems/polyml-5.2.ML Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/ML-Systems/polyml-5.2.ML Tue Aug 24 08:22:17 2010 +0200
@@ -27,4 +27,5 @@
use "ML-Systems/compiler_polyml-5.2.ML";
use "ML-Systems/pp_polyml.ML";
+use "ML-Systems/pp_dummy.ML";
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Pure/ML-Systems/pp_dummy.ML Tue Aug 24 08:22:17 2010 +0200
@@ -0,0 +1,16 @@
+(* Title: Pure/ML-Systems/pp_dummy.ML
+
+Dummy setup for toplevel pretty printing.
+*)
+
+fun ml_pretty _ = raise Fail "ml_pretty dummy";
+fun pretty_ml _ = raise Fail "pretty_ml dummy";
+
+structure PolyML =
+struct
+ fun addPrettyPrinter _ = ();
+ fun prettyRepresentation _ =
+ raise Fail "PolyML.prettyRepresentation dummy";
+ open PolyML;
+end;
+
--- a/src/Pure/ML-Systems/smlnj.ML Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/ML-Systems/smlnj.ML Tue Aug 24 08:22:17 2010 +0200
@@ -18,6 +18,8 @@
use "ML-Systems/bash.ML";
use "ML-Systems/ml_name_space.ML";
use "ML-Systems/ml_pretty.ML";
+structure PolyML = struct end;
+use "ML-Systems/pp_dummy.ML";
use "ML-Systems/use_context.ML";
--- a/src/Pure/PIDE/command.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/PIDE/command.scala Tue Aug 24 08:22:17 2010 +0200
@@ -8,10 +8,6 @@
package isabelle
-import scala.actors.Actor, Actor._
-import scala.collection.mutable
-
-
object Command
{
/** accumulated results from prover **/
@@ -40,7 +36,7 @@
def accumulate(message: XML.Tree): Command.State =
message match {
- case XML.Elem(Markup(Markup.STATUS, _), body) => // FIXME explicit checks!?
+ case XML.Elem(Markup(Markup.STATUS, _), body) => // FIXME explicit body check!?
copy(status = (for (XML.Elem(markup, _) <- body) yield markup) ::: status)
case XML.Elem(Markup(Markup.REPORT, _), msgs) =>
--- a/src/Pure/ROOT.ML Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/ROOT.ML Tue Aug 24 08:22:17 2010 +0200
@@ -179,7 +179,8 @@
use "ML/ml_syntax.ML";
use "ML/ml_env.ML";
use "Isar/runtime.ML";
-if member (op =) ["polyml-5.2", "polyml-5.2.1", "smlnj-110"] ml_system
+if ml_system = "polyml-5.2" orelse ml_system = "polyml-5.2.1" orelse
+ String.isPrefix "smlnj" ml_system
then use "ML/ml_compiler.ML"
else use "ML/ml_compiler_polyml-5.3.ML";
use "ML/ml_context.ML";
--- a/src/Pure/System/isabelle_process.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/System/isabelle_process.scala Tue Aug 24 08:22:17 2010 +0200
@@ -155,7 +155,7 @@
/* raw stdin */
private def stdin_actor(name: String, stream: OutputStream): Actor =
- Library.thread_actor(name) {
+ Simple_Thread.actor(name) {
val writer = new BufferedWriter(new OutputStreamWriter(stream, Standard_System.charset))
var finished = false
while (!finished) {
@@ -184,7 +184,7 @@
/* raw stdout */
private def stdout_actor(name: String, stream: InputStream): Actor =
- Library.thread_actor(name) {
+ Simple_Thread.actor(name) {
val reader = new BufferedReader(new InputStreamReader(stream, Standard_System.charset))
var result = new StringBuilder(100)
@@ -221,7 +221,7 @@
/* command input */
private def input_actor(name: String, raw_stream: OutputStream): Actor =
- Library.thread_actor(name) {
+ Simple_Thread.actor(name) {
val stream = new BufferedOutputStream(raw_stream)
var finished = false
while (!finished) {
@@ -253,7 +253,7 @@
private class Protocol_Error(msg: String) extends Exception(msg)
private def message_actor(name: String, stream: InputStream): Actor =
- Library.thread_actor(name) {
+ Simple_Thread.actor(name) {
val default_buffer = new Array[Byte](65536)
var c = -1
@@ -344,7 +344,7 @@
/* exit process */
- Library.thread_actor("process_exit") {
+ Simple_Thread.actor("process_exit") {
proc match {
case None =>
case Some(p) =>
--- a/src/Pure/System/session.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/System/session.scala Tue Aug 24 08:22:17 2010 +0200
@@ -69,8 +69,8 @@
private case class Started(timeout: Int, args: List[String])
private case object Stop
- private lazy val session_actor = actor {
-
+ private val session_actor = Simple_Thread.actor("session_actor", daemon = true)
+ {
var prover: Isabelle_Process with Isar_Document = null
@@ -199,8 +199,9 @@
/* main loop */
- loop {
- react {
+ var finished = false
+ while (!finished) {
+ receive {
case Started(timeout, args) =>
if (prover == null) {
prover = new Isabelle_Process(system, self, args:_*) with Isar_Document
@@ -211,10 +212,11 @@
}
else reply(None)
- case Stop => // FIXME clarify; synchronous
+ case Stop => // FIXME synchronous!?
if (prover != null) {
prover.kill
prover = null
+ finished = true
}
case change: Document.Change if prover != null =>
@@ -235,7 +237,7 @@
/** buffered command changes (delay_first discipline) **/
- private lazy val command_change_buffer = actor
+ private val command_change_buffer = actor
//{{{
{
import scala.compat.Platform.currentTime
@@ -286,36 +288,33 @@
private case class Edit_Version(edits: List[Document.Node_Text_Edit])
- private val editor_history = new Actor
- {
- @volatile private var history = Document.History.init
+ @volatile private var history = Document.History.init
- def snapshot(name: String, pending_edits: List[Text.Edit]): Document.Snapshot =
- history.snapshot(name, pending_edits, current_state())
+ def snapshot(name: String, pending_edits: List[Text.Edit]): Document.Snapshot =
+ history.snapshot(name, pending_edits, current_state())
- def act
- {
- loop {
- react {
- case Edit_Version(edits) =>
- val prev = history.tip.current
- val result =
- isabelle.Future.fork {
- val previous = prev.join
- val former_assignment = current_state().the_assignment(previous).join // FIXME async!?
- Thy_Syntax.text_edits(Session.this, previous, edits)
- }
- val change = new Document.Change(prev, edits, result)
- history += change
- change.current.map(_ => session_actor ! change)
- reply(())
+ private val editor_history = actor
+ {
+ loop {
+ react {
+ case Edit_Version(edits) =>
+ val prev = history.tip.current
+ val result =
+ // FIXME potential denial-of-service concerning worker pool (!?!?)
+ isabelle.Future.fork {
+ val previous = prev.join
+ val former_assignment = current_state().the_assignment(previous).join // FIXME async!?
+ Thy_Syntax.text_edits(Session.this, previous, edits)
+ }
+ val change = new Document.Change(prev, edits, result)
+ history += change
+ change.current.map(_ => session_actor ! change)
+ reply(())
- case bad => System.err.println("editor_model: ignoring bad message " + bad)
- }
+ case bad => System.err.println("editor_model: ignoring bad message " + bad)
}
}
}
- editor_history.start
@@ -326,8 +325,5 @@
def stop() { session_actor ! Stop }
- def snapshot(name: String, pending_edits: List[Text.Edit]): Document.Snapshot =
- editor_history.snapshot(name, pending_edits)
-
def edit_version(edits: List[Document.Node_Text_Edit]) { editor_history !? Edit_Version(edits) }
}
--- a/src/Pure/build-jars Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/build-jars Tue Aug 24 08:22:17 2010 +0200
@@ -23,6 +23,7 @@
declare -a SOURCES=(
Concurrent/future.scala
+ Concurrent/simple_thread.scala
General/exn.scala
General/linear_set.scala
General/markup.scala
--- a/src/Pure/library.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Pure/library.scala Tue Aug 24 08:22:17 2010 +0200
@@ -7,11 +7,10 @@
package isabelle
-import java.lang.{System, Thread}
+import java.lang.System
import java.awt.Component
import javax.swing.JOptionPane
-import scala.actors.Actor
import scala.swing.ComboBox
import scala.swing.event.SelectionChanged
@@ -139,15 +138,4 @@
((stop - start).toDouble / 1000000) + "ms elapsed time")
Exn.release(result)
}
-
-
- /* thread as actor */
-
- def thread_actor(name: String)(body: => Unit): Actor =
- {
- val actor = Future.promise[Actor]
- val thread = new Thread(name) { override def run() = { actor.fulfill(Actor.self); body } }
- thread.start
- actor.join
- }
}
--- a/src/Tools/jEdit/src/jedit/document_model.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/document_model.scala Tue Aug 24 08:22:17 2010 +0200
@@ -10,7 +10,6 @@
import isabelle._
-import scala.actors.Actor, Actor._
import scala.collection.mutable
import org.gjt.sp.jedit.Buffer
@@ -260,61 +259,63 @@
// FIXME proper synchronization / thread context (!??)
val snapshot = Swing_Thread.now { Document_Model.this.snapshot() }
- val previous = prev.asInstanceOf[Document_Model.Token_Markup.LineContext]
- val line = if (prev == null) 0 else previous.line + 1
- val context = new Document_Model.Token_Markup.LineContext(line, previous)
+ Isabelle.buffer_read_lock(buffer) {
+ val previous = prev.asInstanceOf[Document_Model.Token_Markup.LineContext]
+ val line = if (prev == null) 0 else previous.line + 1
+ val context = new Document_Model.Token_Markup.LineContext(line, previous)
- val start = buffer.getLineStartOffset(line)
- val stop = start + line_segment.count
- val range = Text.Range(start, stop)
- val former_range = snapshot.revert(range)
+ val start = buffer.getLineStartOffset(line)
+ val stop = start + line_segment.count
+ val range = Text.Range(start, stop)
+ val former_range = snapshot.revert(range)
- /* FIXME
- for (text_area <- Isabelle.jedit_text_areas(buffer)
- if Document_View(text_area).isDefined)
- Document_View(text_area).get.set_styles()
- */
+ /* FIXME
+ for (text_area <- Isabelle.jedit_text_areas(buffer)
+ if Document_View(text_area).isDefined)
+ Document_View(text_area).get.set_styles()
+ */
- def handle_token(style: Byte, offset: Text.Offset, length: Int) =
- handler.handleToken(line_segment, style, offset, length, context)
+ def handle_token(style: Byte, offset: Text.Offset, length: Int) =
+ handler.handleToken(line_segment, style, offset, length, context)
- val syntax = session.current_syntax()
- val token_markup: PartialFunction[Text.Info[Any], Byte] =
- {
- case Text.Info(_, XML.Elem(Markup(Markup.COMMAND, List((Markup.NAME, name))), _))
- if syntax.keyword_kind(name).isDefined =>
- Document_Model.Token_Markup.command_style(syntax.keyword_kind(name).get)
+ val syntax = session.current_syntax()
+ val token_markup: PartialFunction[Text.Info[Any], Byte] =
+ {
+ case Text.Info(_, XML.Elem(Markup(Markup.COMMAND, List((Markup.NAME, name))), _))
+ if syntax.keyword_kind(name).isDefined =>
+ Document_Model.Token_Markup.command_style(syntax.keyword_kind(name).get)
- case Text.Info(_, XML.Elem(Markup(name, _), _))
- if Document_Model.Token_Markup.token_style(name) != Token.NULL =>
- Document_Model.Token_Markup.token_style(name)
- }
+ case Text.Info(_, XML.Elem(Markup(name, _), _))
+ if Document_Model.Token_Markup.token_style(name) != Token.NULL =>
+ Document_Model.Token_Markup.token_style(name)
+ }
- var next_x = start
- for {
- (command, command_start) <- snapshot.node.command_range(former_range)
- info <- snapshot.state(command).markup.
- select((former_range - command_start).restrict(command.range))(token_markup)(Token.NULL)
- val Text.Range(abs_start, abs_stop) = snapshot.convert(info.range + command_start)
- if abs_stop > start && abs_start < stop // FIXME abs_range overlaps range (redundant!?)
+ var next_x = start
+ for {
+ (command, command_start) <- snapshot.node.command_range(former_range)
+ info <- snapshot.state(command).markup.
+ select((former_range - command_start).restrict(command.range))(token_markup)(Token.NULL)
+ val Text.Range(abs_start, abs_stop) = snapshot.convert(info.range + command_start)
+ if abs_stop > start && abs_start < stop // FIXME abs_range overlaps range (redundant!?)
+ }
+ {
+ val token_start = (abs_start - start) max 0
+ val token_length =
+ (abs_stop - abs_start) -
+ ((start - abs_start) max 0) -
+ ((abs_stop - stop) max 0)
+ if (start + token_start > next_x) // FIXME ??
+ handle_token(Token.COMMENT1, next_x - start, start + token_start - next_x)
+ handle_token(info.info, token_start, token_length)
+ next_x = start + token_start + token_length
+ }
+ if (next_x < stop) // FIXME ??
+ handle_token(Token.COMMENT1, next_x - start, stop - next_x)
+
+ handle_token(Token.END, line_segment.count, 0)
+ handler.setLineContext(context)
+ context
}
- {
- val token_start = (abs_start - start) max 0
- val token_length =
- (abs_stop - abs_start) -
- ((start - abs_start) max 0) -
- ((abs_stop - stop) max 0)
- if (start + token_start > next_x) // FIXME ??
- handle_token(Token.COMMENT1, next_x - start, start + token_start - next_x)
- handle_token(info.info, token_start, token_length)
- next_x = start + token_start + token_length
- }
- if (next_x < stop) // FIXME ??
- handle_token(Token.COMMENT1, next_x - start, stop - next_x)
-
- handle_token(Token.END, line_segment.count, 0)
- handler.setLineContext(context)
- context
}
}
--- a/src/Tools/jEdit/src/jedit/document_view.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/document_view.scala Tue Aug 24 08:22:17 2010 +0200
@@ -106,8 +106,27 @@
Swing_Thread.now {
// FIXME cover doc states as well!!?
val snapshot = model.snapshot()
- if (changed.exists(snapshot.node.commands.contains))
- full_repaint(snapshot, changed)
+ val buffer = model.buffer
+ Isabelle.buffer_read_lock(buffer) {
+ if (changed.exists(snapshot.node.commands.contains)) {
+ var visible_change = false
+
+ for ((command, start) <- snapshot.node.command_starts) {
+ if (changed(command)) {
+ val stop = start + command.length
+ val line1 = buffer.getLineOfOffset(snapshot.convert(start))
+ val line2 = buffer.getLineOfOffset(snapshot.convert(stop))
+ if (line2 >= text_area.getFirstLine &&
+ line1 <= text_area.getFirstLine + text_area.getVisibleLines)
+ visible_change = true
+ text_area.invalidateLineRange(line1, line2)
+ }
+ }
+ if (visible_change) model.buffer.propertiesChanged()
+
+ overview.repaint() // FIXME paint here!?
+ }
+ }
}
case bad => System.err.println("command_change_actor: ignoring bad message " + bad)
@@ -115,29 +134,6 @@
}
}
- def full_repaint(snapshot: Document.Snapshot, changed: Set[Command])
- {
- Swing_Thread.require()
-
- val buffer = model.buffer
- var visible_change = false
-
- for ((command, start) <- snapshot.node.command_starts) {
- if (changed(command)) {
- val stop = start + command.length
- val line1 = buffer.getLineOfOffset(snapshot.convert(start))
- val line2 = buffer.getLineOfOffset(snapshot.convert(stop))
- if (line2 >= text_area.getFirstLine &&
- line1 <= text_area.getFirstLine + text_area.getVisibleLines)
- visible_change = true
- text_area.invalidateLineRange(line1, line2)
- }
- }
- if (visible_change) model.buffer.propertiesChanged()
-
- overview.repaint() // FIXME paint here!?
- }
-
/* text_area_extension */
@@ -269,20 +265,23 @@
override def paintComponent(gfx: Graphics)
{
super.paintComponent(gfx)
+ Swing_Thread.assert()
val buffer = model.buffer
- val snapshot = model.snapshot()
- val saved_color = gfx.getColor // FIXME needed!?
- try {
- for ((command, start) <- snapshot.node.command_starts if !command.is_ignored) {
- val line1 = buffer.getLineOfOffset(snapshot.convert(start))
- val line2 = buffer.getLineOfOffset(snapshot.convert(start + command.length)) + 1
- val y = line_to_y(line1)
- val height = HEIGHT * (line2 - line1)
- gfx.setColor(Document_View.choose_color(snapshot, command))
- gfx.fillRect(0, y, getWidth - 1, height)
+ Isabelle.buffer_read_lock(buffer) {
+ val snapshot = model.snapshot()
+ val saved_color = gfx.getColor // FIXME needed!?
+ try {
+ for ((command, start) <- snapshot.node.command_starts if !command.is_ignored) {
+ val line1 = buffer.getLineOfOffset(snapshot.convert(start))
+ val line2 = buffer.getLineOfOffset(snapshot.convert(start + command.length)) + 1
+ val y = line_to_y(line1)
+ val height = HEIGHT * (line2 - line1)
+ gfx.setColor(Document_View.choose_color(snapshot, command))
+ gfx.fillRect(0, y, getWidth - 1, height)
+ }
}
+ finally { gfx.setColor(saved_color) }
}
- finally { gfx.setColor(saved_color) }
}
private def line_to_y(line: Int): Int =
--- a/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/isabelle_hyperlinks.scala Tue Aug 24 08:22:17 2010 +0200
@@ -40,6 +40,7 @@
{
def getHyperlink(buffer: Buffer, buffer_offset: Int): Hyperlink =
{
+ // FIXME lock buffer (!??)
Swing_Thread.assert()
Document_Model(buffer) match {
case Some(model) =>
--- a/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/isabelle_sidekick.scala Tue Aug 24 08:22:17 2010 +0200
@@ -44,7 +44,7 @@
stopped = false
- // FIXME lock buffer !??
+ // FIXME lock buffer (!??)
val data = new SideKickParsedData(buffer.getName)
val root = data.root
data.getAsset(root).setEnd(buffer.getLength)
@@ -66,6 +66,7 @@
override def complete(pane: EditPane, caret: Int): SideKickCompletion =
{
+ // FIXME lock buffer (!?)
val buffer = pane.getBuffer
val line = buffer.getLineOfOffset(caret)
--- a/src/Tools/jEdit/src/jedit/plugin.scala Mon Aug 23 16:47:57 2010 +0200
+++ b/src/Tools/jEdit/src/jedit/plugin.scala Tue Aug 24 08:22:17 2010 +0200
@@ -118,6 +118,12 @@
def jedit_text_areas(buffer: JEditBuffer): Iterator[JEditTextArea] =
jedit_text_areas().filter(_.getBuffer == buffer)
+ def buffer_read_lock[A](buffer: JEditBuffer)(body: => A): A =
+ {
+ try { buffer.readLock(); body }
+ finally { buffer.readUnlock() }
+ }
+
/* dockable windows */