# HG changeset patch # User paulson # Date 1523910218 -3600 # Node ID c0ebecf6e3ebc006903a08fa45bf74c15b87d58b # Parent 706f86afff43106d60ff94576b9b55aa7ddc3c15 some more random results diff -r 706f86afff43 -r c0ebecf6e3eb src/HOL/Analysis/Determinants.thy --- a/src/HOL/Analysis/Determinants.thy Mon Apr 16 15:00:46 2018 +0100 +++ b/src/HOL/Analysis/Determinants.thy Mon Apr 16 21:23:38 2018 +0100 @@ -832,6 +832,27 @@ by blast qed +lemma det_nz_iff_inj: + fixes f :: "real^'n \ real^'n" + assumes "linear f" + shows "det (matrix f) \ 0 \ inj f" +proof + assume "det (matrix f) \ 0" + then show "inj f" + using assms invertible_det_nz inj_matrix_vector_mult by force +next + assume "inj f" + show "det (matrix f) \ 0" + using linear_injective_left_inverse [OF assms \inj f\] + by (metis assms invertible_det_nz invertible_left_inverse matrix_compose matrix_id_mat_1) +qed + +lemma det_eq_0_rank: + fixes A :: "real^'n^'n" + shows "det A = 0 \ rank A < CARD('n)" + using invertible_det_nz [of A] + by (auto simp: matrix_left_invertible_injective invertible_left_inverse less_rank_noninjective) + subsubsection\Invertibility of matrices and corresponding linear functions\ lemma matrix_left_invertible: diff -r 706f86afff43 -r c0ebecf6e3eb src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy --- a/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Mon Apr 16 15:00:46 2018 +0100 +++ b/src/HOL/Analysis/Equivalence_Lebesgue_Henstock_Integration.thy Mon Apr 16 21:23:38 2018 +0100 @@ -1287,7 +1287,7 @@ 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) + unfolding negligible_iff_null_sets by (auto simp: null_sets_def) lemma negligible_interval: "negligible (cbox a b) \ box a b = {}" "negligible (box a b) \ box a b = {}" @@ -1422,6 +1422,9 @@ using assms by (metis (full_types) closure_subset completion.complete_sets_sandwich_fmeasurable measure_closure measure_interior) qed +lemma measurable_convex: "\convex S; bounded S\ \ S \ lmeasurable" + by (simp add: measurable_Jordan negligible_convex_frontier) + subsection\Negligibility of image under non-injective linear map\ lemma negligible_Union_nat: @@ -1987,6 +1990,42 @@ qed auto qed + +subsection\Negligibility is a local property\ + +lemma locally_negligible_alt: + "negligible S \ (\x \ S. \U. openin (subtopology euclidean S) U \ x \ U \ negligible U)" + (is "_ = ?rhs") +proof + assume "negligible S" + then show ?rhs + using openin_subtopology_self by blast +next + assume ?rhs + then obtain U where ope: "\x. x \ S \ openin (subtopology euclidean S) (U x)" + and cov: "\x. x \ S \ x \ U x" + and neg: "\x. x \ S \ negligible (U x)" + by metis + obtain \ where "\ \ U ` S" "countable \" and eq: "\\ = UNION S U" + using ope by (force intro: Lindelof_openin [of "U ` S" S]) + then have "negligible (\\)" + by (metis imageE neg negligible_countable_Union subset_eq) + with eq have "negligible (UNION S U)" + by metis + moreover have "S \ UNION S U" + using cov by blast + ultimately show "negligible S" + using negligible_subset by blast +qed + +lemma locally_negligible: + "locally negligible S \ negligible S" + unfolding locally_def + apply safe + apply (meson negligible_subset openin_subtopology_self locally_negligible_alt) + by (meson negligible_subset openin_imp_subset order_refl) + + subsection\Integral bounds\ lemma set_integral_norm_bound: @@ -2564,15 +2603,171 @@ by blast qed + +subsection\Outer and inner approximation of measurable sets by well-behaved sets.\ + +proposition measurable_outer_intervals_bounded: + assumes "S \ lmeasurable" "S \ cbox a b" "e > 0" + obtains \ + where "countable \" + "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" + "pairwise (\A B. interior A \ interior B = {}) \" + "\u v. cbox u v \ \ \ \n. \i \ Basis. v \ i - u \ i = (b \ i - a \ i)/2^n" + "\K. \K \ \; box a b \ {}\ \ interior K \ {}" + "S \ \\" "\\ \ lmeasurable" "measure lebesgue (\\) \ measure lebesgue S + e" +proof (cases "box a b = {}") + case True + show ?thesis + proof (cases "cbox a b = {}") + case True + with assms have [simp]: "S = {}" + by auto + show ?thesis + proof + show "countable {}" + by simp + qed (use \e > 0\ in auto) + next + case False + show ?thesis + proof + show "countable {cbox a b}" + by simp + show "\u v. cbox u v \ {cbox a b} \ \n. \i\Basis. v \ i - u \ i = (b \ i - a \ i)/2 ^ n" + using False by (force simp: eq_cbox intro: exI [where x=0]) + show "measure lebesgue (\{cbox a b}) \ measure lebesgue S + e" + using assms by (simp add: sum_content.box_empty_imp [OF True]) + qed (use assms \cbox a b \ {}\ in auto) + qed +next + case False + let ?\ = "measure lebesgue" + have "S \ cbox a b \ lmeasurable" + using \S \ lmeasurable\ by blast + then have indS_int: "(indicator S has_integral (?\ S)) (cbox a b)" + by (metis integral_indicator \S \ cbox a b\ has_integral_integrable_integral inf.orderE integrable_on_indicator) + with \e > 0\ obtain \ where "gauge \" and \: + "\\. \\ tagged_division_of (cbox a b); \ fine \\ \ norm ((\(x,K)\\. content(K) *\<^sub>R indicator S x) - ?\ S) < e" + by (force simp: has_integral) + have inteq: "integral (cbox a b) (indicat_real S) = integral UNIV (indicator S)" + using assms by (metis has_integral_iff indS_int lmeasure_integral_UNIV) + obtain \ where \: "countable \" "\\ \ cbox a b" + and cbox: "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" + and djointish: "pairwise (\A B. interior A \ interior B = {}) \" + and covered: "\K. K \ \ \ \x \ S \ K. K \ \ x" + and close: "\u v. cbox u v \ \ \ \n. \i \ Basis. v\i - u\i = (b\i - a\i)/2^n" + and covers: "S \ \\" + using covering_lemma [of S a b \] \gauge \\ \box a b \ {}\ assms by force + show ?thesis + proof + show "\K. K \ \ \ K \ cbox a b \ K \ {} \ (\c d. K = cbox c d)" + by (meson Sup_le_iff \(2) cbox interior_empty) + have negl_int: "negligible(K \ L)" if "K \ \" "L \ \" "K \ L" for K L + proof - + have "interior K \ interior L = {}" + using djointish pairwiseD that by fastforce + moreover obtain u v x y where "K = cbox u v" "L = cbox x y" + using cbox \K \ \\ \L \ \\ by blast + ultimately show ?thesis + by (simp add: Int_interval box_Int_box negligible_interval(1)) + qed + have fincase: "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" if "finite \" "\ \ \" for \ + proof - + obtain t where t: "\K. K \ \ \ t K \ S \ K \ K \ \(t K)" + using covered \\ \ \\ subsetD by metis + have "\K \ \. \L \ \. K \ L \ interior K \ interior L = {}" + using that djointish by (simp add: pairwise_def) (metis subsetD) + with cbox that \ have \div: "\ division_of (\\)" + by (fastforce simp: division_of_def dest: cbox) + then have 1: "\\ \ lmeasurable" + by blast + have norme: "\p. \p tagged_division_of cbox a b; \ fine p\ + \ norm ((\(x,K)\p. content K * indicator S x) - integral (cbox a b) (indicator S)) < e" + by (auto simp: lmeasure_integral_UNIV assms inteq dest: \) + have "\x K y L. (x,K) \ (\K. (t K,K)) ` \ \ (y,L) \ (\K. (t K,K)) ` \ \ (x,K) \ (y,L) \ interior K \ interior L = {}" + using that djointish by (clarsimp simp: pairwise_def) (metis subsetD) + with that \ have tagged: "(\K. (t K, K)) ` \ tagged_partial_division_of cbox a b" + by (auto simp: tagged_partial_division_of_def dest: t cbox) + have fine: "\ fine (\K. (t K, K)) ` \" + using t by (auto simp: fine_def) + have *: "y \ ?\ S \ \x - y\ \ e \ x \ ?\ S + e" for x y + by arith + have "?\ (\\) \ ?\ S + e" + proof (rule *) + have "(\K\\. ?\ (K \ S)) = ?\ (\C\\. C \ S)" + apply (rule measure_negligible_finite_Union_image [OF \finite \\, symmetric]) + using \div \S \ lmeasurable\ apply blast + unfolding pairwise_def + by (metis inf.commute inf_sup_aci(3) negligible_Int subsetCE negl_int \\ \ \\) + also have "\ = ?\ (\\ \ S)" + by simp + also have "\ \ ?\ S" + by (simp add: "1" \S \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Int) + finally show "(\K\\. ?\ (K \ S)) \ ?\ S" . + next + have "?\ (\\) = sum ?\ \" + by (metis \div content_division) + also have "\ = (\K\\. content K)" + using \div by (force intro: sum.cong) + also have "\ = (\x\\. content x * indicator S (t x))" + using t by auto + finally have eq1: "?\ (\\) = (\x\\. content x * indicator S (t x))" . + have eq2: "(\K\\. ?\ (K \ S)) = (\K\\. integral K (indicator S))" + apply (rule sum.cong [OF refl]) + by (metis integral_indicator \div \S \ lmeasurable\ division_ofD(4) fmeasurable.Int inf.commute lmeasurable_cbox) + have "\\(x,K)\(\K. (t K, K)) ` \. content K * indicator S x - integral K (indicator S)\ \ e" + using Henstock_lemma_part1 [of "indicator S::'a\real", OF _ \e > 0\ \gauge \\ _ tagged fine] + indS_int norme by auto + then show "\?\ (\\) - (\K\\. ?\ (K \ S))\ \ e" + by (simp add: eq1 eq2 comm_monoid_add_class.sum.reindex inj_on_def sum_subtractf) + qed + with 1 show ?thesis by blast + qed + have "\\ \ lmeasurable \ ?\ (\\) \ ?\ S + e" + proof (cases "finite \") + case True + with fincase show ?thesis + by blast + next + case False + let ?T = "from_nat_into \" + have T: "bij_betw ?T UNIV \" + by (simp add: False \(1) bij_betw_from_nat_into) + have TM: "\n. ?T n \ lmeasurable" + by (metis False cbox finite.emptyI from_nat_into lmeasurable_cbox) + have TN: "\m n. m \ n \ negligible (?T m \ ?T n)" + by (simp add: False \(1) from_nat_into infinite_imp_nonempty negl_int) + have TB: "(\k\n. ?\ (?T k)) \ ?\ S + e" for n + proof - + have "(\k\n. ?\ (?T k)) = ?\ (UNION {..n} ?T)" + by (simp add: pairwise_def TM TN measure_negligible_finite_Union_image) + also have "?\ (UNION {..n} ?T) \ ?\ S + e" + using fincase [of "?T ` {..n}"] T by (auto simp: bij_betw_def) + finally show ?thesis . + qed + have "\\ \ lmeasurable" + by (metis lmeasurable_compact T \(2) bij_betw_def cbox compact_cbox countable_Un_Int(1) fmeasurableD fmeasurableI2 rangeI) + moreover + have "?\ (\x. from_nat_into \ x) \ ?\ S + e" + proof (rule measure_countable_Union_le [OF TM]) + show "?\ (\x\n. from_nat_into \ x) \ ?\ S + e" for n + by (metis (mono_tags, lifting) False fincase finite.emptyI finite_atMost finite_imageI from_nat_into imageE subsetI) + qed + ultimately show ?thesis by (metis T bij_betw_def) + qed + then show "\\ \ lmeasurable" "measure lebesgue (\\) \ ?\ S + e" by blast+ + qed (use \ cbox djointish close covers in auto) +qed + subsection\Lemmas about absolute integrability\ -text\Redundant!\ +text\FIXME Redundant!\ lemma absolutely_integrable_add[intro]: fixes f g :: "'n::euclidean_space \ 'm::euclidean_space" shows "f absolutely_integrable_on s \ g absolutely_integrable_on s \ (\x. f x + g x) absolutely_integrable_on s" by (rule set_integral_add) -text\Redundant!\ +text\FIXME Redundant!\ lemma absolutely_integrable_diff[intro]: fixes f g :: "'n::euclidean_space \ 'm::euclidean_space" shows "f absolutely_integrable_on s \ g absolutely_integrable_on s \ (\x. f x - g x) absolutely_integrable_on s" diff -r 706f86afff43 -r c0ebecf6e3eb src/HOL/Analysis/Lebesgue_Measure.thy --- a/src/HOL/Analysis/Lebesgue_Measure.thy Mon Apr 16 15:00:46 2018 +0100 +++ b/src/HOL/Analysis/Lebesgue_Measure.thy Mon Apr 16 21:23:38 2018 +0100 @@ -1000,9 +1000,12 @@ 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" +lemma lmeasurable_ball [simp]: "ball a r \ lmeasurable" by (simp add: lmeasurable_open) +lemma lmeasurable_cball [simp]: "cball a r \ lmeasurable" + by (simp add: lmeasurable_compact) + lemma lmeasurable_interior: "bounded S \ interior S \ lmeasurable" by (simp add: bounded_interior lmeasurable_open) diff -r 706f86afff43 -r c0ebecf6e3eb src/HOL/Analysis/Starlike.thy --- a/src/HOL/Analysis/Starlike.thy Mon Apr 16 15:00:46 2018 +0100 +++ b/src/HOL/Analysis/Starlike.thy Mon Apr 16 21:23:38 2018 +0100 @@ -8381,6 +8381,33 @@ by auto qed +lemma add_subspaces: + assumes "subspace S" "subspace T" + shows "subspace (S + T)" + unfolding subspace_def +proof (intro conjI ballI allI) + show "0 \ S + T" + by (meson assms set_zero_plus2 subsetCE subspace_0) +next + fix x y + assume "x \ S + T" and "y \ S + T" + then obtain xs xt ys yt where "xs \ S" "xt \ T" "ys \ S" "yt \ T" and eq: "x = xs+xt" "y = ys+yt" + by (meson set_plus_elim) + then have "xs+ys \ S" "xt+yt \ T" + using assms subspace_def by blast+ + then have "(xs + ys) + (xt + yt) \ S + T" + by blast + then show "x + y \ S + T" + by (simp add: eq add.assoc add.left_commute) +next + fix c x + assume "x \ S + T" + then obtain xs xt where "xs \ S" "xt \ T" "x = xs+xt" + by (meson set_plus_elim) + then show "c *\<^sub>R x \ S + T" + by (metis assms scaleR_add_right set_plus_intro subspace_def) +qed + lemma orthogonal_Int_0: assumes "subspace U" shows "U \ U\<^sup>\ = {0}"