added projective_family; generalized generator in product_prob_space to projective_family
--- a/src/HOL/Probability/Infinite_Product_Measure.thy Tue Nov 06 11:03:28 2012 +0100
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy Wed Nov 07 11:33:27 2012 +0100
@@ -5,7 +5,7 @@
header {*Infinite Product Measure*}
theory Infinite_Product_Measure
- imports Probability_Measure Caratheodory
+ imports Probability_Measure Caratheodory Projective_Family
begin
lemma split_merge: "P (merge I J (x,y) i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)"
@@ -76,16 +76,20 @@
by (auto intro!: measurable_restrict_subset simp: space_PiM)
qed
-abbreviation (in product_prob_space)
- "emb L K X \<equiv> prod_emb L M K X"
-
lemma (in product_prob_space) emeasure_prod_emb[simp]:
assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^isub>M J M)"
- shows "emeasure (Pi\<^isub>M L M) (emb L J X) = emeasure (Pi\<^isub>M J M) X"
+ shows "emeasure (Pi\<^isub>M L M) (prod_emb L M J X) = emeasure (Pi\<^isub>M J M) X"
by (subst distr_restrict[OF L])
(simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X)
-lemma (in product_prob_space) prod_emb_injective:
+sublocale product_prob_space \<subseteq> projective_family I "\<lambda>J. PiM J M" M
+proof
+ fix J::"'i set" assume "finite J"
+ interpret f: finite_product_prob_space M J proof qed fact
+ show "emeasure (Pi\<^isub>M J M) (space (Pi\<^isub>M J M)) \<noteq> \<infinity>" by simp
+qed simp_all
+
+lemma (in projective_family) prod_emb_injective:
assumes "J \<noteq> {}" "J \<subseteq> L" "finite J" and sets: "X \<in> sets (Pi\<^isub>M J M)" "Y \<in> sets (Pi\<^isub>M J M)"
assumes "prod_emb L M J X = prod_emb L M J Y"
shows "X = Y"
@@ -100,14 +104,17 @@
using `prod_emb L M J X = prod_emb L M J Y` by (simp add: prod_emb_def)
qed fact
-definition (in product_prob_space) generator :: "('i \<Rightarrow> 'a) set set" where
+abbreviation (in projective_family)
+ "emb L K X \<equiv> prod_emb L M K X"
+
+definition (in projective_family) generator :: "('i \<Rightarrow> 'a) set set" where
"generator = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M))"
-lemma (in product_prob_space) generatorI':
+lemma (in projective_family) generatorI':
"J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> emb I J X \<in> generator"
unfolding generator_def by auto
-lemma (in product_prob_space) algebra_generator:
+lemma (in projective_family) algebra_generator:
assumes "I \<noteq> {}" shows "algebra (\<Pi>\<^isub>E i\<in>I. space (M i)) generator" (is "algebra ?\<Omega> ?G")
unfolding algebra_def algebra_axioms_def ring_of_sets_iff
proof (intro conjI ballI)
@@ -135,7 +142,7 @@
unfolding * using XA XB by (safe intro!: generatorI') auto
qed
-lemma (in product_prob_space) sets_PiM_generator:
+lemma (in projective_family) sets_PiM_generator:
"sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
proof cases
assume "I = {}" then show ?thesis
@@ -154,57 +161,56 @@
qed (auto simp: generator_def space_PiM[symmetric] intro!: sigma_sets_subset)
qed
-
-lemma (in product_prob_space) generatorI:
+lemma (in projective_family) generatorI:
"J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> generator"
unfolding generator_def by auto
-definition (in product_prob_space)
+definition (in projective_family)
"\<mu>G A =
- (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^isub>M J M). A = emb I J X \<longrightarrow> x = emeasure (Pi\<^isub>M J M) X))"
+ (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^isub>M J M). A = emb I J X \<longrightarrow> x = emeasure (PiP J M P) X))"
-lemma (in product_prob_space) \<mu>G_spec:
+lemma (in projective_family) \<mu>G_spec:
assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
- shows "\<mu>G A = emeasure (Pi\<^isub>M J M) X"
+ shows "\<mu>G A = emeasure (PiP J M P) X"
unfolding \<mu>G_def
proof (intro the_equality allI impI ballI)
fix K Y assume K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "A = emb I K Y" "Y \<in> sets (Pi\<^isub>M K M)"
- have "emeasure (Pi\<^isub>M K M) Y = emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) K Y)"
+ have "emeasure (PiP K M P) Y = emeasure (PiP (K \<union> J) M P) (emb (K \<union> J) K Y)"
using K J by simp
also have "emb (K \<union> J) K Y = emb (K \<union> J) J X"
using K J by (simp add: prod_emb_injective[of "K \<union> J" I])
- also have "emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) J X) = emeasure (Pi\<^isub>M J M) X"
+ also have "emeasure (PiP (K \<union> J) M P) (emb (K \<union> J) J X) = emeasure (PiP J M P) X"
using K J by simp
- finally show "emeasure (Pi\<^isub>M J M) X = emeasure (Pi\<^isub>M K M) Y" ..
+ finally show "emeasure (PiP J M P) X = emeasure (PiP K M P) Y" ..
qed (insert J, force)
-lemma (in product_prob_space) \<mu>G_eq:
- "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = emeasure (Pi\<^isub>M J M) X"
+lemma (in projective_family) \<mu>G_eq:
+ "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = emeasure (PiP J M P) X"
by (intro \<mu>G_spec) auto
-lemma (in product_prob_space) generator_Ex:
+lemma (in projective_family) generator_Ex:
assumes *: "A \<in> generator"
- shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A = emb I J X \<and> \<mu>G A = emeasure (Pi\<^isub>M J M) X"
+ shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A = emb I J X \<and> \<mu>G A = emeasure (PiP J M P) X"
proof -
from * obtain J X where J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
unfolding generator_def by auto
with \<mu>G_spec[OF this] show ?thesis by auto
qed
-lemma (in product_prob_space) generatorE:
+lemma (in projective_family) generatorE:
assumes A: "A \<in> generator"
- obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" "\<mu>G A = emeasure (Pi\<^isub>M J M) X"
+ obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" "\<mu>G A = emeasure (PiP J M P) X"
proof -
from generator_Ex[OF A] obtain X J where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A"
- "\<mu>G A = emeasure (Pi\<^isub>M J M) X" by auto
+ "\<mu>G A = emeasure (PiP J M P) X" by auto
then show thesis by (intro that) auto
qed
-lemma (in product_prob_space) merge_sets:
+lemma (in projective_family) merge_sets:
"J \<inter> K = {} \<Longrightarrow> A \<in> sets (Pi\<^isub>M (J \<union> K) M) \<Longrightarrow> x \<in> space (Pi\<^isub>M J M) \<Longrightarrow> (\<lambda>y. merge J K (x,y)) -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)"
by simp
-lemma (in product_prob_space) merge_emb:
+lemma (in projective_family) merge_emb:
assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^isub>M J M)"
shows "((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^isub>M I M)) =
emb I (K - J) ((\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M))"
@@ -221,7 +227,7 @@
auto
qed
-lemma (in product_prob_space) positive_\<mu>G:
+lemma (in projective_family) positive_\<mu>G:
assumes "I \<noteq> {}"
shows "positive generator \<mu>G"
proof -
@@ -241,7 +247,7 @@
qed
qed
-lemma (in product_prob_space) additive_\<mu>G:
+lemma (in projective_family) additive_\<mu>G:
assumes "I \<noteq> {}"
shows "additive generator \<mu>G"
proof -
@@ -263,7 +269,7 @@
using J K by simp_all
then have "\<mu>G (A \<union> B) = \<mu>G (emb I (J \<union> K) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y))"
by simp
- also have "\<dots> = emeasure (Pi\<^isub>M (J \<union> K) M) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"
+ also have "\<dots> = emeasure (PiP (J \<union> K) M P) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"
using JK J(1, 4) K(1, 4) by (simp add: \<mu>G_eq Un del: prod_emb_Un)
also have "\<dots> = \<mu>G A + \<mu>G B"
using J K JK_disj by (simp add: plus_emeasure[symmetric])
@@ -271,6 +277,10 @@
qed
qed
+lemma (in product_prob_space) PiP_PiM_finite[simp]:
+ assumes "J \<noteq> {}" "finite J" "J \<subseteq> I" shows "PiP J M (\<lambda>J. PiM J M) = PiM J M"
+ using assms by (simp add: PiP_finite)
+
lemma (in product_prob_space) emeasure_PiM_emb_not_empty:
assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)"
shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)"
@@ -295,7 +305,6 @@
ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^isub>M K M)" "Z = emb I K X"
"K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^isub>M K M) X"
by (auto simp: subset_insertI)
-
let ?M = "\<lambda>y. (\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)"
{ fix y assume y: "y \<in> space (Pi\<^isub>M J M)"
note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X]
@@ -360,7 +369,7 @@
using A positive_\<mu>G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def)
ultimately have "0 < ?a" by auto
- have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (Pi\<^isub>M J M) X"
+ have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (PiP J M (\<lambda>J. (Pi\<^isub>M J M))) X"
using A by (intro allI generator_Ex) auto
then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^isub>M (J' n) M)"
and A': "\<And>n. A n = emb I (J' n) (X' n)"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Projective_Family.thy Wed Nov 07 11:33:27 2012 +0100
@@ -0,0 +1,113 @@
+theory Projective_Family
+imports Finite_Product_Measure Probability_Measure
+begin
+
+definition
+ PiP :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i set \<Rightarrow> ('i \<Rightarrow> 'a) measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
+ "PiP I M P = extend_measure (\<Pi>\<^isub>E i\<in>I. space (M i))
+ {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
+ (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^isub>E j\<in>J. X j))
+ (\<lambda>(J, X). emeasure (P J) (Pi\<^isub>E J X))"
+
+lemma space_PiP[simp]: "space (PiP I M P) = space (PiM I M)"
+ by (auto simp add: PiP_def space_PiM prod_emb_def intro!: space_extend_measure)
+
+lemma sets_PiP[simp]: "sets (PiP I M P) = sets (PiM I M)"
+ by (auto simp add: PiP_def sets_PiM prod_algebra_def prod_emb_def intro!: sets_extend_measure)
+
+lemma measurable_PiP1[simp]: "measurable (PiP I M P) M' = measurable (\<Pi>\<^isub>M i\<in>I. M i) M'"
+ unfolding measurable_def by auto
+
+lemma measurable_PiP2[simp]: "measurable M' (PiP I M P) = measurable M' (\<Pi>\<^isub>M i\<in>I. M i)"
+ unfolding measurable_def by auto
+
+locale projective_family =
+ fixes I::"'i set" and P::"'i set \<Rightarrow> ('i \<Rightarrow> 'a) measure" and M::"('i \<Rightarrow> 'a measure)"
+ assumes projective: "\<And>J H X. J \<noteq> {} \<Longrightarrow> J \<subseteq> H \<Longrightarrow> H \<subseteq> I \<Longrightarrow> finite H \<Longrightarrow> X \<in> sets (PiM J M) \<Longrightarrow>
+ (P H) (prod_emb H M J X) = (P J) X"
+ assumes proj_space: "\<And>J. finite J \<Longrightarrow> space (P J) = space (PiM J M)"
+ assumes proj_sets: "\<And>J. finite J \<Longrightarrow> sets (P J) = sets (PiM J M)"
+ assumes proj_finite_measure: "\<And>J. finite J \<Longrightarrow> emeasure (P J) (space (PiM J M)) \<noteq> \<infinity>"
+ assumes prob_space: "\<And>i. prob_space (M i)"
+begin
+
+lemma emeasure_PiP:
+ assumes "J \<noteq> {}"
+ assumes "finite J"
+ assumes "J \<subseteq> I"
+ assumes A: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> sets (M i)"
+ shows "emeasure (PiP J M P) (Pi\<^isub>E J A) = emeasure (P J) (Pi\<^isub>E J A)"
+proof -
+ have "Pi\<^isub>E J (restrict A J) \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"
+ proof safe
+ fix x j assume "x \<in> Pi J (restrict A J)" "j \<in> J"
+ hence "x j \<in> restrict A J j" by (auto simp: Pi_def)
+ also have "\<dots> \<subseteq> space (M j)" using sets_into_space A `j \<in> J` by auto
+ finally show "x j \<in> space (M j)" .
+ qed
+ hence "emeasure (PiP J M P) (Pi\<^isub>E J A) =
+ emeasure (PiP J M P) (prod_emb J M J (Pi\<^isub>E J A))"
+ using assms(1-3) sets_into_space by (auto simp add: prod_emb_id Pi_def)
+ also have "\<dots> = emeasure (P J) (Pi\<^isub>E J A)"
+ proof (rule emeasure_extend_measure[OF PiP_def, where i="(J, A)", simplified,
+ of J M "P J" P])
+ show "positive (sets (PiM J M)) (P J)" unfolding positive_def by auto
+ show "countably_additive (sets (PiM J M)) (P J)" unfolding countably_additive_def
+ by (auto simp: suminf_emeasure proj_sets[OF `finite J`])
+ show "(\<lambda>(Ja, X). prod_emb J M Ja (Pi\<^isub>E Ja X)) ` {(Ja, X). (Ja = {} \<longrightarrow> J = {}) \<and>
+ finite Ja \<and> Ja \<subseteq> J \<and> X \<in> (\<Pi> j\<in>Ja. sets (M j))} \<subseteq> Pow (\<Pi> i\<in>J. space (M i)) \<and>
+ (\<lambda>(Ja, X). prod_emb J M Ja (Pi\<^isub>E Ja X)) `
+ {(Ja, X). (Ja = {} \<longrightarrow> J = {}) \<and> finite Ja \<and> Ja \<subseteq> J \<and> X \<in> (\<Pi> j\<in>Ja. sets (M j))} \<subseteq>
+ Pow (extensional J)" by (auto simp: prod_emb_def)
+ show "(J = {} \<longrightarrow> J = {}) \<and> finite J \<and> J \<subseteq> J \<and> A \<in> (\<Pi> j\<in>J. sets (M j))"
+ using assms by auto
+ fix i
+ assume
+ "case i of (Ja, X) \<Rightarrow> (Ja = {} \<longrightarrow> J = {}) \<and> finite Ja \<and> Ja \<subseteq> J \<and> X \<in> (\<Pi> j\<in>Ja. sets (M j))"
+ thus "emeasure (P J) (case i of (Ja, X) \<Rightarrow> prod_emb J M Ja (Pi\<^isub>E Ja X)) =
+ (case i of (J, X) \<Rightarrow> emeasure (P J) (Pi\<^isub>E J X))" using assms
+ by (cases i) (auto simp add: intro!: projective sets_PiM_I_finite)
+ qed
+ finally show ?thesis .
+qed
+
+lemma PiP_finite:
+ assumes "J \<noteq> {}"
+ assumes "finite J"
+ assumes "J \<subseteq> I"
+ shows "PiP J M P = P J" (is "?P = _")
+proof (rule measure_eqI_generator_eq)
+ let ?J = "{Pi\<^isub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
+ let ?F = "\<lambda>i. \<Pi>\<^isub>E k\<in>J. space (M k)"
+ let ?\<Omega> = "(\<Pi>\<^isub>E k\<in>J. space (M k))"
+ show "Int_stable ?J"
+ by (rule Int_stable_PiE)
+ interpret finite_measure "P J" using proj_finite_measure `finite J`
+ by (intro finite_measureI) (simp add: proj_space)
+ show "emeasure ?P (?F _) \<noteq> \<infinity>" using assms `finite J` by (auto simp: emeasure_PiP)
+ show "?J \<subseteq> Pow ?\<Omega>" by (auto simp: Pi_iff dest: sets_into_space)
+ show "sets (PiP J M P) = sigma_sets ?\<Omega> ?J" "sets (P J) = sigma_sets ?\<Omega> ?J"
+ using `finite J` proj_sets by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)
+ fix X assume "X \<in> ?J"
+ then obtain E where X: "X = Pi\<^isub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
+ with `finite J` have "X \<in> sets (PiP J M P)" by simp
+ have emb_self: "prod_emb J M J (Pi\<^isub>E J E) = Pi\<^isub>E J E"
+ using E sets_into_space
+ by (auto intro!: prod_emb_PiE_same_index)
+ show "emeasure (PiP J M P) X = emeasure (P J) X"
+ unfolding X using E
+ by (intro emeasure_PiP assms) simp
+qed (insert `finite J`, auto intro!: prod_algebraI_finite)
+
+lemma emeasure_fun_emb[simp]:
+ assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" "L \<subseteq> I" and X: "X \<in> sets (PiM J M)"
+ shows "emeasure (PiP L M P) (prod_emb L M J X) = emeasure (PiP J M P) X"
+ using assms
+ by (subst PiP_finite) (auto simp: PiP_finite finite_subset projective)
+
+end
+
+sublocale projective_family \<subseteq> M: prob_space "M i" for i
+ by (rule prob_space)
+
+end