# HG changeset patch # User hoelzl # Date 1475167954 -7200 # Node ID f77dca1abf1b6845fe96607abe39660287a3fd64 # Parent 02de4a58e210ec0209b272749de9672dd668d32f HOL-Analysis: prove that a starlike set is negligible (based on HOL Light proof ported by L. C. Paulson) diff -r 02de4a58e210 -r f77dca1abf1b src/HOL/Analysis/Complete_Measure.thy --- a/src/HOL/Analysis/Complete_Measure.thy Thu Sep 29 13:54:57 2016 +0200 +++ b/src/HOL/Analysis/Complete_Measure.thy Thu Sep 29 18:52:34 2016 +0200 @@ -794,6 +794,27 @@ lemma (in complete_measure) complete2: "A \ B \ B \ null_sets M \ A \ null_sets M" using complete[of A B] null_sets_subset[of A B M] by simp +lemma (in complete_measure) AE_iff_null_sets: "(AE x in M. P x) \ {x\space M. \ P x} \ null_sets M" + unfolding eventually_ae_filter by (auto intro: complete2) + +lemma (in complete_measure) null_sets_iff_AE: "A \ null_sets M \ ((AE x in M. x \ A) \ A \ space M)" + unfolding AE_iff_null_sets by (auto cong: rev_conj_cong dest: sets.sets_into_space simp: subset_eq) + +lemma (in complete_measure) in_sets_AE: + assumes ae: "AE x in M. x \ A \ x \ B" and A: "A \ sets M" and B: "B \ space M" + shows "B \ sets M" +proof - + have "(AE x in M. x \ B - A \ x \ A - B)" + using ae by eventually_elim auto + then have "B - A \ null_sets M" "A - B \ null_sets M" + using A B unfolding null_sets_iff_AE by (auto dest: sets.sets_into_space) + then have "A \ (B - A) - (A - B) \ sets M" + using A by blast + also have "A \ (B - A) - (A - B) = B" + by auto + finally show "B \ sets M" . +qed + lemma (in complete_measure) vimage_null_part_null_sets: assumes f: "f \ M \\<^sub>M N" and eq: "null_sets N \ null_sets (distr M N f)" and A: "A \ completion N" @@ -955,6 +976,16 @@ by simp qed (auto intro!: bexI[of _ S]) +lemma (in complete_measure) null_sets_outer: + "S \ null_sets M \ (\e>0. \T\fmeasurable M. S \ T \ measure M T < e)" +proof - + have "S \ null_sets M \ (S \ fmeasurable M \ 0 = measure M S)" + by (auto simp: null_sets_def emeasure_eq_measure2 intro: fmeasurableI) (simp add: measure_def) + also have "\ = (\e>0. \T\fmeasurable M. S \ T \ measure M T < e)" + unfolding fmeasurable_measure_inner_outer by auto + finally show ?thesis . +qed + lemma (in cld_measure) notin_sets_outer_measure_of_cover: assumes E: "E \ space M" "E \ sets M" shows "\B\sets M. 0 < emeasure M B \ emeasure M B < \ \ diff -r 02de4a58e210 -r f77dca1abf1b src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Sep 29 13:54:57 2016 +0200 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Thu Sep 29 18:52:34 2016 +0200 @@ -959,6 +959,27 @@ lemma lmeasure_integral: "S \ lmeasurable \ measure lebesgue S = integral S (\x. 1::real)" by (auto simp add: lmeasure_integral_UNIV indicator_def[abs_def] lmeasurable_iff_integrable_on) +lemma + assumes \: "\ division_of S" + shows lmeasurable_division: "S \ lmeasurable" (is ?l) + and content_divsion: "(\k\\. measure lebesgue k) = measure lebesgue S" (is ?m) +proof - + { fix d1 d2 assume *: "d1 \ \" "d2 \ \" "d1 \ d2" + then obtain a b c d where "d1 = cbox a b" "d2 = cbox c d" + using division_ofD(4)[OF \] by blast + with division_ofD(5)[OF \ *] + have "d1 \ sets lborel" "d2 \ sets lborel" "d1 \ d2 \ (cbox a b - box a b) \ (cbox c d - box c d)" + by auto + moreover have "(cbox a b - box a b) \ (cbox c d - box c d) \ null_sets lborel" + by (intro null_sets.Un null_sets_cbox_Diff_box) + ultimately have "d1 \ d2 \ null_sets lborel" + by (blast intro: null_sets_subset) } + then show ?l ?m + unfolding division_ofD(6)[OF \, symmetric] + using division_ofD(1,4)[OF \] + by (auto intro!: measure_Union_AE[symmetric] simp: completion.AE_iff_null_sets Int_def[symmetric] pairwise_def null_sets_def) +qed + text \This should be an abbreviation for negligible.\ lemma negligible_iff_null_sets: "negligible S \ S \ null_sets lebesgue" proof @@ -982,6 +1003,127 @@ qed auto qed +lemma starlike_negligible: + assumes "closed S" + and eq1: "\c x. \(a + c *\<^sub>R x) \ S; 0 \ c; a + x \ S\ \ c = 1" + shows "negligible S" +proof - + have "negligible (op + (-a) ` S)" + proof (subst negligible_on_intervals, intro allI) + fix u v + show "negligible (op + (- a) ` S \ cbox u v)" + unfolding negligible_iff_null_sets + apply (rule starlike_negligible_compact) + apply (simp add: assms closed_translation closed_Int_compact, clarify) + by (metis eq1 minus_add_cancel) + qed + then show ?thesis + by (rule negligible_translation_rev) +qed + +lemma starlike_negligible_strong: + assumes "closed S" + and star: "\c x. \0 \ c; c < 1; a+x \ S\ \ a + c *\<^sub>R x \ S" + shows "negligible S" +proof - + show ?thesis + proof (rule starlike_negligible [OF \closed S\, of a]) + fix c x + assume cx: "a + c *\<^sub>R x \ S" "0 \ c" "a + x \ S" + with star have "~ (c < 1)" by auto + moreover have "~ (c > 1)" + using star [of "1/c" "c *\<^sub>R x"] cx by force + ultimately show "c = 1" by arith + qed +qed + +subsection\Applications\ + +lemma negligible_hyperplane: + assumes "a \ 0 \ b \ 0" shows "negligible {x. a \ x = b}" +proof - + obtain x where x: "a \ x \ b" + using assms + apply auto + apply (metis inner_eq_zero_iff inner_zero_right) + using inner_zero_right by fastforce + show ?thesis + apply (rule starlike_negligible [OF closed_hyperplane, of x]) + using x apply (auto simp: algebra_simps) + done +qed + +lemma negligible_lowdim: + fixes S :: "'N :: euclidean_space set" + assumes "dim S < DIM('N)" + shows "negligible S" +proof - + obtain a where "a \ 0" and a: "span S \ {x. a \ x = 0}" + using lowdim_subset_hyperplane [OF assms] by blast + have "negligible (span S)" + using \a \ 0\ a negligible_hyperplane by (blast intro: negligible_subset) + then show ?thesis + using span_inc by (blast intro: negligible_subset) +qed + +proposition negligible_convex_frontier: + fixes S :: "'N :: euclidean_space set" + assumes "convex S" + shows "negligible(frontier S)" +proof - + have nf: "negligible(frontier S)" if "convex S" "0 \ S" for S :: "'N set" + proof - + obtain B where "B \ S" and indB: "independent B" + and spanB: "S \ span B" and cardB: "card B = dim S" + by (metis basis_exists) + consider "dim S < DIM('N)" | "dim S = DIM('N)" + using dim_subset_UNIV le_eq_less_or_eq by blast + then show ?thesis + proof cases + case 1 + show ?thesis + by (rule negligible_subset [of "closure S"]) + (simp_all add: Diff_subset frontier_def negligible_lowdim 1) + next + case 2 + obtain a where a: "a \ interior S" + apply (rule interior_simplex_nonempty [OF indB]) + apply (simp add: indB independent_finite) + apply (simp add: cardB 2) + apply (metis \B \ S\ \0 \ S\ \convex S\ insert_absorb insert_subset interior_mono subset_hull) + done + show ?thesis + proof (rule starlike_negligible_strong [where a=a]) + fix c::real and x + have eq: "a + c *\<^sub>R x = (a + x) - (1 - c) *\<^sub>R ((a + x) - a)" + by (simp add: algebra_simps) + assume "0 \ c" "c < 1" "a + x \ frontier S" + then show "a + c *\<^sub>R x \ frontier S" + apply (clarsimp simp: frontier_def) + apply (subst eq) + apply (rule mem_interior_closure_convex_shrink [OF \convex S\ a, of _ "1-c"], auto) + done + qed auto + qed + qed + show ?thesis + proof (cases "S = {}") + case True then show ?thesis by auto + next + case False + then obtain a where "a \ S" by auto + show ?thesis + using nf [of "(\x. -a + x) ` S"] + by (metis \a \ S\ add.left_inverse assms convex_translation_eq frontier_translation + image_eqI negligible_translation_rev) + qed +qed + +corollary negligible_sphere: "negligible (sphere a e)" + using frontier_cball negligible_convex_frontier convex_cball + by (blast intro: negligible_subset) + + lemma non_negligible_UNIV [simp]: "\ negligible UNIV" unfolding negligible_iff_null_sets by (auto simp: null_sets_def emeasure_lborel_UNIV) diff -r 02de4a58e210 -r f77dca1abf1b src/HOL/Analysis/Lebesgue_Measure.thy --- a/src/HOL/Analysis/Lebesgue_Measure.thy Thu Sep 29 13:54:57 2016 +0200 +++ b/src/HOL/Analysis/Lebesgue_Measure.thy Thu Sep 29 18:52:34 2016 +0200 @@ -8,7 +8,7 @@ section \Lebesgue measure\ theory Lebesgue_Measure - imports Finite_Product_Measure Bochner_Integration Caratheodory Complete_Measure + imports Finite_Product_Measure Bochner_Integration Caratheodory Complete_Measure Summation_Tests begin subsection \Every right continuous and nondecreasing function gives rise to a measure\ @@ -695,6 +695,19 @@ finally show "lebesgue = density (distr lebesgue lebesgue T) (\_. (\j\Basis. \c j\))" . qed +lemma lebesgue_measurable_scaling[measurable]: "op *\<^sub>R x \ lebesgue \\<^sub>M lebesgue" +proof cases + assume "x = 0" + then have "op *\<^sub>R x = (\x. 0::'a)" + by (auto simp: fun_eq_iff) + then show ?thesis by auto +next + assume "x \ 0" then show ?thesis + using lebesgue_affine_measurable[of "\_. x" 0] + unfolding scaleR_scaleR[symmetric] scaleR_setsum_right[symmetric] euclidean_representation + by (auto simp add: ac_simps) +qed + lemma fixes m :: real and \ :: "'a::euclidean_space" defines "T r d x \ r *\<^sub>R x + d" @@ -884,4 +897,93 @@ and lmeasurable_box [iff]: "box a b \ lmeasurable" by (auto simp: fmeasurable_def emeasure_lborel_box_eq emeasure_lborel_cbox_eq) +lemma lmeasurable_compact: "compact S \ S \ lmeasurable" + using emeasure_compact_finite[of S] by (intro fmeasurableI) (auto simp: borel_compact) + +lemma lmeasurable_open: "bounded S \ open S \ S \ lmeasurable" + using emeasure_bounded_finite[of S] by (intro fmeasurableI) (auto simp: borel_open) + +lemma lmeasurable_ball: "ball a r \ lmeasurable" + by (simp add: lmeasurable_open) + +lemma lmeasurable_interior: "bounded S \ interior S \ lmeasurable" + by (simp add: bounded_interior lmeasurable_open) + +lemma null_sets_cbox_Diff_box: "cbox a b - box a b \ null_sets lborel" +proof - + have "emeasure lborel (cbox a b - box a b) = 0" + by (subst emeasure_Diff) (auto simp: emeasure_lborel_cbox_eq emeasure_lborel_box_eq box_subset_cbox) + then have "cbox a b - box a b \ null_sets lborel" + by (auto simp: null_sets_def) + then show ?thesis + by (auto dest!: AE_not_in) +qed +subsection\ A nice lemma for negligibility proofs.\ + +lemma summable_iff_suminf_neq_top: "(\n. f n \ 0) \ \ summable f \ (\i. ennreal (f i)) = top" + by (metis summable_suminf_not_top) + +proposition starlike_negligible_bounded_gmeasurable: + fixes S :: "'a :: euclidean_space set" + assumes S: "S \ sets lebesgue" and "bounded S" + and eq1: "\c x. \(c *\<^sub>R x) \ S; 0 \ c; x \ S\ \ c = 1" + shows "S \ null_sets lebesgue" +proof - + obtain M where "0 < M" "S \ ball 0 M" + using \bounded S\ by (auto dest: bounded_subset_ballD) + + let ?f = "\n. root DIM('a) (Suc n)" + + have vimage_eq_image: "op *\<^sub>R (?f n) -` S = op *\<^sub>R (1 / ?f n) ` S" for n + apply safe + subgoal for x by (rule image_eqI[of _ _ "?f n *\<^sub>R x"]) auto + subgoal by auto + done + + have eq: "(1 / ?f n) ^ DIM('a) = 1 / Suc n" for n + by (simp add: field_simps) + + { fix n x assume x: "root DIM('a) (1 + real n) *\<^sub>R x \ S" + have "1 * norm x \ root DIM('a) (1 + real n) * norm x" + by (rule mult_mono) auto + also have "\ < M" + using x \S \ ball 0 M\ by auto + finally have "norm x < M" by simp } + note less_M = this + + have "(\n. ennreal (1 / Suc n)) = top" + using not_summable_harmonic[where 'a=real] summable_Suc_iff[where f="\n. 1 / (real n)"] + by (intro summable_iff_suminf_neq_top) (auto simp add: inverse_eq_divide) + then have "top * emeasure lebesgue S = (\n. (1 / ?f n)^DIM('a) * emeasure lebesgue S)" + unfolding ennreal_suminf_multc eq by simp + also have "\ = (\n. emeasure lebesgue (op *\<^sub>R (?f n) -` S))" + unfolding vimage_eq_image using emeasure_lebesgue_affine[of "1 / ?f n" 0 S for n] by simp + also have "\ = emeasure lebesgue (\n. op *\<^sub>R (?f n) -` S)" + proof (intro suminf_emeasure) + show "disjoint_family (\n. op *\<^sub>R (?f n) -` S)" + unfolding disjoint_family_on_def + proof safe + fix m n :: nat and x assume "m \ n" "?f m *\<^sub>R x \ S" "?f n *\<^sub>R x \ S" + with eq1[of "?f m / ?f n" "?f n *\<^sub>R x"] show "x \ {}" + by auto + qed + have "op *\<^sub>R (?f i) -` S \ sets lebesgue" for i + using measurable_sets[OF lebesgue_measurable_scaling[of "?f i"] S] by auto + then show "range (\i. op *\<^sub>R (?f i) -` S) \ sets lebesgue" + by auto + qed + also have "\ \ emeasure lebesgue (ball 0 M :: 'a set)" + using less_M by (intro emeasure_mono) auto + also have "\ < top" + using lmeasurable_ball by (auto simp: fmeasurable_def) + finally have "emeasure lebesgue S = 0" + by (simp add: ennreal_top_mult split: if_split_asm) + then show "S \ null_sets lebesgue" + unfolding null_sets_def using \S \ sets lebesgue\ by auto +qed + +corollary starlike_negligible_compact: + "compact S \ (\c x. \(c *\<^sub>R x) \ S; 0 \ c; x \ S\ \ c = 1) \ S \ null_sets lebesgue" + using starlike_negligible_bounded_gmeasurable[of S] by (auto simp: compact_eq_bounded_closed) + end diff -r 02de4a58e210 -r f77dca1abf1b src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy Thu Sep 29 13:54:57 2016 +0200 +++ b/src/HOL/Analysis/Measure_Space.thy Thu Sep 29 18:52:34 2016 +0200 @@ -1127,6 +1127,12 @@ unfolding eventually_ae_filter by auto qed auto +lemma pairwise_alt: "pairwise R S \ (\x\S. \y\S-{x}. R x y)" + by (auto simp add: pairwise_def) + +lemma AE_pairwise: "countable F \ pairwise (\A B. AE x in M. R x A B) F \ (AE x in M. pairwise (R x) F)" + unfolding pairwise_alt by (simp add: AE_ball_countable) + lemma AE_discrete_difference: assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" @@ -1453,6 +1459,12 @@ by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top del: real_of_ereal_enn2ereal) +lemma measure_eq_AE: + assumes iff: "AE x in M. x \ A \ x \ B" + assumes A: "A \ sets M" and B: "B \ sets M" + shows "measure M A = measure M B" + using assms emeasure_eq_AE[OF assms] by (simp add: measure_def) + lemma measure_Union: "emeasure M A \ \ \ emeasure M B \ \ \ A \ sets M \ B \ sets M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" @@ -1554,6 +1566,12 @@ done qed +lemma measure_Un_null_set: "A \ sets M \ B \ null_sets M \ measure M (A \ B) = measure M A" + by (simp add: measure_def emeasure_Un_null_set) + +lemma measure_Diff_null_set: "A \ sets M \ B \ null_sets M \ measure M (A - B) = measure M A" + by (simp add: measure_def emeasure_Diff_null_set) + lemma measure_eq_setsum_singleton: "finite S \ (\x. x \ S \ {x} \ sets M) \ (\x. x \ S \ emeasure M {x} \ \) \ measure M S = (\x\S. measure M {x})" @@ -1595,6 +1613,9 @@ lemma fmeasurableD[dest, measurable_dest]: "A \ fmeasurable M \ A \ sets M" by (auto simp: fmeasurable_def) +lemma fmeasurableD2: "A \ fmeasurable M \ emeasure M A \ top" + by (auto simp: fmeasurable_def) + lemma fmeasurableI: "A \ sets M \ emeasure M A < \ \ A \ fmeasurable M" by (auto simp: fmeasurable_def) @@ -1648,6 +1669,86 @@ using assms by (intro sets.countable_INT') auto qed +lemma measure_Un2: + "A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M (B - A)" + using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff) + +lemma measure_Un3: + assumes "A \ fmeasurable M" "B \ fmeasurable M" + shows "measure M (A \ B) = measure M A + measure M B - measure M (A \ B)" +proof - + have "measure M (A \ B) = measure M A + measure M (B - A)" + using assms by (rule measure_Un2) + also have "B - A = B - (A \ B)" + by auto + also have "measure M (B - (A \ B)) = measure M B - measure M (A \ B)" + using assms by (intro measure_Diff) (auto simp: fmeasurable_def) + finally show ?thesis + by simp +qed + +lemma measure_Un_AE: + "AE x in M. x \ A \ x \ B \ A \ fmeasurable M \ B \ fmeasurable M \ + measure M (A \ B) = measure M A + measure M B" + by (subst measure_Un2) (auto intro!: measure_eq_AE) + +lemma measure_UNION_AE: + assumes I: "finite I" + shows "(\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. AE x in M. x \ F i \ x \ F j) I \ + measure M (\i\I. F i) = (\i\I. measure M (F i))" + unfolding AE_pairwise[OF countable_finite, OF I] + using I + apply (induction I rule: finite_induct) + apply simp + apply (simp add: pairwise_insert) + apply (subst measure_Un_AE) + apply auto + done + +lemma measure_UNION': + "finite I \ (\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. disjnt (F i) (F j)) I \ + measure M (\i\I. F i) = (\i\I. measure M (F i))" + by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually) + +lemma measure_Union_AE: + "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise (\S T. AE x in M. x \ S \ x \ T) F \ + measure M (\F) = (\S\F. measure M S)" + using measure_UNION_AE[of F "\x. x" M] by simp + +lemma measure_Union': + "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise disjnt F \ measure M (\F) = (\S\F. measure M S)" + using measure_UNION'[of F "\x. x" M] by simp + +lemma measure_Un_le: + assumes "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" +proof cases + assume "A \ fmeasurable M \ B \ fmeasurable M" + with measure_subadditive[of A M B] assms show ?thesis + by (auto simp: fmeasurableD2) +next + assume "\ (A \ fmeasurable M \ B \ fmeasurable M)" + then have "A \ B \ fmeasurable M" + using fmeasurableI2[of "A \ B" M A] fmeasurableI2[of "A \ B" M B] assms by auto + with assms show ?thesis + by (auto simp: fmeasurable_def measure_def less_top[symmetric]) +qed + +lemma measure_UNION_le: + "finite I \ (\i. i \ I \ F i \ sets M) \ measure M (\i\I. F i) \ (\i\I. measure M (F i))" +proof (induction I rule: finite_induct) + case (insert i I) + then have "measure M (\i\insert i I. F i) \ measure M (F i) + measure M (\i\I. F i)" + by (auto intro!: measure_Un_le) + also have "measure M (\i\I. F i) \ (\i\I. measure M (F i))" + using insert by auto + finally show ?case + using insert by simp +qed simp + +lemma measure_Union_le: + "finite F \ (\S. S \ F \ S \ sets M) \ measure M (\F) \ (\S\F. measure M S)" + using measure_UNION_le[of F "\x. x" M] by simp + subsection \Measure spaces with @{term "emeasure M (space M) < \"}\ locale finite_measure = sigma_finite_measure M for M + diff -r 02de4a58e210 -r f77dca1abf1b src/HOL/Set.thy --- a/src/HOL/Set.thy Thu Sep 29 13:54:57 2016 +0200 +++ b/src/HOL/Set.thy Thu Sep 29 18:52:34 2016 +0200 @@ -1847,7 +1847,7 @@ text \Misc\ definition pairwise :: "('a \ 'a \ bool) \ 'a set \ bool" - where "pairwise R S \ (\x \ S. \y\ S. x \ y \ R x y)" + where "pairwise R S \ (\x \ S. \y \ S. x \ y \ R x y)" lemma pairwise_subset: "pairwise P S \ T \ S \ pairwise P T" by (force simp: pairwise_def)