# HG changeset patch # User hoelzl # Date 1305637239 -7200 # Node ID 760094e49a2c0c37921ba71a5e070a6c9be3e1cc # Parent b0746bd57a41a994371353df03c8f426d365a21b Collect intro-rules for sigma-algebras * * * sets_Collect shouldn't be intro rules diff -r b0746bd57a41 -r 760094e49a2c src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Tue May 17 14:36:54 2011 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Tue May 17 15:00:39 2011 +0200 @@ -77,6 +77,42 @@ lemma (in ring_of_sets) Int_space_eq2 [simp]: "x \ sets M \ x \ space M = x" by (metis Int_absorb2 sets_into_space) +lemma (in ring_of_sets) sets_Collect_conj: + assumes "{x\space M. P x} \ sets M" "{x\space M. Q x} \ sets M" + shows "{x\space M. Q x \ P x} \ sets M" +proof - + have "{x\space M. Q x \ P x} = {x\space M. Q x} \ {x\space M. P x}" + by auto + with assms show ?thesis by auto +qed + +lemma (in ring_of_sets) sets_Collect_disj: + assumes "{x\space M. P x} \ sets M" "{x\space M. Q x} \ sets M" + shows "{x\space M. Q x \ P x} \ sets M" +proof - + have "{x\space M. Q x \ P x} = {x\space M. Q x} \ {x\space M. P x}" + by auto + with assms show ?thesis by auto +qed + +lemma (in ring_of_sets) sets_Collect_finite_All: + assumes "\i. i \ S \ {x\space M. P i x} \ sets M" "finite S" "S \ {}" + shows "{x\space M. \i\S. P i x} \ sets M" +proof - + have "{x\space M. \i\S. P i x} = (\i\S. {x\space M. P i x})" + using `S \ {}` by auto + with assms show ?thesis by auto +qed + +lemma (in ring_of_sets) sets_Collect_finite_Ex: + assumes "\i. i \ S \ {x\space M. P i x} \ sets M" "finite S" + shows "{x\space M. \i\S. P i x} \ sets M" +proof - + have "{x\space M. \i\S. P i x} = (\i\S. {x\space M. P i x})" + by auto + with assms show ?thesis by auto +qed + locale algebra = ring_of_sets + assumes top [iff]: "space M \ sets M" @@ -134,6 +170,22 @@ qed qed +lemma (in algebra) sets_Collect_neg: + assumes "{x\space M. P x} \ sets M" + shows "{x\space M. \ P x} \ sets M" +proof - + have "{x\space M. \ P x} = space M - {x\space M. P x}" by auto + with assms show ?thesis by auto +qed + +lemma (in algebra) sets_Collect_imp: + "{x\space M. P x} \ sets M \ {x\space M. Q x} \ sets M \ {x\space M. Q x \ P x} \ sets M" + unfolding imp_conv_disj by (intro sets_Collect_disj sets_Collect_neg) + +lemma (in algebra) sets_Collect_const: + "{x\space M. P} \ sets M" + by (cases P) auto + section {* Restricted algebras *} abbreviation (in algebra) @@ -212,6 +264,26 @@ algebra M \ (\A. range A \ sets M \ (\i::nat. A i) \ sets M)" by (simp add: sigma_algebra_def sigma_algebra_axioms_def) +lemma (in sigma_algebra) sets_Collect_countable_All: + assumes "\i. {x\space M. P i x} \ sets M" + shows "{x\space M. \i::'i::countable. P i x} \ sets M" +proof - + have "{x\space M. \i::'i::countable. P i x} = (\i. {x\space M. P i x})" by auto + with assms show ?thesis by auto +qed + +lemma (in sigma_algebra) sets_Collect_countable_Ex: + assumes "\i. {x\space M. P i x} \ sets M" + shows "{x\space M. \i::'i::countable. P i x} \ sets M" +proof - + have "{x\space M. \i::'i::countable. P i x} = (\i. {x\space M. P i x})" by auto + with assms show ?thesis by auto +qed + +lemmas (in sigma_algebra) sets_Collect = + sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const + sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All + subsection {* Binary Unions *} definition binary :: "'a \ 'a \ nat \ 'a"