# HG changeset patch # User Angeliki KoutsoukouArgyraki # Date 1548191603 0 # Node ID be6634e99e09c1a7b34454d7ec7784476132c953 # Parent 749aaeb4078886a993aaa2049abf4caac4222998 redid tagging for 3 theories i.e. Determinants, Change_of_Vars, Finite_Cartesian_Product diff -r 749aaeb40788 -r be6634e99e09 src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy Fri Dec 14 14:33:26 2018 +0100 +++ b/src/HOL/Analysis/Change_Of_Vars.thy Tue Jan 22 21:13:23 2019 +0000 @@ -9,9 +9,9 @@ begin -subsection%unimportant \Orthogonal Transformation of Balls\ +subsection \Orthogonal Transformation of Balls\ -lemma%unimportant image_orthogonal_transformation_ball: +lemma image_orthogonal_transformation_ball: fixes f :: "'a::euclidean_space \ 'a" assumes "orthogonal_transformation f" shows "f ` ball x r = ball (f x) r" @@ -27,7 +27,7 @@ by (auto simp: orthogonal_transformation_isometry) qed -lemma%unimportant image_orthogonal_transformation_cball: +lemma image_orthogonal_transformation_cball: fixes f :: "'a::euclidean_space \ 'a" assumes "orthogonal_transformation f" shows "f ` cball x r = cball (f x) r" @@ -46,14 +46,14 @@ subsection \Measurable Shear and Stretch\ -proposition%important +proposition fixes a :: "real^'n" assumes "m \ n" and ab_ne: "cbox a b \ {}" and an: "0 \ a$n" shows measurable_shear_interval: "(\x. \ i. if i = m then x$m + x$n else x$i) ` (cbox a b) \ lmeasurable" (is "?f ` _ \ _") and measure_shear_interval: "measure lebesgue ((\x. \ i. if i = m then x$m + x$n else x$i) ` cbox a b) = measure lebesgue (cbox a b)" (is "?Q") -proof%unimportant - +proof - have lin: "linear ?f" by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps) show fab: "?f ` cbox a b \ lmeasurable" @@ -155,13 +155,13 @@ qed -proposition%important +proposition fixes S :: "(real^'n) set" assumes "S \ lmeasurable" shows measurable_stretch: "((\x. \ k. m k * x$k) ` S) \ lmeasurable" (is "?f ` S \ _") and measure_stretch: "measure lebesgue ((\x. \ k. m k * x$k) ` S) = \prod m UNIV\ * measure lebesgue S" (is "?MEQ") -proof%unimportant - +proof - have "(?f ` S) \ lmeasurable \ ?MEQ" proof (cases "\k. m k \ 0") case True @@ -259,12 +259,12 @@ qed -proposition%important +proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "linear f" "S \ lmeasurable" shows measurable_linear_image: "(f ` S) \ lmeasurable" and measure_linear_image: "measure lebesgue (f ` S) = \det (matrix f)\ * measure lebesgue S" (is "?Q f S") -proof%unimportant - +proof - have "\S \ lmeasurable. (f ` S) \ lmeasurable \ ?Q f S" proof (rule induct_linear_elementary [OF \linear f\]; intro ballI) fix f g and S :: "(real,'n) vec set" @@ -407,7 +407,7 @@ qed -lemma%unimportant +lemma fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes f: "orthogonal_transformation f" and S: "S \ lmeasurable" shows measurable_orthogonal_image: "f ` S \ lmeasurable" @@ -430,10 +430,10 @@ inductive%important gdelta :: "'a::topological_space set \ bool" where "(\n::nat. open (F n)) \ gdelta (\(F ` UNIV))" -lemma%important fsigma_Union_compact: +lemma fsigma_Union_compact: fixes S :: "'a::{real_normed_vector,heine_borel} set" shows "fsigma S \ (\F::nat \ 'a set. range F \ Collect compact \ S = \(F ` UNIV))" -proof%unimportant safe +proof safe assume "fsigma S" then obtain F :: "nat \ 'a set" where F: "range F \ Collect closed" "S = \(F ` UNIV)" by (meson fsigma.cases image_subsetI mem_Collect_eq) @@ -464,7 +464,7 @@ by (simp add: compact_imp_closed fsigma.intros image_subset_iff) qed -lemma%unimportant gdelta_imp_fsigma: "gdelta S \ fsigma (- S)" +lemma gdelta_imp_fsigma: "gdelta S \ fsigma (- S)" proof (induction rule: gdelta.induct) case (1 F) have "- \(F ` UNIV) = (\i. -(F i))" @@ -473,7 +473,7 @@ by (simp add: fsigma.intros closed_Compl 1) qed -lemma%unimportant fsigma_imp_gdelta: "fsigma S \ gdelta (- S)" +lemma fsigma_imp_gdelta: "fsigma S \ gdelta (- S)" proof (induction rule: fsigma.induct) case (1 F) have "- \(F ` UNIV) = (\i. -(F i))" @@ -482,11 +482,11 @@ by (simp add: 1 gdelta.intros open_closed) qed -lemma%unimportant gdelta_complement: "gdelta(- S) \ fsigma S" +lemma gdelta_complement: "gdelta(- S) \ fsigma S" using fsigma_imp_gdelta gdelta_imp_fsigma by force text\A Lebesgue set is almost an \F_sigma\ or \G_delta\.\ -lemma%unimportant lebesgue_set_almost_fsigma: +lemma lebesgue_set_almost_fsigma: assumes "S \ sets lebesgue" obtains C T where "fsigma C" "negligible T" "C \ T = S" "disjnt C T" proof - @@ -521,7 +521,7 @@ qed qed -lemma%unimportant lebesgue_set_almost_gdelta: +lemma lebesgue_set_almost_gdelta: assumes "S \ sets lebesgue" obtains C T where "gdelta C" "negligible T" "S \ T = C" "disjnt S T" proof - @@ -539,12 +539,12 @@ qed -proposition%important measure_semicontinuous_with_hausdist_explicit: +proposition measure_semicontinuous_with_hausdist_explicit: assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0" obtains d where "d > 0" "\T. \T \ lmeasurable; \y. y \ T \ \x. x \ S \ dist x y < d\ \ measure lebesgue T < measure lebesgue S + e" -proof%unimportant (cases "S = {}") +proof (cases "S = {}") case True with that \e > 0\ show ?thesis by force next @@ -608,10 +608,10 @@ qed qed -proposition%important lebesgue_regular_inner: +proposition lebesgue_regular_inner: assumes "S \ sets lebesgue" obtains K C where "negligible K" "\n::nat. compact(C n)" "S = (\n. C n) \ K" -proof%unimportant - +proof - have "\T. closed T \ T \ S \ (S - T) \ lmeasurable \ measure lebesgue (S - T) < (1/2)^n" for n using sets_lebesgue_inner_closed assms by (metis sets_lebesgue_inner_closed zero_less_divide_1_iff zero_less_numeral zero_less_power) @@ -661,7 +661,7 @@ qed -lemma%unimportant sets_lebesgue_continuous_image: +lemma sets_lebesgue_continuous_image: assumes T: "T \ sets lebesgue" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" and "T \ S" shows "f ` T \ sets lebesgue" @@ -686,7 +686,7 @@ by (simp add: Teq image_Un image_Union) qed -lemma%unimportant differentiable_image_in_sets_lebesgue: +lemma differentiable_image_in_sets_lebesgue: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" shows "f`S \ sets lebesgue" @@ -698,7 +698,7 @@ by (auto simp: intro!: negligible_differentiable_image_negligible [OF dim]) qed auto -lemma%unimportant sets_lebesgue_on_continuous_image: +lemma sets_lebesgue_on_continuous_image: assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and contf: "continuous_on S f" and negim: "\T. \negligible T; T \ S\ \ negligible(f ` T)" shows "f ` X \ sets (lebesgue_on (f ` S))" @@ -713,7 +713,7 @@ by (auto simp: sets_restrict_space_iff) qed -lemma%unimportant differentiable_image_in_sets_lebesgue_on: +lemma differentiable_image_in_sets_lebesgue_on: fixes f :: "'m::euclidean_space \ 'n::euclidean_space" assumes S: "S \ sets lebesgue" and X: "X \ sets (lebesgue_on S)" and dim: "DIM('m) \ DIM('n)" and f: "f differentiable_on S" @@ -727,7 +727,7 @@ qed -proposition%important +proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" @@ -737,7 +737,7 @@ "f ` S \ lmeasurable" and measure_bounded_differentiable_image: "measure lebesgue (f ` S) \ B * measure lebesgue S" (is "?M") -proof%unimportant - +proof - have "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ B * measure lebesgue S" proof (cases "B < 0") case True @@ -963,13 +963,13 @@ then show "f ` S \ lmeasurable" ?M by blast+ qed -lemma%important +lemma (*FIXME needs name? *) fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows m_diff_image_weak: "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" -proof%unimportant - +proof - let ?\ = "measure lebesgue" have aint_S: "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int unfolding absolutely_integrable_on_def by auto @@ -1161,7 +1161,7 @@ qed -theorem%important +theorem(* FIXME needs name? *) fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" @@ -1169,7 +1169,7 @@ shows measurable_differentiable_image: "f ` S \ lmeasurable" and measure_differentiable_image: "measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" (is "?M") -proof%unimportant - +proof - let ?I = "\n::nat. cbox (vec (-n)) (vec n) \ S" let ?\ = "measure lebesgue" have "x \ cbox (vec (- real (nat \norm x\))) (vec (real (nat \norm x\)))" for x :: "real^'n::_" @@ -1214,7 +1214,7 @@ qed -lemma%unimportant borel_measurable_simple_function_limit_increasing: +lemma borel_measurable_simple_function_limit_increasing: fixes f :: "'a::euclidean_space \ real" shows "(f \ borel_measurable lebesgue \ (\x. 0 \ f x)) \ (\g. (\n x. 0 \ g n x \ g n x \ f x) \ (\n x. g n x \ (g(Suc n) x)) \ @@ -1413,7 +1413,7 @@ subsection\Borel measurable Jacobian determinant\ -lemma%unimportant lemma_partial_derivatives0: +lemma lemma_partial_derivatives0: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim0: "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x. x \ S - {0} \ norm x < e \ k * norm x \ \v \ x\)" @@ -1486,7 +1486,7 @@ mem_Collect_eq module_hom_zero span_base span_raw_def) qed -lemma%unimportant lemma_partial_derivatives: +lemma lemma_partial_derivatives: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim: "((\x. f (x - a) /\<^sub>R norm (x - a)) \ 0) (at a within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x \ S - {a}. norm(a - x) < e \ k * norm(a - x) \ \v \ (x - a)\)" @@ -1504,12 +1504,12 @@ qed -proposition%important borel_measurable_partial_derivatives: +proposition borel_measurable_partial_derivatives: fixes f :: "real^'m::{finite,wellorder} \ real^'n" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. (matrix(f' x)$m$n)) \ borel_measurable (lebesgue_on S)" -proof%unimportant - +proof - have contf: "continuous_on S f" using continuous_on_eq_continuous_within f has_derivative_continuous by blast have "{x \ S. (matrix (f' x)$m$n) \ b} \ sets lebesgue" for b @@ -2095,7 +2095,7 @@ qed -theorem%important borel_measurable_det_Jacobian: +theorem borel_measurable_det_Jacobian: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. det(matrix(f' x))) \ borel_measurable (lebesgue_on S)" @@ -2104,12 +2104,12 @@ text\The localisation wrt S uses the same argument for many similar results.\ (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*) -lemma%important borel_measurable_lebesgue_on_preimage_borel: +theorem borel_measurable_lebesgue_on_preimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S) \ (\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" -proof%unimportant - +proof - have "{x. (if x \ S then f x else 0) \ T} \ sets lebesgue \ {x \ S. f x \ T} \ sets lebesgue" if "T \ sets borel" for T proof (cases "0 \ T") @@ -2131,7 +2131,7 @@ by blast qed -lemma%unimportant sets_lebesgue_almost_borel: +lemma sets_lebesgue_almost_borel: assumes "S \ sets lebesgue" obtains B N where "B \ sets borel" "negligible N" "B \ N = S" proof - @@ -2141,7 +2141,7 @@ by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that) qed -lemma%unimportant double_lebesgue_sets: +lemma double_lebesgue_sets: assumes S: "S \ sets lebesgue" and T: "T \ sets lebesgue" and fim: "f ` S \ T" shows "(\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue) \ f \ borel_measurable (lebesgue_on S) \ @@ -2186,7 +2186,7 @@ subsection\Simplest case of Sard's theorem (we don't need continuity of derivative)\ -lemma%important Sard_lemma00: +lemma Sard_lemma00: fixes P :: "'b::euclidean_space set" assumes "a \ 0" and a: "a *\<^sub>R i \ 0" and i: "i \ Basis" and P: "P \ {x. a *\<^sub>R i \ x = 0}" @@ -2194,7 +2194,7 @@ obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (DIM('b) - 1)" -proof%unimportant - +proof - have "a > 0" using assms by simp let ?v = "(\j\Basis. (if j = i then e else m) *\<^sub>R j)" @@ -2222,7 +2222,7 @@ qed text\As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\ -lemma%unimportant Sard_lemma0: +lemma Sard_lemma0: fixes P :: "(real^'n::{finite,wellorder}) set" assumes "a \ 0" and P: "P \ {x. a \ x = 0}" and "0 \ m" "0 \ e" @@ -2282,13 +2282,13 @@ qed text\As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\ -lemma%important Sard_lemma1: +lemma Sard_lemma1: fixes P :: "(real^'n::{finite,wellorder}) set" assumes P: "dim P < CARD('n)" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm(z - w) \ m \ (\t \ P. norm(z - w - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" -proof%unimportant - +proof - obtain a where "a \ 0" "P \ {x. a \ x = 0}" using lowdim_subset_hyperplane [of P] P span_base by auto then obtain S where S: "S \ lmeasurable" @@ -2306,7 +2306,7 @@ qed qed -lemma%important Sard_lemma2: +lemma Sard_lemma2: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" (is "?m \ ?n") and "B > 0" "bounded S" @@ -2314,7 +2314,7 @@ and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" and B: "\x. x \ S \ onorm(f' x) \ B" shows "negligible(f ` S)" -proof%unimportant - +proof - have lin_f': "\x. x \ S \ linear(f' x)" using derS has_derivative_linear by blast show ?thesis @@ -2521,13 +2521,13 @@ qed -theorem%important baby_Sard: +theorem baby_Sard: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" and der: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" shows "negligible(f ` S)" -proof%unimportant - +proof - let ?U = "\n. {x \ S. norm(x) \ n \ onorm(f' x) \ real n}" have "\x. x \ S \ \n. norm x \ real n \ onorm (f' x) \ real n" by (meson linear order_trans real_arch_simple) @@ -2549,7 +2549,7 @@ subsection\A one-way version of change-of-variables not assuming injectivity. \ -lemma%important integral_on_image_ubound_weak: +lemma integral_on_image_ubound_weak: fixes f :: "real^'n::{finite,wellorder} \ real" assumes S: "S \ sets lebesgue" and f: "f \ borel_measurable (lebesgue_on (g ` S))" @@ -2560,7 +2560,7 @@ shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") -proof%unimportant - +proof - let ?D = "\x. \det (matrix (g' x))\" have cont_g: "continuous_on S g" using der_g has_derivative_continuous_on by blast @@ -2738,14 +2738,14 @@ qed -lemma%important integral_on_image_ubound_nonneg: +lemma integral_on_image_ubound_nonneg: fixes f :: "real^'n::{finite,wellorder} \ real" assumes nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") -proof%unimportant - +proof - let ?D = "\x. det (matrix (g' x))" define S' where "S' \ {x \ S. ?D x * f(g x) \ 0}" then have der_gS': "\x. x \ S' \ (g has_derivative g' x) (at x within S')" @@ -2868,7 +2868,7 @@ qed -lemma%unimportant absolutely_integrable_on_image_real: +lemma absolutely_integrable_on_image_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S" @@ -2943,7 +2943,7 @@ qed -proposition%important absolutely_integrable_on_image: +proposition absolutely_integrable_on_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" @@ -2951,7 +2951,7 @@ apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]]) using%unimportant absolutely_integrable_component [OF intS] by%unimportant auto -proposition%important integral_on_image_ubound: +proposition integral_on_image_ubound: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes"\x. x \ S \ 0 \ f(g x)" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" @@ -2965,7 +2965,7 @@ text\The classic change-of-variables theorem. We have two versions with quite general hypotheses, the first that the transforming function has a continuous inverse, the second that the base set is Lebesgue measurable.\ -lemma%unimportant cov_invertible_nonneg_le: +lemma cov_invertible_nonneg_le: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" @@ -3036,7 +3036,7 @@ qed -lemma%unimportant cov_invertible_nonneg_eq: +lemma cov_invertible_nonneg_eq: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "\y. y \ T \ (h has_derivative h' y) (at y within T)" @@ -3049,7 +3049,7 @@ by (simp add: has_integral_iff) (meson eq_iff) -lemma%unimportant cov_invertible_real: +lemma cov_invertible_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" @@ -3215,7 +3215,7 @@ qed -lemma%unimportant cv_inv_version3: +lemma cv_inv_version3: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" @@ -3241,7 +3241,7 @@ qed -lemma%unimportant cv_inv_version4: +lemma cv_inv_version4: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S) \ invertible(matrix(g' x))" and hg: "\x. x \ S \ continuous_on (g ` S) h \ h(g x) = x" @@ -3266,7 +3266,7 @@ qed -proposition%important has_absolute_integral_change_of_variables_invertible: +theorem has_absolute_integral_change_of_variables_invertible: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and hg: "\x. x \ S \ h(g x) = x" @@ -3274,7 +3274,7 @@ shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" (is "?lhs = ?rhs") -proof%unimportant - +proof - let ?S = "{x \ S. invertible (matrix (g' x))}" and ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have *: "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b" @@ -3310,7 +3310,7 @@ -lemma%unimportant has_absolute_integral_change_of_variables_compact: +theorem has_absolute_integral_change_of_variables_compact: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "compact S" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" @@ -3328,7 +3328,7 @@ qed -lemma%unimportant has_absolute_integral_change_of_variables_compact_family: +lemma has_absolute_integral_change_of_variables_compact_family: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes compact: "\n::nat. compact (F n)" and der_g: "\x. x \ (\n. F n) \ (g has_derivative g' x) (at x within (\n. F n))" @@ -3507,7 +3507,7 @@ qed -proposition%important has_absolute_integral_change_of_variables: +theorem has_absolute_integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" @@ -3515,7 +3515,7 @@ shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" -proof%unimportant - +proof - obtain C N where "fsigma C" "negligible N" and CNS: "C \ N = S" and "disjnt C N" using lebesgue_set_almost_fsigma [OF S] . then obtain F :: "nat \ (real^'m::_) set" @@ -3569,7 +3569,7 @@ qed -corollary%important absolutely_integrable_change_of_variables: +corollary absolutely_integrable_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" @@ -3578,7 +3578,7 @@ \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" using%unimportant assms has_absolute_integral_change_of_variables by%unimportant blast -corollary%important integral_change_of_variables: +corollary integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" @@ -3589,7 +3589,7 @@ using%unimportant has_absolute_integral_change_of_variables [OF S der_g inj] disj by%unimportant blast -lemma%unimportant has_absolute_integral_change_of_variables_1: +lemma has_absolute_integral_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" @@ -3628,7 +3628,7 @@ qed -corollary%important absolutely_integrable_change_of_variables_1: +corollary absolutely_integrable_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" @@ -3640,7 +3640,7 @@ subsection\Change of variables for integrals: special case of linear function\ -lemma%unimportant has_absolute_integral_change_of_variables_linear: +lemma has_absolute_integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ @@ -3665,14 +3665,14 @@ qed (use h in auto) qed -lemma%unimportant absolutely_integrable_change_of_variables_linear: +lemma absolutely_integrable_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ f absolutely_integrable_on (g ` S)" using assms has_absolute_integral_change_of_variables_linear by blast -lemma%unimportant absolutely_integrable_on_linear_image: +lemma absolutely_integrable_on_linear_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "f absolutely_integrable_on (g ` S) @@ -3680,12 +3680,12 @@ unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff by (auto simp: set_integrable_def) -lemma%important integral_change_of_variables_linear: +lemma integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" and "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S" shows "integral (g ` S) f = \det (matrix g)\ *\<^sub>R integral S (f \ g)" -proof%unimportant - +proof - have "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S) \ (f absolutely_integrable_on g ` S)" using absolutely_integrable_on_linear_image assms by blast moreover @@ -3699,18 +3699,18 @@ subsection\Change of variable for measure\ -lemma%unimportant has_measure_differentiable_image: +lemma has_measure_differentiable_image: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) = m \ ((\x. \det (matrix (f' x))\) has_integral m) S" - using%unimportant has_absolute_integral_change_of_variables [OF assms, of "\x. (1::real^1)" "vec m"] - unfolding%unimportant absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def - by%unimportant (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral) + using has_absolute_integral_change_of_variables [OF assms, of "\x. (1::real^1)" "vec m"] + unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def + by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral) -lemma%unimportant measurable_differentiable_image_eq: +lemma measurable_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" @@ -3719,23 +3719,23 @@ using has_measure_differentiable_image [OF assms] by blast -lemma%important measurable_differentiable_image_alt: +lemma measurable_differentiable_image_alt: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) absolutely_integrable_on S" - using%unimportant measurable_differentiable_image_eq [OF assms] - by%unimportant (simp only: absolutely_integrable_on_iff_nonneg) + using measurable_differentiable_image_eq [OF assms] + by (simp only: absolutely_integrable_on_iff_nonneg) -lemma%important measure_differentiable_image_eq: +lemma measure_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and der_f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and inj: "inj_on f S" and intS: "(\x. \det (matrix (f' x))\) integrable_on S" shows "measure lebesgue (f ` S) = integral S (\x. \det (matrix (f' x))\)" - using%unimportant measurable_differentiable_image_eq [OF S der_f inj] - assms has_measure_differentiable_image by%unimportant blast + using measurable_differentiable_image_eq [OF S der_f inj] + assms has_measure_differentiable_image by blast end diff -r 749aaeb40788 -r be6634e99e09 src/HOL/Analysis/Determinants.thy --- a/src/HOL/Analysis/Determinants.thy Fri Dec 14 14:33:26 2018 +0100 +++ b/src/HOL/Analysis/Determinants.thy Tue Jan 22 21:13:23 2019 +0000 @@ -15,25 +15,25 @@ definition%important trace :: "'a::semiring_1^'n^'n \ 'a" where "trace A = sum (\i. ((A$i)$i)) (UNIV::'n set)" -lemma%unimportant trace_0: "trace (mat 0) = 0" +lemma trace_0: "trace (mat 0) = 0" by (simp add: trace_def mat_def) -lemma%unimportant trace_I: "trace (mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))" +lemma trace_I: "trace (mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))" by (simp add: trace_def mat_def) -lemma%unimportant trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B" +lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B" by (simp add: trace_def sum.distrib) -lemma%unimportant trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B" +lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B" by (simp add: trace_def sum_subtractf) -lemma%important trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)" +lemma trace_mul_sym: "trace ((A::'a::comm_semiring_1^'n^'m) ** B) = trace (B**A)" apply (simp add: trace_def matrix_matrix_mult_def) apply (subst sum.swap) apply (simp add: mult.commute) done -subsubsection \Definition of determinant\ +subsubsection%important \Definition of determinant\ definition%important det:: "'a::comm_ring_1^'n^'n \ 'a" where "det A = @@ -42,8 +42,8 @@ text \Basic determinant properties\ -lemma%important det_transpose [simp]: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)" -proof%unimportant - +lemma det_transpose [simp]: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)" +proof - let ?di = "\A i j. A$i$j" let ?U = "(UNIV :: 'n set)" have fU: "finite ?U" by simp @@ -80,7 +80,7 @@ by (subst sum_permutations_inverse) (blast intro: sum.cong) qed -lemma%unimportant det_lowerdiagonal: +lemma det_lowerdiagonal: fixes A :: "'a::comm_ring_1^('n::{finite,wellorder})^('n::{finite,wellorder})" assumes ld: "\i j. i < j \ A$i$j = 0" shows "det A = prod (\i. A$i$i) (UNIV:: 'n set)" @@ -107,11 +107,11 @@ unfolding det_def by (simp add: sign_id) qed -lemma%important det_upperdiagonal: +lemma det_upperdiagonal: fixes A :: "'a::comm_ring_1^'n::{finite,wellorder}^'n::{finite,wellorder}" assumes ld: "\i j. i > j \ A$i$j = 0" shows "det A = prod (\i. A$i$i) (UNIV:: 'n set)" -proof%unimportant - +proof - let ?U = "UNIV:: 'n set" let ?PU = "{p. p permutes ?U}" let ?pp = "(\p. of_int (sign p) * prod (\i. A$i$p i) (UNIV :: 'n set))" @@ -134,11 +134,11 @@ unfolding det_def by (simp add: sign_id) qed -lemma%important det_diagonal: +proposition det_diagonal: fixes A :: "'a::comm_ring_1^'n^'n" assumes ld: "\i j. i \ j \ A$i$j = 0" shows "det A = prod (\i. A$i$i) (UNIV::'n set)" -proof%unimportant - +proof - let ?U = "UNIV:: 'n set" let ?PU = "{p. p permutes ?U}" let ?pp = "\p. of_int (sign p) * prod (\i. A$i$p i) (UNIV :: 'n set)" @@ -161,13 +161,13 @@ unfolding det_def by (simp add: sign_id) qed -lemma%unimportant det_I [simp]: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1" +lemma det_I [simp]: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1" by (simp add: det_diagonal mat_def) -lemma%unimportant det_0 [simp]: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0" +lemma det_0 [simp]: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0" by (simp add: det_def prod_zero power_0_left) -lemma%unimportant det_permute_rows: +lemma det_permute_rows: fixes A :: "'a::comm_ring_1^'n^'n" assumes p: "p permutes (UNIV :: 'n::finite set)" shows "det (\ i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A" @@ -204,11 +204,11 @@ done qed -lemma%important det_permute_columns: +lemma det_permute_columns: fixes A :: "'a::comm_ring_1^'n^'n" assumes p: "p permutes (UNIV :: 'n set)" shows "det(\ i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A" -proof%unimportant - +proof - let ?Ap = "\ i j. A$i$ p j :: 'a^'n^'n" let ?At = "transpose A" have "of_int (sign p) * det A = det (transpose (\ i. transpose A $ p i))" @@ -220,7 +220,7 @@ by simp qed -lemma%unimportant det_identical_columns: +lemma det_identical_columns: fixes A :: "'a::comm_ring_1^'n^'n" assumes jk: "j \ k" and r: "column j A = column k A" @@ -300,24 +300,24 @@ finally show "det A = 0" by simp qed -lemma%unimportant det_identical_rows: +lemma det_identical_rows: fixes A :: "'a::comm_ring_1^'n^'n" assumes ij: "i \ j" and r: "row i A = row j A" shows "det A = 0" by (metis column_transpose det_identical_columns det_transpose ij r) -lemma%unimportant det_zero_row: +lemma det_zero_row: fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m" shows "row i A = 0 \ det A = 0" and "row j F = 0 \ det F = 0" by (force simp: row_def det_def vec_eq_iff sign_nz intro!: sum.neutral)+ -lemma%unimportant det_zero_column: +lemma det_zero_column: fixes A :: "'a::{idom, ring_char_0}^'n^'n" and F :: "'b::{field}^'m^'m" shows "column i A = 0 \ det A = 0" and "column j F = 0 \ det F = 0" unfolding atomize_conj atomize_imp by (metis det_transpose det_zero_row row_transpose) -lemma%unimportant det_row_add: +lemma det_row_add: fixes a b c :: "'n::finite \ _ ^ 'n" shows "det((\ i. if i = k then a i + b i else c i)::'a::comm_ring_1^'n^'n) = det((\ i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) + @@ -358,7 +358,7 @@ by (simp add: field_simps) qed auto -lemma%unimportant det_row_mul: +lemma det_row_mul: fixes a b :: "'n::finite \ _ ^ 'n" shows "det((\ i. if i = k then c *s a i else b i)::'a::comm_ring_1^'n^'n) = c * det((\ i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)" @@ -395,7 +395,7 @@ by (simp add: field_simps) qed auto -lemma%unimportant det_row_0: +lemma det_row_0: fixes b :: "'n::finite \ _ ^ 'n" shows "det((\ i. if i = k then 0 else b i)::'a::comm_ring_1^'n^'n) = 0" using det_row_mul[of k 0 "\i. 1" b] @@ -403,11 +403,11 @@ apply (simp only: vector_smult_lzero) done -lemma%important det_row_operation: +lemma det_row_operation: fixes A :: "'a::{comm_ring_1}^'n^'n" assumes ij: "i \ j" shows "det (\ k. if k = i then row i A + c *s row j A else row k A) = det A" -proof%unimportant - +proof - let ?Z = "(\ k. if k = i then row j A else row k A) :: 'a ^'n^'n" have th: "row i ?Z = row j ?Z" by (vector row_def) have th2: "((\ k. if k = i then row i A else row k A) :: 'a^'n^'n) = A" @@ -417,7 +417,7 @@ by simp qed -lemma%unimportant det_row_span: +lemma det_row_span: fixes A :: "'a::{field}^'n^'n" assumes x: "x \ vec.span {row j A |j. j \ i}" shows "det (\ k. if k = i then row i A + x else row k A) = det A" @@ -449,10 +449,10 @@ unfolding scalar_mult_eq_scaleR . qed -lemma%unimportant matrix_id [simp]: "det (matrix id) = 1" +lemma matrix_id [simp]: "det (matrix id) = 1" by (simp add: matrix_id_mat_1) -lemma%important det_matrix_scaleR [simp]: "det (matrix (((*\<^sub>R) r)) :: real^'n^'n) = r ^ CARD('n::finite)" +proposition det_matrix_scaleR [simp]: "det (matrix (((*\<^sub>R) r)) :: real^'n^'n) = r ^ CARD('n::finite)" apply (subst det_diagonal) apply (auto simp: matrix_def mat_def) apply (simp add: cart_eq_inner_axis inner_axis_axis) @@ -463,7 +463,7 @@ exact duplicates by considering the rows/columns as a set. \ -lemma%unimportant det_dependent_rows: +lemma det_dependent_rows: fixes A:: "'a::{field}^'n^'n" assumes d: "vec.dependent (rows A)" shows "det A = 0" @@ -491,23 +491,23 @@ qed qed -lemma%unimportant det_dependent_columns: +lemma det_dependent_columns: assumes d: "vec.dependent (columns (A::real^'n^'n))" shows "det A = 0" by (metis d det_dependent_rows rows_transpose det_transpose) text \Multilinearity and the multiplication formula\ -lemma%unimportant Cart_lambda_cong: "(\x. f x = g x) \ (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)" +lemma Cart_lambda_cong: "(\x. f x = g x) \ (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)" by auto -lemma%unimportant det_linear_row_sum: +lemma det_linear_row_sum: assumes fS: "finite S" shows "det ((\ i. if i = k then sum (a i) S else c i)::'a::comm_ring_1^'n^'n) = sum (\j. det ((\ i. if i = k then a i j else c i)::'a^'n^'n)) S" using fS by (induct rule: finite_induct; simp add: det_row_0 det_row_add cong: if_cong) -lemma%unimportant finite_bounded_functions: +lemma finite_bounded_functions: assumes fS: "finite S" shows "finite {f. (\i \ {1.. (k::nat)}. f i \ S) \ (\i. i \ {1 .. k} \ f i = i)}" proof (induct k) @@ -532,14 +532,14 @@ qed -lemma%important det_linear_rows_sum_lemma: +lemma det_linear_rows_sum_lemma: assumes fS: "finite S" and fT: "finite T" shows "det ((\ i. if i \ T then sum (a i) S else c i):: 'a::comm_ring_1^'n^'n) = sum (\f. det((\ i. if i \ T then a i (f i) else c i)::'a^'n^'n)) {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" using fT -proof%unimportant (induct T arbitrary: a c set: finite) +proof (induct T arbitrary: a c set: finite) case empty have th0: "\x y. (\ i. if i \ {} then x i else y i) = (\ i. y i)" by vector @@ -577,7 +577,7 @@ (auto intro!: cong[OF refl[of det]] simp: vec_eq_iff) qed -lemma%unimportant det_linear_rows_sum: +lemma det_linear_rows_sum: fixes S :: "'n::finite set" assumes fS: "finite S" shows "det (\ i. sum (a i) S) = @@ -589,12 +589,12 @@ show ?thesis by simp qed -lemma%unimportant matrix_mul_sum_alt: +lemma matrix_mul_sum_alt: fixes A B :: "'a::comm_ring_1^'n^'n" shows "A ** B = (\ i. sum (\k. A$i$k *s B $ k) (UNIV :: 'n set))" by (vector matrix_matrix_mult_def sum_component) -lemma%unimportant det_rows_mul: +lemma det_rows_mul: "det((\ i. c i *s a i)::'a::comm_ring_1^'n^'n) = prod (\i. c i) (UNIV:: 'n set) * det((\ i. a i)::'a^'n^'n)" proof (simp add: det_def sum_distrib_left cong add: prod.cong, rule sum.cong) @@ -612,10 +612,10 @@ by (simp add: field_simps) qed rule -lemma%important det_mul: +proposition det_mul: fixes A B :: "'a::comm_ring_1^'n^'n" shows "det (A ** B) = det A * det B" -proof%unimportant - +proof - let ?U = "UNIV :: 'n set" let ?F = "{f. (\i \ ?U. f i \ ?U) \ (\i. i \ ?U \ f i = i)}" let ?PU = "{p. p permutes ?U}" @@ -714,10 +714,10 @@ subsection \Relation to invertibility\ -lemma%important invertible_det_nz: +proposition invertible_det_nz: fixes A::"'a::{field}^'n^'n" shows "invertible A \ det A \ 0" -proof%unimportant (cases "invertible A") +proof (cases "invertible A") case True then obtain B :: "'a^'n^'n" where B: "A ** B = mat 1" unfolding invertible_right_inverse by blast @@ -748,7 +748,7 @@ qed -lemma%unimportant det_nz_iff_inj_gen: +lemma det_nz_iff_inj_gen: fixes f :: "'a::field^'n \ 'a::field^'n" assumes "Vector_Spaces.linear (*s) (*s) f" shows "det (matrix f) \ 0 \ inj f" @@ -763,26 +763,26 @@ by (metis assms invertible_det_nz invertible_left_inverse matrix_compose_gen matrix_id_mat_1) qed -lemma%unimportant det_nz_iff_inj: +lemma det_nz_iff_inj: fixes f :: "real^'n \ real^'n" assumes "linear f" shows "det (matrix f) \ 0 \ inj f" using det_nz_iff_inj_gen[of f] assms unfolding linear_matrix_vector_mul_eq . -lemma%unimportant det_eq_0_rank: +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\ +subsubsection%important \Invertibility of matrices and corresponding linear functions\ -lemma%important matrix_left_invertible_gen: +lemma matrix_left_invertible_gen: fixes f :: "'a::field^'m \ 'a::field^'n" assumes "Vector_Spaces.linear (*s) (*s) f" shows "((\B. B ** matrix f = mat 1) \ (\g. Vector_Spaces.linear (*s) (*s) g \ g \ f = id))" -proof%unimportant safe +proof safe fix B assume 1: "B ** matrix f = mat 1" show "\g. Vector_Spaces.linear (*s) (*s) g \ g \ f = id" @@ -801,12 +801,12 @@ then show "\B. B ** matrix f = mat 1" .. qed -lemma%unimportant matrix_left_invertible: +lemma matrix_left_invertible: "linear f \ ((\B. B ** matrix f = mat 1) \ (\g. linear g \ g \ f = id))" for f::"real^'m \ real^'n" using matrix_left_invertible_gen[of f] by (auto simp: linear_matrix_vector_mul_eq) -lemma%unimportant matrix_right_invertible_gen: +lemma matrix_right_invertible_gen: fixes f :: "'a::field^'m \ 'a^'n" assumes "Vector_Spaces.linear (*s) (*s) f" shows "((\B. matrix f ** B = mat 1) \ (\g. Vector_Spaces.linear (*s) (*s) g \ f \ g = id))" @@ -829,12 +829,12 @@ then show "\B. matrix f ** B = mat 1" .. qed -lemma%unimportant matrix_right_invertible: +lemma matrix_right_invertible: "linear f \ ((\B. matrix f ** B = mat 1) \ (\g. linear g \ f \ g = id))" for f::"real^'m \ real^'n" using matrix_right_invertible_gen[of f] by (auto simp: linear_matrix_vector_mul_eq) -lemma%unimportant matrix_invertible_gen: +lemma matrix_invertible_gen: fixes f :: "'a::field^'m \ 'a::field^'n" assumes "Vector_Spaces.linear (*s) (*s) f" shows "invertible (matrix f) \ (\g. Vector_Spaces.linear (*s) (*s) g \ f \ g = id \ g \ f = id)" @@ -847,13 +847,13 @@ by (metis assms invertible_def matrix_compose_gen matrix_id_mat_1) qed -lemma%unimportant matrix_invertible: +lemma matrix_invertible: "linear f \ invertible (matrix f) \ (\g. linear g \ f \ g = id \ g \ f = id)" for f::"real^'m \ real^'n" using matrix_invertible_gen[of f] by (auto simp: linear_matrix_vector_mul_eq) -lemma%unimportant invertible_eq_bij: +lemma invertible_eq_bij: fixes m :: "'a::field^'m^'n" shows "invertible m \ bij ((*v) m)" using matrix_invertible_gen[OF matrix_vector_mul_linear_gen, of m, simplified matrix_of_matrix_vector_mul] @@ -863,13 +863,13 @@ subsection \Cramer's rule\ -lemma%important cramer_lemma_transpose: +lemma cramer_lemma_transpose: fixes A:: "'a::{field}^'n^'n" and x :: "'a::{field}^'n" shows "det ((\ i. if i = k then sum (\i. x$i *s row i A) (UNIV::'n set) else row i A)::'a::{field}^'n^'n) = x$k * det A" (is "?lhs = ?rhs") -proof%unimportant - +proof - let ?U = "UNIV :: 'n set" let ?Uk = "?U - {k}" have U: "?U = insert k ?Uk" @@ -899,10 +899,10 @@ done qed -lemma%important cramer_lemma: +proposition cramer_lemma: fixes A :: "'a::{field}^'n^'n" shows "det((\ i j. if j = k then (A *v x)$i else A$i$j):: 'a::{field}^'n^'n) = x$k * det A" -proof%unimportant - +proof - let ?U = "UNIV :: 'n set" have *: "\c. sum (\i. c i *s row i (transpose A)) ?U = sum (\i. c i *s column i A) ?U" by (auto intro: sum.cong) @@ -916,11 +916,11 @@ done qed -lemma%important cramer: +proposition cramer: fixes A ::"'a::{field}^'n^'n" assumes d0: "det A \ 0" shows "A *v x = b \ x = (\ k. det(\ i j. if j=k then b$i else A$i$j) / det A)" -proof%unimportant - +proof - from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" unfolding invertible_det_nz[symmetric] invertible_def by blast @@ -941,10 +941,10 @@ by auto qed -lemma%unimportant det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" +lemma det_1: "det (A::'a::comm_ring_1^1^1) = A$1$1" by (simp add: det_def sign_id) -lemma%unimportant det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" +lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" proof - have f12: "finite {2::2}" "1 \ {2::2}" by auto show ?thesis @@ -954,7 +954,7 @@ by (simp add: sign_swap_id sign_id swap_id_eq) qed -lemma%unimportant det_3: +lemma det_3: "det (A::'a::comm_ring_1^3^3) = A$1$1 * A$2$2 * A$3$3 + A$1$2 * A$2$3 * A$3$1 + @@ -976,11 +976,11 @@ by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) qed -lemma%important det_orthogonal_matrix: +proposition det_orthogonal_matrix: fixes Q:: "'a::linordered_idom^'n^'n" assumes oQ: "orthogonal_matrix Q" shows "det Q = 1 \ det Q = - 1" -proof%unimportant - +proof - have "Q ** transpose Q = mat 1" by (metis oQ orthogonal_matrix_def) then have "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)" @@ -991,7 +991,7 @@ by (simp add: square_eq_1_iff) qed -lemma%important orthogonal_transformation_det [simp]: +proposition orthogonal_transformation_det [simp]: fixes f :: "real^'n \ real^'n" shows "orthogonal_transformation f \ \det (matrix f)\ = 1" using%unimportant det_orthogonal_matrix orthogonal_transformation_matrix by fastforce @@ -1001,14 +1001,14 @@ definition%important "rotation_matrix Q \ orthogonal_matrix Q \ det Q = 1" definition%important "rotoinversion_matrix Q \ orthogonal_matrix Q \ det Q = - 1" -lemma%important orthogonal_rotation_or_rotoinversion: +lemma orthogonal_rotation_or_rotoinversion: fixes Q :: "'a::linordered_idom^'n^'n" shows " orthogonal_matrix Q \ rotation_matrix Q \ rotoinversion_matrix Q" - by%unimportant (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) + by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) text\ Slightly stronger results giving rotation, but only in two or more dimensions\ -lemma%unimportant rotation_matrix_exists_basis: +lemma rotation_matrix_exists_basis: fixes a :: "real^'n" assumes 2: "2 \ CARD('n)" and "norm a = 1" obtains A where "rotation_matrix A" "A *v (axis k 1) = a" @@ -1049,7 +1049,7 @@ qed qed -lemma%unimportant rotation_exists_1: +lemma rotation_exists_1: fixes a :: "real^'n" assumes "2 \ CARD('n)" "norm a = 1" "norm b = 1" obtains f where "orthogonal_transformation f" "det(matrix f) = 1" "f a = b" @@ -1072,7 +1072,7 @@ qed qed -lemma%unimportant rotation_exists: +lemma rotation_exists: fixes a :: "real^'n" assumes 2: "2 \ CARD('n)" and eq: "norm a = norm b" obtains f where "orthogonal_transformation f" "det(matrix f) = 1" "f a = b" @@ -1094,7 +1094,7 @@ with f show thesis .. qed -lemma%unimportant rotation_rightward_line: +lemma rotation_rightward_line: fixes a :: "real^'n" obtains f where "orthogonal_transformation f" "2 \ CARD('n) \ det(matrix f) = 1" "f(norm a *\<^sub>R axis k 1) = a" diff -r 749aaeb40788 -r be6634e99e09 src/HOL/Analysis/Finite_Cartesian_Product.thy --- a/src/HOL/Analysis/Finite_Cartesian_Product.thy Fri Dec 14 14:33:26 2018 +0100 +++ b/src/HOL/Analysis/Finite_Cartesian_Product.thy Tue Jan 22 21:13:23 2019 +0000 @@ -2,7 +2,7 @@ Author: Amine Chaieb, University of Cambridge *) -section%important \Definition of Finite Cartesian Product Type\ +section \Definition of Finite Cartesian Product Type\ theory Finite_Cartesian_Product imports @@ -104,10 +104,10 @@ by (auto elim!: countableE) qed -lemma%important infinite_UNIV_vec: +lemma infinite_UNIV_vec: assumes "infinite (UNIV :: 'a set)" shows "infinite (UNIV :: ('a^'b) set)" -proof%unimportant (subst bij_betw_finite) +proof (subst bij_betw_finite) show "bij_betw vec_nth UNIV (Pi (UNIV :: 'b set) (\_. UNIV :: 'a set))" by (intro bij_betwI[of _ _ _ vec_lambda]) (auto simp: vec_eq_iff) have "infinite (PiE (UNIV :: 'b set) (\_. UNIV :: 'a set))" (is "infinite ?A") @@ -125,9 +125,9 @@ finally show "infinite (Pi (UNIV :: 'b set) (\_. UNIV :: 'a set))" . qed -lemma%important CARD_vec [simp]: +proposition CARD_vec [simp]: "CARD('a^'b) = CARD('a) ^ CARD('b)" -proof%unimportant (cases "finite (UNIV :: 'a set)") +proof (cases "finite (UNIV :: 'a set)") case True show ?thesis proof (subst bij_betw_same_card) @@ -143,7 +143,7 @@ qed qed (simp_all add: infinite_UNIV_vec) -lemma%unimportant countable_vector: +lemma countable_vector: fixes B:: "'n::finite \ 'a set" assumes "\i. countable (B i)" shows "countable {V. \i::'n::finite. V $ i \ B i}" @@ -299,7 +299,7 @@ definition%important "scaleR \ (\ r x. (\ i. scaleR r (x$i)))" -lemma%important vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)" +lemma vector_scaleR_component [simp]: "(scaleR r x)$i = scaleR r (x$i)" unfolding scaleR_vec_def by simp instance%unimportant @@ -318,7 +318,7 @@ (\x\S. \A. (\i. open (A i) \ x$i \ A i) \ (\y. (\i. y$i \ A i) \ y \ S))" -instance proof%unimportant +instance%unimportant proof show "open (UNIV :: ('a ^ 'b) set)" unfolding open_vec_def by auto next @@ -346,30 +346,30 @@ end -lemma%unimportant open_vector_box: "\i. open (S i) \ open {x. \i. x $ i \ S i}" +lemma open_vector_box: "\i. open (S i) \ open {x. \i. x $ i \ S i}" unfolding open_vec_def by auto -lemma%unimportant open_vimage_vec_nth: "open S \ open ((\x. x $ i) -` S)" +lemma open_vimage_vec_nth: "open S \ open ((\x. x $ i) -` S)" unfolding open_vec_def apply clarify apply (rule_tac x="\k. if k = i then S else UNIV" in exI, simp) done -lemma%unimportant closed_vimage_vec_nth: "closed S \ closed ((\x. x $ i) -` S)" +lemma closed_vimage_vec_nth: "closed S \ closed ((\x. x $ i) -` S)" unfolding closed_open vimage_Compl [symmetric] by (rule open_vimage_vec_nth) -lemma%unimportant closed_vector_box: "\i. closed (S i) \ closed {x. \i. x $ i \ S i}" +lemma closed_vector_box: "\i. closed (S i) \ closed {x. \i. x $ i \ S i}" proof - have "{x. \i. x $ i \ S i} = (\i. (\x. x $ i) -` S i)" by auto thus "\i. closed (S i) \ closed {x. \i. x $ i \ S i}" by (simp add: closed_INT closed_vimage_vec_nth) qed -lemma%important tendsto_vec_nth [tendsto_intros]: +lemma tendsto_vec_nth [tendsto_intros]: assumes "((\x. f x) \ a) net" shows "((\x. f x $ i) \ a $ i) net" -proof%unimportant (rule topological_tendstoI) +proof (rule topological_tendstoI) fix S assume "open S" "a $ i \ S" then have "open ((\y. y $ i) -` S)" "a \ ((\y. y $ i) -` S)" by (simp_all add: open_vimage_vec_nth) @@ -379,13 +379,13 @@ by simp qed -lemma%unimportant isCont_vec_nth [simp]: "isCont f a \ isCont (\x. f x $ i) a" +lemma isCont_vec_nth [simp]: "isCont f a \ isCont (\x. f x $ i) a" unfolding isCont_def by (rule tendsto_vec_nth) -lemma%important vec_tendstoI: +lemma vec_tendstoI: assumes "\i. ((\x. f x $ i) \ a $ i) net" shows "((\x. f x) \ a) net" -proof%unimportant (rule topological_tendstoI) +proof (rule topological_tendstoI) fix S assume "open S" and "a \ S" then obtain A where A: "\i. open (A i)" "\i. a $ i \ A i" and S: "\y. \i. y $ i \ A i \ y \ S" @@ -398,13 +398,13 @@ by (rule eventually_mono, simp add: S) qed -lemma%unimportant tendsto_vec_lambda [tendsto_intros]: +lemma tendsto_vec_lambda [tendsto_intros]: assumes "\i. ((\x. f x i) \ a i) net" shows "((\x. \ i. f x i) \ (\ i. a i)) net" using assms by (simp add: vec_tendstoI) -lemma%important open_image_vec_nth: assumes "open S" shows "open ((\x. x $ i) ` S)" -proof%unimportant (rule openI) +lemma open_image_vec_nth: assumes "open S" shows "open ((\x. x $ i) ` S)" +proof (rule openI) fix a assume "a \ (\x. x $ i) ` S" then obtain z where "a = z $ i" and "z \ S" .. then obtain A where A: "\i. open (A i) \ z $ i \ A i" @@ -418,7 +418,7 @@ then show "\T. open T \ a \ T \ T \ (\x. x $ i) ` S" by - (rule exI) qed -instance vec :: (perfect_space, finite) perfect_space +instance%unimportant vec :: (perfect_space, finite) perfect_space proof fix x :: "'a ^ 'b" show "\ open {x}" proof @@ -458,10 +458,10 @@ instantiation%unimportant vec :: (metric_space, finite) metric_space begin -lemma%important dist_vec_nth_le: "dist (x $ i) (y $ i) \ dist x y" +proposition dist_vec_nth_le: "dist (x $ i) (y $ i) \ dist x y" unfolding dist_vec_def by (rule member_le_L2_set) simp_all -instance proof%unimportant +instance proof fix x y :: "'a ^ 'b" show "dist x y = 0 \ x = y" unfolding dist_vec_def @@ -520,15 +520,15 @@ end -lemma%important Cauchy_vec_nth: +lemma Cauchy_vec_nth: "Cauchy (\n. X n) \ Cauchy (\n. X n $ i)" unfolding Cauchy_def by (fast intro: le_less_trans [OF dist_vec_nth_le]) -lemma%important vec_CauchyI: +lemma vec_CauchyI: fixes X :: "nat \ 'a::metric_space ^ 'n" assumes X: "\i. Cauchy (\n. X n $ i)" shows "Cauchy (\n. X n)" -proof%unimportant (rule metric_CauchyI) +proof (rule metric_CauchyI) fix r :: real assume "0 < r" hence "0 < r / of_nat CARD('n)" (is "0 < ?s") by simp define N where "N i = (LEAST N. \m\N. \n\N. dist (X m $ i) (X n $ i) < ?s)" for i @@ -557,7 +557,7 @@ then show "\M. \m\M. \n\M. dist (X m) (X n) < r" .. qed -instance vec :: (complete_space, finite) complete_space +instance%unimportant vec :: (complete_space, finite) complete_space proof fix X :: "nat \ 'a ^ 'b" assume "Cauchy X" have "\i. (\n. X n $ i) \ lim (\n. X n $ i)" @@ -579,7 +579,7 @@ definition%important "sgn (x::'a^'b) = scaleR (inverse (norm x)) x" -instance proof%unimportant +instance%unimportant proof fix a :: real and x y :: "'a ^ 'b" show "norm x = 0 \ x = 0" unfolding norm_vec_def @@ -601,32 +601,32 @@ end -lemma%unimportant norm_nth_le: "norm (x $ i) \ norm x" +lemma norm_nth_le: "norm (x $ i) \ norm x" unfolding norm_vec_def by (rule member_le_L2_set) simp_all -lemma%important norm_le_componentwise_cart: +lemma norm_le_componentwise_cart: fixes x :: "'a::real_normed_vector^'n" assumes "\i. norm(x$i) \ norm(y$i)" shows "norm x \ norm y" unfolding%unimportant norm_vec_def by%unimportant (rule L2_set_mono) (auto simp: assms) -lemma%unimportant component_le_norm_cart: "\x$i\ \ norm x" +lemma component_le_norm_cart: "\x$i\ \ norm x" apply (simp add: norm_vec_def) apply (rule member_le_L2_set, simp_all) done -lemma%unimportant norm_bound_component_le_cart: "norm x \ e ==> \x$i\ \ e" +lemma norm_bound_component_le_cart: "norm x \ e ==> \x$i\ \ e" by (metis component_le_norm_cart order_trans) -lemma%unimportant norm_bound_component_lt_cart: "norm x < e ==> \x$i\ < e" +lemma norm_bound_component_lt_cart: "norm x < e ==> \x$i\ < e" by (metis component_le_norm_cart le_less_trans) -lemma%unimportant norm_le_l1_cart: "norm x \ sum(\i. \x$i\) UNIV" +lemma norm_le_l1_cart: "norm x \ sum(\i. \x$i\) UNIV" by (simp add: norm_vec_def L2_set_le_sum) -lemma%unimportant bounded_linear_vec_nth[intro]: "bounded_linear (\x. x $ i)" +lemma bounded_linear_vec_nth[intro]: "bounded_linear (\x. x $ i)" apply standard apply (rule vector_add_component) apply (rule vector_scaleR_component) @@ -643,7 +643,7 @@ definition%important "inner x y = sum (\i. inner (x$i) (y$i)) UNIV" -instance proof%unimportant +instance%unimportant proof fix r :: real and x y z :: "'a ^ 'b" show "inner x y = inner y x" unfolding inner_vec_def @@ -674,13 +674,13 @@ definition%important "axis k x = (\ i. if i = k then x else 0)" -lemma%unimportant axis_nth [simp]: "axis i x $ i = x" +lemma axis_nth [simp]: "axis i x $ i = x" unfolding axis_def by simp -lemma%unimportant axis_eq_axis: "axis i x = axis j y \ x = y \ i = j \ x = 0 \ y = 0" +lemma axis_eq_axis: "axis i x = axis j y \ x = y \ i = j \ x = 0 \ y = 0" unfolding axis_def vec_eq_iff by auto -lemma%unimportant inner_axis_axis: +lemma inner_axis_axis: "inner (axis i x) (axis j y) = (if i = j then inner x y else 0)" unfolding inner_vec_def apply (cases "i = j") @@ -690,10 +690,10 @@ apply (rule sum.neutral, simp add: axis_def) done -lemma%unimportant inner_axis: "inner x (axis i y) = inner (x $ i) y" +lemma inner_axis: "inner x (axis i y) = inner (x $ i) y" by (simp add: inner_vec_def axis_def sum.remove [where x=i]) -lemma%unimportant inner_axis': "inner(axis i y) x = inner y (x $ i)" +lemma inner_axis': "inner(axis i y) x = inner y (x $ i)" by (simp add: inner_axis inner_commute) instantiation%unimportant vec :: (euclidean_space, finite) euclidean_space @@ -701,7 +701,7 @@ definition%important "Basis = (\i. \u\Basis. {axis i u})" -instance proof%unimportant +instance%unimportant proof show "(Basis :: ('a ^ 'b) set) \ {}" unfolding Basis_vec_def by simp next @@ -720,8 +720,8 @@ by (simp add: inner_axis euclidean_all_zero_iff vec_eq_iff) qed -lemma%important DIM_cart [simp]: "DIM('a^'b) = CARD('b) * DIM('a)" -proof%unimportant - +proposition DIM_cart [simp]: "DIM('a^'b) = CARD('b) * DIM('a)" +proof - have "card (\i::'b. \u::'a\Basis. {axis i u}) = (\i::'b\UNIV. card (\u::'a\Basis. {axis i u}))" by (rule card_UN_disjoint) (auto simp: axis_eq_axis) also have "... = CARD('b) * DIM('a)" @@ -732,30 +732,30 @@ end -lemma%unimportant norm_axis_1 [simp]: "norm (axis m (1::real)) = 1" +lemma norm_axis_1 [simp]: "norm (axis m (1::real)) = 1" by (simp add: inner_axis' norm_eq_1) -lemma%important sum_norm_allsubsets_bound_cart: +lemma sum_norm_allsubsets_bound_cart: fixes f:: "'a \ real ^'n" assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (sum f Q) \ e" shows "sum (\x. norm (f x)) P \ 2 * real CARD('n) * e" - using%unimportant sum_norm_allsubsets_bound[OF assms] - by%unimportant simp + using sum_norm_allsubsets_bound[OF assms] + by simp -lemma%unimportant cart_eq_inner_axis: "a $ i = inner a (axis i 1)" +lemma cart_eq_inner_axis: "a $ i = inner a (axis i 1)" by (simp add: inner_axis) -lemma%unimportant axis_eq_0_iff [simp]: +lemma axis_eq_0_iff [simp]: shows "axis m x = 0 \ x = 0" by (simp add: axis_def vec_eq_iff) -lemma%unimportant axis_in_Basis_iff [simp]: "axis i a \ Basis \ a \ Basis" +lemma axis_in_Basis_iff [simp]: "axis i a \ Basis \ a \ Basis" by (auto simp: Basis_vec_def axis_eq_axis) text\Mapping each basis element to the corresponding finite index\ definition axis_index :: "('a::comm_ring_1)^'n \ 'n" where "axis_index v \ SOME i. v = axis i 1" -lemma%unimportant axis_inverse: +lemma axis_inverse: fixes v :: "real^'n" assumes "v \ Basis" shows "\i. v = axis i 1" @@ -766,20 +766,20 @@ by (force simp add: vec_eq_iff) qed -lemma%unimportant axis_index: +lemma axis_index: fixes v :: "real^'n" assumes "v \ Basis" shows "v = axis (axis_index v) 1" by (metis (mono_tags) assms axis_inverse axis_index_def someI_ex) -lemma%unimportant axis_index_axis [simp]: +lemma axis_index_axis [simp]: fixes UU :: "real^'n" shows "(axis_index (axis u 1 :: real^'n)) = (u::'n)" by (simp add: axis_eq_axis axis_index_def) subsection%unimportant \A naive proof procedure to lift really trivial arithmetic stuff from the basis of the vector space\ -lemma%unimportant sum_cong_aux: +lemma sum_cong_aux: "(\x. x \ A \ f x = g x) \ sum f A = sum g A" by (auto intro: sum.cong) @@ -811,22 +811,22 @@ end \ "lift trivial vector statements to real arith statements" -lemma%unimportant vec_0[simp]: "vec 0 = 0" by vector -lemma%unimportant vec_1[simp]: "vec 1 = 1" by vector +lemma vec_0[simp]: "vec 0 = 0" by vector +lemma vec_1[simp]: "vec 1 = 1" by vector -lemma%unimportant vec_inj[simp]: "vec x = vec y \ x = y" by vector +lemma vec_inj[simp]: "vec x = vec y \ x = y" by vector -lemma%unimportant vec_in_image_vec: "vec x \ (vec ` S) \ x \ S" by auto +lemma vec_in_image_vec: "vec x \ (vec ` S) \ x \ S" by auto -lemma%unimportant vec_add: "vec(x + y) = vec x + vec y" by vector -lemma%unimportant vec_sub: "vec(x - y) = vec x - vec y" by vector -lemma%unimportant vec_cmul: "vec(c * x) = c *s vec x " by vector -lemma%unimportant vec_neg: "vec(- x) = - vec x " by vector +lemma vec_add: "vec(x + y) = vec x + vec y" by vector +lemma vec_sub: "vec(x - y) = vec x - vec y" by vector +lemma vec_cmul: "vec(c * x) = c *s vec x " by vector +lemma vec_neg: "vec(- x) = - vec x " by vector -lemma%unimportant vec_scaleR: "vec(c * x) = c *\<^sub>R vec x" +lemma vec_scaleR: "vec(c * x) = c *\<^sub>R vec x" by vector -lemma%unimportant vec_sum: +lemma vec_sum: assumes "finite S" shows "vec(sum f S) = sum (vec \ f) S" using assms @@ -840,16 +840,16 @@ text\Obvious "component-pushing".\ -lemma%unimportant vec_component [simp]: "vec x $ i = x" +lemma vec_component [simp]: "vec x $ i = x" by vector -lemma%unimportant vector_mult_component [simp]: "(x * y)$i = x$i * y$i" +lemma vector_mult_component [simp]: "(x * y)$i = x$i * y$i" by vector -lemma%unimportant vector_smult_component [simp]: "(c *s y)$i = c * (y$i)" +lemma vector_smult_component [simp]: "(c *s y)$i = c * (y$i)" by vector -lemma%unimportant cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector +lemma cond_component: "(if b then x else y)$i = (if b then x$i else y$i)" by vector lemmas%unimportant vector_component = vec_component vector_add_component vector_mult_component @@ -999,7 +999,7 @@ definition%important map_matrix::"('a \ 'b) \ (('a, 'i::finite)vec, 'j::finite) vec \ (('b, 'i)vec, 'j) vec" where "map_matrix f x = (\ i j. f (x $ i $ j))" -lemma%unimportant nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)" +lemma nth_map_matrix[simp]: "map_matrix f x $ i $ j = f (x $ i $ j)" by (simp add: map_matrix_def) definition%important matrix_matrix_mult :: "('a::semiring_1) ^'n^'m \ 'a ^'p^'n \ 'a ^ 'p ^'m" @@ -1022,17 +1022,17 @@ definition%unimportant "rows(A::'a^'n^'m) = { row i A | i. i \ (UNIV :: 'm set)}" definition%unimportant "columns(A::'a^'n^'m) = { column i A | i. i \ (UNIV :: 'n set)}" -lemma%unimportant times0_left [simp]: "(0::'a::semiring_1^'n^'m) ** (A::'a ^'p^'n) = 0" +lemma times0_left [simp]: "(0::'a::semiring_1^'n^'m) ** (A::'a ^'p^'n) = 0" by (simp add: matrix_matrix_mult_def zero_vec_def) -lemma%unimportant times0_right [simp]: "(A::'a::semiring_1^'n^'m) ** (0::'a ^'p^'n) = 0" +lemma times0_right [simp]: "(A::'a::semiring_1^'n^'m) ** (0::'a ^'p^'n) = 0" by (simp add: matrix_matrix_mult_def zero_vec_def) -lemma%unimportant mat_0[simp]: "mat 0 = 0" by (vector mat_def) -lemma%unimportant matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)" +lemma mat_0[simp]: "mat 0 = 0" by (vector mat_def) +lemma matrix_add_ldistrib: "(A ** (B + C)) = (A ** B) + (A ** C)" by (vector matrix_matrix_mult_def sum.distrib[symmetric] field_simps) -lemma%unimportant matrix_mul_lid [simp]: +lemma matrix_mul_lid [simp]: fixes A :: "'a::semiring_1 ^ 'm ^ 'n" shows "mat 1 ** A = A" apply (simp add: matrix_matrix_mult_def mat_def) @@ -1041,7 +1041,7 @@ mult_1_left mult_zero_left if_True UNIV_I) done -lemma%unimportant matrix_mul_rid [simp]: +lemma matrix_mul_rid [simp]: fixes A :: "'a::semiring_1 ^ 'm ^ 'n" shows "A ** mat 1 = A" apply (simp add: matrix_matrix_mult_def mat_def) @@ -1050,47 +1050,47 @@ mult_1_right mult_zero_right if_True UNIV_I cong: if_cong) done -lemma%unimportant matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" +proposition matrix_mul_assoc: "A ** (B ** C) = (A ** B) ** C" apply (vector matrix_matrix_mult_def sum_distrib_left sum_distrib_right mult.assoc) apply (subst sum.swap) apply simp done -lemma%unimportant matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" +proposition matrix_vector_mul_assoc: "A *v (B *v x) = (A ** B) *v x" apply (vector matrix_matrix_mult_def matrix_vector_mult_def sum_distrib_left sum_distrib_right mult.assoc) apply (subst sum.swap) apply simp done -lemma%unimportant scalar_matrix_assoc: +proposition scalar_matrix_assoc: fixes A :: "('a::real_algebra_1)^'m^'n" shows "k *\<^sub>R (A ** B) = (k *\<^sub>R A) ** B" by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff scaleR_sum_right) -lemma%unimportant matrix_scalar_ac: +proposition matrix_scalar_ac: fixes A :: "('a::real_algebra_1)^'m^'n" shows "A ** (k *\<^sub>R B) = k *\<^sub>R A ** B" by (simp add: matrix_matrix_mult_def sum_distrib_left mult_ac vec_eq_iff) -lemma%unimportant matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)" +lemma matrix_vector_mul_lid [simp]: "mat 1 *v x = (x::'a::semiring_1 ^ 'n)" apply (vector matrix_vector_mult_def mat_def) apply (simp add: if_distrib if_distribR sum.delta' cong del: if_weak_cong) done -lemma%unimportant matrix_transpose_mul: +lemma matrix_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)" by (simp add: matrix_matrix_mult_def transpose_def vec_eq_iff mult.commute) -lemma%unimportant matrix_mult_transpose_dot_column: +lemma matrix_mult_transpose_dot_column: shows "transpose A ** A = (\ i j. inner (column i A) (column j A))" by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def column_def inner_vec_def) -lemma%unimportant matrix_mult_transpose_dot_row: +lemma matrix_mult_transpose_dot_row: shows "A ** transpose A = (\ i j. inner (row i A) (row j A))" by (simp add: matrix_matrix_mult_def vec_eq_iff transpose_def row_def inner_vec_def) -lemma%unimportant matrix_eq: +lemma matrix_eq: fixes A B :: "'a::semiring_1 ^ 'n ^ 'm" shows "A = B \ (\x. A *v x = B *v x)" (is "?lhs \ ?rhs") apply auto @@ -1103,49 +1103,49 @@ sum.delta[OF finite] cong del: if_weak_cong) done -lemma%unimportant matrix_vector_mul_component: "(A *v x)$k = inner (A$k) x" +lemma matrix_vector_mul_component: "(A *v x)$k = inner (A$k) x" by (simp add: matrix_vector_mult_def inner_vec_def) -lemma%unimportant dot_lmul_matrix: "inner ((x::real ^_) v* A) y = inner x (A *v y)" +lemma dot_lmul_matrix: "inner ((x::real ^_) v* A) y = inner x (A *v y)" apply (simp add: inner_vec_def matrix_vector_mult_def vector_matrix_mult_def sum_distrib_right sum_distrib_left ac_simps) apply (subst sum.swap) apply simp done -lemma%unimportant transpose_mat [simp]: "transpose (mat n) = mat n" +lemma transpose_mat [simp]: "transpose (mat n) = mat n" by (vector transpose_def mat_def) -lemma%unimportant transpose_transpose [simp]: "transpose(transpose A) = A" +lemma transpose_transpose [simp]: "transpose(transpose A) = A" by (vector transpose_def) -lemma%unimportant row_transpose [simp]: "row i (transpose A) = column i A" +lemma row_transpose [simp]: "row i (transpose A) = column i A" by (simp add: row_def column_def transpose_def vec_eq_iff) -lemma%unimportant column_transpose [simp]: "column i (transpose A) = row i A" +lemma column_transpose [simp]: "column i (transpose A) = row i A" by (simp add: row_def column_def transpose_def vec_eq_iff) -lemma%unimportant rows_transpose [simp]: "rows(transpose A) = columns A" +lemma rows_transpose [simp]: "rows(transpose A) = columns A" by (auto simp add: rows_def columns_def intro: set_eqI) -lemma%unimportant columns_transpose [simp]: "columns(transpose A) = rows A" +lemma columns_transpose [simp]: "columns(transpose A) = rows A" by (metis transpose_transpose rows_transpose) -lemma%unimportant transpose_scalar: "transpose (k *\<^sub>R A) = k *\<^sub>R transpose A" +lemma transpose_scalar: "transpose (k *\<^sub>R A) = k *\<^sub>R transpose A" unfolding transpose_def by (simp add: vec_eq_iff) -lemma%unimportant transpose_iff [iff]: "transpose A = transpose B \ A = B" +lemma transpose_iff [iff]: "transpose A = transpose B \ A = B" by (metis transpose_transpose) -lemma%unimportant matrix_mult_sum: +lemma matrix_mult_sum: "(A::'a::comm_semiring_1^'n^'m) *v x = sum (\i. (x$i) *s column i A) (UNIV:: 'n set)" by (simp add: matrix_vector_mult_def vec_eq_iff column_def mult.commute) -lemma%unimportant vector_componentwise: +lemma vector_componentwise: "(x::'a::ring_1^'n) = (\ j. \i\UNIV. (x$i) * (axis i 1 :: 'a^'n) $ j)" by (simp add: axis_def if_distrib sum.If_cases vec_eq_iff) -lemma%unimportant basis_expansion: "sum (\i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)" +lemma basis_expansion: "sum (\i. (x$i) *s axis i 1) UNIV = (x::('a::ring_1) ^'n)" by (auto simp add: axis_def vec_eq_iff if_distrib sum.If_cases cong del: if_weak_cong) @@ -1154,51 +1154,51 @@ definition%important matrix :: "('a::{plus,times, one, zero}^'m \ 'a ^ 'n) \ 'a^'m^'n" where "matrix f = (\ i j. (f(axis j 1))$i)" -lemma%unimportant matrix_id_mat_1: "matrix id = mat 1" +lemma matrix_id_mat_1: "matrix id = mat 1" by (simp add: mat_def matrix_def axis_def) -lemma%unimportant matrix_scaleR: "(matrix ((*\<^sub>R) r)) = mat r" +lemma matrix_scaleR: "(matrix ((*\<^sub>R) r)) = mat r" by (simp add: mat_def matrix_def axis_def if_distrib cong: if_cong) -lemma%unimportant matrix_vector_mul_linear[intro, simp]: "linear (\x. A *v (x::'a::real_algebra_1 ^ _))" +lemma matrix_vector_mul_linear[intro, simp]: "linear (\x. A *v (x::'a::real_algebra_1 ^ _))" by (simp add: linear_iff matrix_vector_mult_def vec_eq_iff field_simps sum_distrib_left sum.distrib scaleR_right.sum) -lemma%unimportant vector_matrix_left_distrib [algebra_simps]: +lemma vector_matrix_left_distrib [algebra_simps]: shows "(x + y) v* A = x v* A + y v* A" unfolding vector_matrix_mult_def by (simp add: algebra_simps sum.distrib vec_eq_iff) -lemma%unimportant matrix_vector_right_distrib [algebra_simps]: +lemma matrix_vector_right_distrib [algebra_simps]: "A *v (x + y) = A *v x + A *v y" by (vector matrix_vector_mult_def sum.distrib distrib_left) -lemma%unimportant matrix_vector_mult_diff_distrib [algebra_simps]: +lemma matrix_vector_mult_diff_distrib [algebra_simps]: fixes A :: "'a::ring_1^'n^'m" shows "A *v (x - y) = A *v x - A *v y" by (vector matrix_vector_mult_def sum_subtractf right_diff_distrib) -lemma%unimportant matrix_vector_mult_scaleR[algebra_simps]: +lemma matrix_vector_mult_scaleR[algebra_simps]: fixes A :: "real^'n^'m" shows "A *v (c *\<^sub>R x) = c *\<^sub>R (A *v x)" using linear_iff matrix_vector_mul_linear by blast -lemma%unimportant matrix_vector_mult_0_right [simp]: "A *v 0 = 0" +lemma matrix_vector_mult_0_right [simp]: "A *v 0 = 0" by (simp add: matrix_vector_mult_def vec_eq_iff) -lemma%unimportant matrix_vector_mult_0 [simp]: "0 *v w = 0" +lemma matrix_vector_mult_0 [simp]: "0 *v w = 0" by (simp add: matrix_vector_mult_def vec_eq_iff) -lemma%unimportant matrix_vector_mult_add_rdistrib [algebra_simps]: +lemma matrix_vector_mult_add_rdistrib [algebra_simps]: "(A + B) *v x = (A *v x) + (B *v x)" by (vector matrix_vector_mult_def sum.distrib distrib_right) -lemma%unimportant matrix_vector_mult_diff_rdistrib [algebra_simps]: +lemma matrix_vector_mult_diff_rdistrib [algebra_simps]: fixes A :: "'a :: ring_1^'n^'m" shows "(A - B) *v x = (A *v x) - (B *v x)" by (vector matrix_vector_mult_def sum_subtractf left_diff_distrib) -lemma%unimportant matrix_vector_column: +lemma matrix_vector_column: "(A::'a::comm_semiring_1^'n^_) *v x = sum (\i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)" by (simp add: matrix_vector_mult_def transpose_def vec_eq_iff mult.commute) @@ -1211,17 +1211,17 @@ "matrix_inv(A:: 'a::semiring_1^'n^'m) = (SOME A'::'a^'m^'n. A ** A' = mat 1 \ A' ** A = mat 1)" -lemma%unimportant inj_matrix_vector_mult: +lemma inj_matrix_vector_mult: fixes A::"'a::field^'n^'m" assumes "invertible A" shows "inj ((*v) A)" by (metis assms inj_on_inverseI invertible_def matrix_vector_mul_assoc matrix_vector_mul_lid) -lemma%important scalar_invertible: +lemma scalar_invertible: fixes A :: "('a::real_algebra_1)^'m^'n" assumes "k \ 0" and "invertible A" shows "invertible (k *\<^sub>R A)" -proof%unimportant - +proof - obtain A' where "A ** A' = mat 1" and "A' ** A = mat 1" using assms unfolding invertible_def by auto with \k \ 0\ @@ -1231,50 +1231,50 @@ unfolding invertible_def by auto qed -lemma%unimportant scalar_invertible_iff: +proposition scalar_invertible_iff: fixes A :: "('a::real_algebra_1)^'m^'n" assumes "k \ 0" and "invertible A" shows "invertible (k *\<^sub>R A) \ k \ 0 \ invertible A" by (simp add: assms scalar_invertible) -lemma%unimportant vector_transpose_matrix [simp]: "x v* transpose A = A *v x" +lemma vector_transpose_matrix [simp]: "x v* transpose A = A *v x" unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def by simp -lemma%unimportant transpose_matrix_vector [simp]: "transpose A *v x = x v* A" +lemma transpose_matrix_vector [simp]: "transpose A *v x = x v* A" unfolding transpose_def vector_matrix_mult_def matrix_vector_mult_def by simp -lemma%unimportant vector_scalar_commute: +lemma vector_scalar_commute: fixes A :: "'a::{field}^'m^'n" shows "A *v (c *s x) = c *s (A *v x)" by (simp add: vector_scalar_mult_def matrix_vector_mult_def mult_ac sum_distrib_left) -lemma%unimportant scalar_vector_matrix_assoc: +lemma scalar_vector_matrix_assoc: fixes k :: "'a::{field}" and x :: "'a::{field}^'n" and A :: "'a^'m^'n" shows "(k *s x) v* A = k *s (x v* A)" by (metis transpose_matrix_vector vector_scalar_commute) -lemma%unimportant vector_matrix_mult_0 [simp]: "0 v* A = 0" +lemma vector_matrix_mult_0 [simp]: "0 v* A = 0" unfolding vector_matrix_mult_def by (simp add: zero_vec_def) -lemma%unimportant vector_matrix_mult_0_right [simp]: "x v* 0 = 0" +lemma vector_matrix_mult_0_right [simp]: "x v* 0 = 0" unfolding vector_matrix_mult_def by (simp add: zero_vec_def) -lemma%unimportant vector_matrix_mul_rid [simp]: +lemma vector_matrix_mul_rid [simp]: fixes v :: "('a::semiring_1)^'n" shows "v v* mat 1 = v" by (metis matrix_vector_mul_lid transpose_mat vector_transpose_matrix) -lemma%unimportant scaleR_vector_matrix_assoc: +lemma scaleR_vector_matrix_assoc: fixes k :: real and x :: "real^'n" and A :: "real^'m^'n" shows "(k *\<^sub>R x) v* A = k *\<^sub>R (x v* A)" by (metis matrix_vector_mult_scaleR transpose_matrix_vector) -lemma%important vector_scaleR_matrix_ac: +proposition vector_scaleR_matrix_ac: fixes k :: real and x :: "real^'n" and A :: "real^'m^'n" shows "x v* (k *\<^sub>R A) = k *\<^sub>R (x v* A)" -proof%unimportant - +proof - have "x v* (k *\<^sub>R A) = (k *\<^sub>R x) v* A" unfolding vector_matrix_mult_def by (simp add: algebra_simps)