# HG changeset patch # User hoelzl # Date 1291230581 -3600 # Node ID 251df82c0088d4195f12398c5bc9eb288cdfdbfd # Parent ab83ba2cd5d1cbc7d9b399c503d5a19570c987eb Replace algebra_eqI by algebra.equality; Move sets_sigma_subset to Sigma_Algebra; diff -r ab83ba2cd5d1 -r 251df82c0088 src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Wed Dec 01 19:42:09 2010 +0100 +++ b/src/HOL/Probability/Borel_Space.thy Wed Dec 01 20:09:41 2010 +0100 @@ -6,12 +6,6 @@ imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis begin -lemma (in sigma_algebra) sets_sigma_subset: - assumes "space N = space M" - assumes "sets N \ sets M" - shows "sets (sigma N) \ sets M" - by (unfold assms sets_sigma, rule sigma_sets_subset, rule assms) - lemma LIMSEQ_max: "u ----> (x::real) \ (\i. max (u i) 0) ----> max x 0" by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D) @@ -612,13 +606,10 @@ then show ?thesis by (intro sets_sigma_subset) auto qed -lemma algebra_eqI: assumes "sets A = sets (B::'a algebra)" "space A = space B" - shows "A = B" using assms by auto - lemma borel_eq_atMost: "borel = (sigma \space=UNIV, sets=range (\ a. {.. a::'a\ordered_euclidean_space})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets borel \ sets ?SIGMA" using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace by auto @@ -629,7 +620,7 @@ lemma borel_eq_atLeastAtMost: "borel = (sigma \space=UNIV, sets=range (\ (a :: 'a\ordered_euclidean_space, b). {a .. b})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets borel \ sets ?SIGMA" using atMost_span_atLeastAtMost halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace @@ -641,7 +632,7 @@ lemma borel_eq_greaterThan: "borel = (sigma \space=UNIV, sets=range (\ (a :: 'a\ordered_euclidean_space). {a <..})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets borel \ sets ?SIGMA" using halfspace_le_span_greaterThan halfspace_span_halfspace_le open_span_halfspace @@ -653,7 +644,7 @@ lemma borel_eq_lessThan: "borel = (sigma \space=UNIV, sets=range (\ (a :: 'a\ordered_euclidean_space). {..< a})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets borel \ sets ?SIGMA" using halfspace_le_span_lessThan halfspace_span_halfspace_ge open_span_halfspace @@ -665,7 +656,7 @@ lemma borel_eq_greaterThanLessThan: "borel = (sigma \space=UNIV, sets=range (\ (a, b). {a <..< (b :: 'a \ ordered_euclidean_space)})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets ?SIGMA \ sets borel" by (rule borel.sets_sigma_subset) auto show "sets borel \ sets ?SIGMA" @@ -686,7 +677,7 @@ lemma borel_eq_halfspace_le: "borel = (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. x$$i \ a})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets borel \ sets ?SIGMA" using open_span_halfspace halfspace_span_halfspace_le by auto show "sets ?SIGMA \ sets borel" @@ -696,7 +687,7 @@ lemma borel_eq_halfspace_less: "borel = (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. x$$i < a})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets borel \ sets ?SIGMA" using open_span_halfspace . show "sets ?SIGMA \ sets borel" @@ -706,7 +697,7 @@ lemma borel_eq_halfspace_gt: "borel = (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. a < x$$i})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets borel \ sets ?SIGMA" using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto show "sets ?SIGMA \ sets borel" @@ -716,7 +707,7 @@ lemma borel_eq_halfspace_ge: "borel = (sigma \space=UNIV, sets=range (\ (a, i). {x::'a::ordered_euclidean_space. a \ x$$i})\)" (is "_ = ?SIGMA") -proof (rule algebra_eqI, rule antisym) +proof (intro algebra.equality antisym) show "sets borel \ sets ?SIGMA" using halfspace_span_halfspace_ge open_span_halfspace by auto show "sets ?SIGMA \ sets borel" diff -r ab83ba2cd5d1 -r 251df82c0088 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Wed Dec 01 19:42:09 2010 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Wed Dec 01 20:09:41 2010 +0100 @@ -397,6 +397,12 @@ by (auto intro: sigma_sets.Empty sigma_sets_top) qed +lemma (in sigma_algebra) sets_sigma_subset: + assumes "space N = space M" + assumes "sets N \ sets M" + shows "sets (sigma N) \ sets M" + by (unfold assms sets_sigma, rule sigma_sets_subset, rule assms) + section {* Measurable functions *} definition @@ -1149,7 +1155,6 @@ section {* Dynkin systems *} - locale dynkin_system = fixes M :: "'a algebra" assumes space_closed: "sets M \ Pow (space M)"