# HG changeset patch # User hoelzl # Date 1353073583 -3600 # Node ID a3bede207a0488a088a080c3485ce795aa94ef5a # Parent 9af8721ecd20e6128c3896c365bc86dcb3dd465e renamed prob_space to proj_prob_space as it clashed with Probability_Measure.prob_space diff -r 9af8721ecd20 -r a3bede207a04 src/HOL/Probability/Projective_Family.thy --- a/src/HOL/Probability/Projective_Family.thy Fri Nov 16 14:46:23 2012 +0100 +++ b/src/HOL/Probability/Projective_Family.thy Fri Nov 16 14:46:23 2012 +0100 @@ -81,7 +81,7 @@ fixes I::"'i set" and P::"'i set \ ('i \ 'a) measure" and M::"('i \ 'a measure)" assumes projective: "\J H X. J \ {} \ J \ H \ H \ I \ finite H \ X \ sets (PiM J M) \ (P H) (prod_emb H M J X) = (P J) X" - assumes prob_space: "\J. finite J \ prob_space (P J)" + assumes proj_prob_space: "\J. finite J \ prob_space (P J)" assumes proj_space: "\J. finite J \ space (P J) = space (PiM J M)" assumes proj_sets: "\J. finite J \ sets (P J) = sets (PiM J M)" begin @@ -133,7 +133,7 @@ let ?\ = "(\\<^isub>E k\J. space (M k))" show "Int_stable ?J" by (rule Int_stable_PiE) - interpret prob_space "P J" using prob_space `finite J` by simp + interpret prob_space "P J" using proj_prob_space `finite J` by simp show "emeasure ?P (?F _) \ \" using assms `finite J` by (auto simp: emeasure_limP) show "?J \ Pow ?\" by (auto simp: Pi_iff dest: sets_into_space) show "sets (limP J M P) = sigma_sets ?\ ?J" "sets (P J) = sigma_sets ?\ ?J" @@ -165,7 +165,7 @@ have "\i\L. \x. x \ space (M i)" proof fix i assume "i \ L" - interpret prob_space "P {i}" using prob_space by simp + interpret prob_space "P {i}" using proj_prob_space by simp from not_empty show "\x. x \ space (M i)" by (auto simp add: proj_space space_PiM) qed from bchoice[OF this] diff -r 9af8721ecd20 -r a3bede207a04 src/HOL/Probability/Projective_Limit.thy --- a/src/HOL/Probability/Projective_Limit.thy Fri Nov 16 14:46:23 2012 +0100 +++ b/src/HOL/Probability/Projective_Limit.thy Fri Nov 16 14:46:23 2012 +0100 @@ -207,7 +207,7 @@ OF `I \ {}`, OF `I \ {}`]) fix A assume "A \ ?G" with generatorE guess J X . note JX = this - interpret prob_space "P J" using prob_space[OF `finite J`] . + interpret prob_space "P J" using proj_prob_space[OF `finite J`] . show "\G A \ \" using JX by (simp add: limP_finite) next fix Z assume Z: "range Z \ ?G" "decseq Z" "(\i. Z i) = {}" @@ -241,7 +241,7 @@ note [simp] = `\n. finite (J n)` from J Z_emb have Z_eq: "\n. Z n = emb I (J n) (B n)" "\n. Z n \ ?G" unfolding J_def B_def by (subst prod_emb_trans) (insert Z, auto) - interpret prob_space "P (J i)" for i using prob_space by simp + interpret prob_space "P (J i)" for i using proj_prob_space by simp have "?a \ \G (Z 0)" by (auto intro: INF_lower) also have "\ < \" using J by (auto simp: Z_eq \G_eq limP_finite proj_sets) finally have "?a \ \" by simp @@ -636,13 +636,13 @@ show "emeasure (lim\<^isub>B I P) (space (lim\<^isub>B I P)) = 1" proof cases assume "I = {}" - interpret prob_space "P {}" using prob_space by simp + interpret prob_space "P {}" using proj_prob_space by simp show ?thesis by (simp add: space_PiM_empty limP_finite emeasure_space_1 `I = {}`) next assume "I \ {}" then obtain i where "i \ I" by auto - interpret prob_space "P {i}" using prob_space by simp + interpret prob_space "P {i}" using proj_prob_space by simp have R: "(space (lim\<^isub>B I P)) = (emb I {i} (Pi\<^isub>E {i} (\_. space borel)))" by (auto simp: prod_emb_def space_PiM) moreover have "extensional {i} = space (P {i})" by (simp add: proj_space space_PiM) @@ -660,7 +660,7 @@ assumes X: "J \ I" "finite J" "\i\J. B i \ sets borel" shows "emeasure (lim\<^isub>B I P) (emb I J (Pi\<^isub>E J B)) = emeasure (P J) (Pi\<^isub>E J B)" proof cases - interpret prob_space "P {}" using prob_space by simp + interpret prob_space "P {}" using proj_prob_space by simp assume "J = {}" moreover have "emb I {} {\x. undefined} = space (lim\<^isub>B I P)" by (auto simp: space_PiM prod_emb_def) @@ -677,7 +677,7 @@ assumes "J \ I" "finite J" "\i\J. B i \ sets borel" shows "measure (lim\<^isub>B I P) (emb I J (Pi\<^isub>E J B)) = measure (P J) (Pi\<^isub>E J B)" proof - - interpret prob_space "P J" using prob_space assms by simp + interpret prob_space "P J" using proj_prob_space assms by simp show ?thesis using emeasure_limB_emb[OF assms] unfolding emeasure_eq_measure limP_finite[OF `finite J` `J \ I`] P.emeasure_eq_measure