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)