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]