diff -r 30f69046f120 -r ffdad62bc235 src/HOL/Analysis/Cartesian_Euclidean_Space.thy --- a/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Sat May 06 12:42:10 2023 +0100 +++ b/src/HOL/Analysis/Cartesian_Euclidean_Space.thy Sun May 07 14:52:53 2023 +0100 @@ -175,39 +175,6 @@ finally show "norm (A *v x) \ CARD('m) * real (CARD('n)) * B * norm x" . qed -lemma rational_approximation: - assumes "e > 0" - obtains r::real where "r \ \" "\r - x\ < e" - using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto - -lemma Rats_closure_real: "closure \ = (UNIV::real set)" -proof - - have "\x::real. x \ closure \" - by (metis closure_approachable dist_real_def rational_approximation) - then show ?thesis by auto -qed - -proposition matrix_rational_approximation: - fixes A :: "real^'n^'m" - assumes "e > 0" - obtains B where "\i j. B$i$j \ \" "onorm(\x. (A - B) *v x) < e" -proof - - have "\i j. \q \ \. \q - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" - using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"]) - then obtain B where B: "\i j. B$i$j \ \" and Bclo: "\i j. \B$i$j - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" - by (auto simp: lambda_skolem Bex_def) - show ?thesis - proof - have "onorm ((*v) (A - B)) \ real CARD('m) * real CARD('n) * - (e / (2 * real CARD('m) * real CARD('n)))" - apply (rule onorm_le_matrix_component) - using Bclo by (simp add: abs_minus_commute less_imp_le) - also have "\ < e" - using \0 < e\ by (simp add: field_split_simps) - finally show "onorm ((*v) (A - B)) < e" . - qed (use B in auto) -qed - lemma vector_sub_project_orthogonal_cart: "(b::real^'n) \ (x - ((b \ x) / (b \ b)) *s b) = 0" unfolding inner_simps scalar_mult_eq_scaleR by auto @@ -478,6 +445,42 @@ qed *) +subsection\Arbitrarily good rational approximations\ + +lemma rational_approximation: + assumes "e > 0" + obtains r::real where "r \ \" "\r - x\ < e" + using Rats_dense_in_real [of "x - e/2" "x + e/2"] assms by auto + +lemma Rats_closure_real: "closure \ = (UNIV::real set)" +proof - + have "\x::real. x \ closure \" + by (metis closure_approachable dist_real_def rational_approximation) + then show ?thesis by auto +qed + +proposition matrix_rational_approximation: + fixes A :: "real^'n^'m" + assumes "e > 0" + obtains B where "\i j. B$i$j \ \" "onorm(\x. (A - B) *v x) < e" +proof - + have "\i j. \q \ \. \q - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" + using assms by (force intro: rational_approximation [of "e / (2 * CARD('m) * CARD('n))"]) + then obtain B where B: "\i j. B$i$j \ \" and Bclo: "\i j. \B$i$j - A $ i $ j\ < e / (2 * CARD('m) * CARD('n))" + by (auto simp: lambda_skolem Bex_def) + show ?thesis + proof + have "onorm ((*v) (A - B)) \ real CARD('m) * real CARD('n) * + (e / (2 * real CARD('m) * real CARD('n)))" + apply (rule onorm_le_matrix_component) + using Bclo by (simp add: abs_minus_commute less_imp_le) + also have "\ < e" + using \0 < e\ by (simp add: field_split_simps) + finally show "onorm ((*v) (A - B)) < e" . + qed (use B in auto) +qed + + subsection "Derivative" definition\<^marker>\tag important\ "jacobian f net = matrix(frechet_derivative f net)" @@ -547,8 +550,7 @@ lemma interval_split_cart: "{a..b::real^'n} \ {x. x$k \ c} = {a .. (\ i. if i = k then min (b$k) c else b$i)}" "cbox a b \ {x. x$k \ c} = {(\ i. if i = k then max (a$k) c else a$i) .. b}" - apply (rule_tac[!] set_eqI) - unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart + unfolding Int_iff mem_box_cart mem_Collect_eq interval_cbox_cart set_eq_iff unfolding vec_lambda_beta by auto