# HG changeset patch # User wenzelm # Date 1377726081 -7200 # Node ID 220f306f5c4ead7792efcaeda423d1b2762a58d4 # Parent 4766fbe322b56bda21001d4734293995074d2088 tuned proofs; diff -r 4766fbe322b5 -r 220f306f5c4e src/HOL/Multivariate_Analysis/Determinants.thy --- a/src/HOL/Multivariate_Analysis/Determinants.thy Wed Aug 28 22:50:23 2013 +0200 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy Wed Aug 28 23:41:21 2013 +0200 @@ -11,14 +11,18 @@ begin subsection{* First some facts about products*} -lemma setprod_insert_eq: "finite A \ setprod f (insert a A) = (if a \ A then setprod f A else f a * setprod f A)" -apply clarsimp -by(subgoal_tac "insert a A = A", auto) + +lemma setprod_insert_eq: + "finite A \ setprod f (insert a A) = (if a \ A then setprod f A else f a * setprod f A)" + apply clarsimp + apply (subgoal_tac "insert a A = A") + apply auto + done lemma setprod_add_split: assumes mn: "(m::nat) <= n + 1" shows "setprod f {m.. n+p} = setprod f {m .. n} * setprod f {n+1..n+p}" -proof- +proof - let ?A = "{m .. n+p}" let ?B = "{m .. n}" let ?C = "{n+1..n+p}" @@ -30,47 +34,56 @@ lemma setprod_offset: "setprod f {(m::nat) + p .. n + p} = setprod (\i. f (i + p)) {m..n}" -apply (rule setprod_reindex_cong[where f="op + p"]) -apply (auto simp add: image_iff Bex_def inj_on_def) -apply arith -apply (rule ext) -apply (simp add: add_commute) -done + apply (rule setprod_reindex_cong[where f="op + p"]) + apply (auto simp add: image_iff Bex_def inj_on_def) + apply presburger + apply (rule ext) + apply (simp add: add_commute) + done -lemma setprod_singleton: "setprod f {x} = f x" by simp - -lemma setprod_singleton_nat_seg: "setprod f {n..n} = f (n::'a::order)" by simp +lemma setprod_singleton: "setprod f {x} = f x" + by simp -lemma setprod_numseg: "setprod f {m..0} = (if m=0 then f 0 else 1)" - "setprod f {m .. Suc n} = (if m \ Suc n then f (Suc n) * setprod f {m..n} - else setprod f {m..n})" +lemma setprod_singleton_nat_seg: "setprod f {n..n} = f (n::'a::order)" + by simp + +lemma setprod_numseg: + "setprod f {m..0} = (if m = 0 then f 0 else 1)" + "setprod f {m .. Suc n} = + (if m \ Suc n then f (Suc n) * setprod f {m..n} else setprod f {m..n})" by (auto simp add: atLeastAtMostSuc_conv) -lemma setprod_le: assumes fS: "finite S" and fg: "\x\S. f x \ 0 \ f x \ (g x :: 'a::linordered_idom)" +lemma setprod_le: + assumes fS: "finite S" + and fg: "\x\S. f x \ 0 \ f x \ (g x :: 'a::linordered_idom)" shows "setprod f S \ setprod g S" -using fS fg -apply(induct S) -apply simp -apply auto -apply (rule mult_mono) -apply (auto intro: setprod_nonneg) -done + using fS fg + apply (induct S) + apply simp + apply auto + apply (rule mult_mono) + apply (auto intro: setprod_nonneg) + done (* FIXME: In Finite_Set there is a useless further assumption *) -lemma setprod_inversef: "finite A ==> setprod (inverse \ f) A = (inverse (setprod f A) :: 'a:: field_inverse_zero)" +lemma setprod_inversef: + "finite A \ setprod (inverse \ f) A = (inverse (setprod f A) :: 'a:: field_inverse_zero)" apply (erule finite_induct) apply (simp) apply simp done -lemma setprod_le_1: assumes fS: "finite S" and f: "\x\S. f x \ 0 \ f x \ (1::'a::linordered_idom)" +lemma setprod_le_1: + assumes fS: "finite S" + and f: "\x\S. f x \ 0 \ f x \ (1::'a::linordered_idom)" shows "setprod f S \ 1" -using setprod_le[OF fS f] unfolding setprod_1 . + using setprod_le[OF fS f] unfolding setprod_1 . -subsection{* Trace *} + +subsection {* Trace *} -definition trace :: "'a::semiring_1^'n^'n \ 'a" where - "trace A = setsum (\i. ((A$i)$i)) (UNIV::'n set)" +definition trace :: "'a::semiring_1^'n^'n \ 'a" + where "trace A = setsum (\i. ((A$i)$i)) (UNIV::'n set)" lemma trace_0: "trace(mat 0) = 0" by (simp add: trace_def mat_def) @@ -87,14 +100,17 @@ 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 setsum_commute) - by (simp add: mult_commute) + apply (simp add: mult_commute) + done (* ------------------------------------------------------------------------- *) (* Definition of determinant. *) (* ------------------------------------------------------------------------- *) definition det:: "'a::comm_ring_1^'n^'n \ 'a" where - "det A = setsum (\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)) {p. p permutes (UNIV :: 'n set)}" + "det A = + setsum (\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)) + {p. p permutes (UNIV :: 'n set)}" (* ------------------------------------------------------------------------- *) (* A few general lemmas we need below. *) @@ -105,7 +121,8 @@ shows "setprod f S = setprod (f o p) S" using assms by (fact setprod.permute) -lemma setproduct_permute_nat_interval: "p permutes {m::nat .. n} ==> setprod f {m..n} = setprod (f o p) {m..n}" +lemma setproduct_permute_nat_interval: + "p permutes {m::nat .. n} ==> setprod f {m..n} = setprod (f o p) {m..n}" by (blast intro!: setprod_permute) (* ------------------------------------------------------------------------- *) @@ -113,52 +130,71 @@ (* ------------------------------------------------------------------------- *) lemma det_transpose: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)" -proof- +proof - let ?di = "\A i j. A$i$j" let ?U = "(UNIV :: 'n set)" have fU: "finite ?U" by simp - {fix p assume p: "p \ {p. p permutes ?U}" + { + fix p + assume p: "p \ {p. p permutes ?U}" from p have pU: "p permutes ?U" by blast have sth: "sign (inv p) = sign p" by (metis sign_inverse fU p mem_Collect_eq permutation_permutes) from permutes_inj[OF pU] have pi: "inj_on p ?U" by (blast intro: subset_inj_on) from permutes_image[OF pU] - have "setprod (\i. ?di (transpose A) i (inv p i)) ?U = setprod (\i. ?di (transpose A) i (inv p i)) (p ` ?U)" by simp + have "setprod (\i. ?di (transpose A) i (inv p i)) ?U = + setprod (\i. ?di (transpose A) i (inv p i)) (p ` ?U)" by simp also have "\ = setprod ((\i. ?di (transpose A) i (inv p i)) o p) ?U" unfolding setprod_reindex[OF pi] .. also have "\ = setprod (\i. ?di A i (p i)) ?U" - proof- - {fix i assume i: "i \ ?U" + proof - + { + fix i + assume i: "i \ ?U" from i permutes_inv_o[OF pU] permutes_in_image[OF pU] have "((\i. ?di (transpose A) i (inv p i)) o p) i = ?di A i (p i)" - unfolding transpose_def by (simp add: fun_eq_iff)} - then show "setprod ((\i. ?di (transpose A) i (inv p i)) o p) ?U = setprod (\i. ?di A i (p i)) ?U" by (auto intro: setprod_cong) + unfolding transpose_def by (simp add: fun_eq_iff) + } + then show "setprod ((\i. ?di (transpose A) i (inv p i)) o p) ?U = + setprod (\i. ?di A i (p i)) ?U" by (auto intro: setprod_cong) qed - finally have "of_int (sign (inv p)) * (setprod (\i. ?di (transpose A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\i. ?di A i (p i)) ?U)" using sth - by simp} - then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse) - apply (rule setsum_cong2) by blast + finally have "of_int (sign (inv p)) * (setprod (\i. ?di (transpose A) i (inv p i)) ?U) = + of_int (sign p) * (setprod (\i. ?di A i (p i)) ?U)" using sth by simp + } + then show ?thesis + unfolding det_def + apply (subst setsum_permutations_inverse) + apply (rule setsum_cong2) + apply blast + done qed 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 = setprod (\i. A$i$i) (UNIV:: 'n set)" -proof- +proof - let ?U = "UNIV:: 'n set" let ?PU = "{p. p permutes ?U}" let ?pp = "\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)" have fU: "finite ?U" by simp from finite_permutations[OF fU] have fPU: "finite ?PU" . have id0: "{id} \ ?PU" by (auto simp add: permutes_id) - {fix p assume p: "p \ ?PU -{id}" - from p have pU: "p permutes ?U" and pid: "p \ id" by blast+ - from permutes_natset_le[OF pU] pid obtain i where - i: "p i > i" by (metis not_le) - from ld[OF i] have ex:"\i \ ?U. A$i$p i = 0" by blast - from setprod_zero[OF fU ex] have "?pp p = 0" by simp} - then have p0: "\p \ ?PU -{id}. ?pp p = 0" by blast + { + fix p + assume p: "p \ ?PU -{id}" + from p have pU: "p permutes ?U" and pid: "p \ id" + by blast+ + from permutes_natset_le[OF pU] pid obtain i where i: "p i > i" + by (metis not_le) + from ld[OF i] have ex:"\i \ ?U. A$i$p i = 0" + by blast + from setprod_zero[OF fU ex] have "?pp p = 0" + by simp + } + then have p0: "\p \ ?PU -{id}. ?pp p = 0" + by blast from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis unfolding det_def by (simp add: sign_id) qed @@ -167,21 +203,26 @@ 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 = setprod (\i. A$i$i) (UNIV:: 'n set)" -proof- +proof - let ?U = "UNIV:: 'n set" let ?PU = "{p. p permutes ?U}" let ?pp = "(\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set))" have fU: "finite ?U" by simp from finite_permutations[OF fU] have fPU: "finite ?PU" . have id0: "{id} \ ?PU" by (auto simp add: permutes_id) - {fix p assume p: "p \ ?PU -{id}" - from p have pU: "p permutes ?U" and pid: "p \ id" by blast+ - from permutes_natset_ge[OF pU] pid obtain i where - i: "p i < i" by (metis not_le) + { + fix p + assume p: "p \ ?PU -{id}" + from p have pU: "p permutes ?U" and pid: "p \ id" + by blast+ + from permutes_natset_ge[OF pU] pid obtain i where i: "p i < i" + by (metis not_le) from ld[OF i] have ex:"\i \ ?U. A$i$p i = 0" by blast - from setprod_zero[OF fU ex] have "?pp p = 0" by simp} - then have p0: "\p \ ?PU -{id}. ?pp p = 0" by blast - from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis + from setprod_zero[OF fU ex] have "?pp p = 0" by simp + } + then have p0: "\p \ ?PU -{id}. ?pp p = 0" + by blast + from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis unfolding det_def by (simp add: sign_id) qed @@ -189,14 +230,16 @@ fixes A :: "'a::comm_ring_1^'n^'n" assumes ld: "\i j. i \ j \ A$i$j = 0" shows "det A = setprod (\i. A$i$i) (UNIV::'n set)" -proof- +proof - let ?U = "UNIV:: 'n set" let ?PU = "{p. p permutes ?U}" let ?pp = "\p. of_int (sign p) * setprod (\i. A$i$p i) (UNIV :: 'n set)" have fU: "finite ?U" by simp from finite_permutations[OF fU] have fPU: "finite ?PU" . have id0: "{id} \ ?PU" by (auto simp add: permutes_id) - {fix p assume p: "p \ ?PU - {id}" + { + fix p + assume p: "p \ ?PU - {id}" then have "p \ id" by simp then obtain i where i: "p i \ i" unfolding fun_eq_iff by auto from ld [OF i [symmetric]] have ex:"\i \ ?U. A$i$p i = 0" by blast @@ -207,16 +250,22 @@ qed lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1" -proof- +proof - let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n" let ?U = "UNIV :: 'n set" let ?f = "\i j. ?A$i$j" - {fix i assume i: "i \ ?U" - have "?f i i = 1" using i by (vector mat_def)} - hence th: "setprod (\i. ?f i i) ?U = setprod (\x. 1) ?U" + { + fix i + assume i: "i \ ?U" + have "?f i i = 1" using i by (vector mat_def) + } + then have th: "setprod (\i. ?f i i) ?U = setprod (\x. 1) ?U" by (auto intro: setprod_cong) - {fix i j assume i: "i \ ?U" and j: "j \ ?U" and ij: "i \ j" - have "?f i j = 0" using i j ij by (vector mat_def) } + { + fix i j + assume i: "i \ ?U" and j: "j \ ?U" and ij: "i \ j" + have "?f i j = 0" using i j ij by (vector mat_def) + } then have "det ?A = setprod (\i. ?f i i) ?U" using det_diagonal by blast also have "\ = 1" unfolding th setprod_1 .. @@ -232,23 +281,27 @@ shows "det(\ i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A" apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric]) apply (subst sum_permutations_compose_right[OF p]) -proof(rule setsum_cong2) +proof (rule setsum_cong2) let ?U = "UNIV :: 'n set" let ?PU = "{p. p permutes ?U}" - fix q assume qPU: "q \ ?PU" + fix q + assume qPU: "q \ ?PU" have fU: "finite ?U" by simp - from qPU have q: "q permutes ?U" by blast + from qPU have q: "q permutes ?U" + by blast from p q have pp: "permutation p" and qp: "permutation q" by (metis fU permutation_permutes)+ from permutes_inv[OF p] have ip: "inv p permutes ?U" . - have "setprod (\i. A$p i$ (q o p) i) ?U = setprod ((\i. A$p i$(q o p) i) o inv p) ?U" - by (simp only: setprod_permute[OF ip, symmetric]) - also have "\ = setprod (\i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U" - by (simp only: o_def) - also have "\ = setprod (\i. A$i$q i) ?U" by (simp only: o_def permutes_inverses[OF p]) - finally have thp: "setprod (\i. A$p i$ (q o p) i) ?U = setprod (\i. A$i$q i) ?U" - by blast - show "of_int (sign (q o p)) * setprod (\i. A$ p i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\i. A$i$q i) ?U" + have "setprod (\i. A$p i$ (q o p) i) ?U = setprod ((\i. A$p i$(q o p) i) o inv p) ?U" + by (simp only: setprod_permute[OF ip, symmetric]) + also have "\ = setprod (\i. A $ (p o inv p) i $ (q o (p o inv p)) i) ?U" + by (simp only: o_def) + also have "\ = setprod (\i. A$i$q i) ?U" + by (simp only: o_def permutes_inverses[OF p]) + finally have thp: "setprod (\i. A$p i$ (q o p) i) ?U = setprod (\i. A$i$q i) ?U" + by blast + show "of_int (sign (q o p)) * setprod (\i. A$ p i$ (q o p) i) ?U = + of_int (sign p) * of_int (sign q) * setprod (\i. A$i$q i) ?U" by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult) qed @@ -256,7 +309,7 @@ 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- +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))" @@ -270,16 +323,16 @@ lemma det_identical_rows: fixes A :: "'a::linordered_idom^'n^'n" assumes ij: "i \ j" - and r: "row i A = row j A" + and r: "row i A = row j A" shows "det A = 0" proof- - have tha: "\(a::'a) b. a = b ==> b = - a ==> a = 0" + have tha: "\(a::'a) b. a = b \ b = - a \ a = 0" by simp have th1: "of_int (-1) = - 1" by simp let ?p = "Fun.swap i j id" let ?A = "\ i. A $ ?p i" from r have "A = ?A" by (simp add: vec_eq_iff row_def swap_def) - hence "det A = det ?A" by simp + then have "det A = det ?A" by simp moreover have "det A = - det ?A" by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1) ultimately show "det A = 0" by (metis tha) @@ -288,21 +341,22 @@ lemma det_identical_columns: fixes A :: "'a::linordered_idom^'n^'n" assumes ij: "i \ j" - and r: "column i A = column j A" + and r: "column i A = column j A" shows "det A = 0" -apply (subst det_transpose[symmetric]) -apply (rule det_identical_rows[OF ij]) -by (metis row_transpose r) + apply (subst det_transpose[symmetric]) + apply (rule det_identical_rows[OF ij]) + apply (metis row_transpose r) + done lemma det_zero_row: fixes A :: "'a::{idom, ring_char_0}^'n^'n" assumes r: "row i A = 0" shows "det A = 0" -using r -apply (simp add: row_def det_def vec_eq_iff) -apply (rule setsum_0') -apply (auto simp: sign_nz) -done + using r + apply (simp add: row_def det_def vec_eq_iff) + apply (rule setsum_0') + apply (auto simp: sign_nz) + done lemma det_zero_column: fixes A :: "'a::{idom,ring_char_0}^'n^'n" @@ -310,27 +364,32 @@ shows "det A = 0" apply (subst det_transpose[symmetric]) apply (rule det_zero_row [of i]) - by (metis row_transpose r) + apply (metis row_transpose r) + done 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) + - det((\ i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)" -unfolding det_def vec_lambda_beta setsum_addf[symmetric] + det((\ i. if i = k then a i else c i)::'a::comm_ring_1^'n^'n) + + det((\ i. if i = k then b i else c i)::'a::comm_ring_1^'n^'n)" + unfolding det_def vec_lambda_beta setsum_addf[symmetric] proof (rule setsum_cong2) let ?U = "UNIV :: 'n set" let ?pU = "{p. p permutes ?U}" let ?f = "(\i. if i = k then a i + b i else c i)::'n \ 'a::comm_ring_1^'n" let ?g = "(\ i. if i = k then a i else c i)::'n \ 'a::comm_ring_1^'n" let ?h = "(\ i. if i = k then b i else c i)::'n \ 'a::comm_ring_1^'n" - fix p assume p: "p \ ?pU" + fix p + assume p: "p \ ?pU" let ?Uk = "?U - {k}" from p have pU: "p permutes ?U" by blast have kU: "?U = insert k ?Uk" by blast - {fix j assume j: "j \ ?Uk" + { + fix j + assume j: "j \ ?Uk" from j have "?f j $ p j = ?g j $ p j" and "?f j $ p j= ?h j $ p j" - by simp_all} + by simp_all + } then have th1: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?g i $ p i) ?Uk" and th2: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?h i $ p i) ?Uk" apply - @@ -342,36 +401,45 @@ also have "\ = ?f k $ p k * setprod (\i. ?f i $ p i) ?Uk" apply (rule setprod_insert) apply simp - by blast - also have "\ = (a k $ p k * setprod (\i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\i. ?f i $ p i) ?Uk)" by (simp add: field_simps) - also have "\ = (a k $ p k * setprod (\i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\i. ?h i $ p i) ?Uk)" by (metis th1 th2) + apply blast + done + also have "\ = (a k $ p k * setprod (\i. ?f i $ p i) ?Uk) + (b k$ p k * setprod (\i. ?f i $ p i) ?Uk)" + by (simp add: field_simps) + also have "\ = (a k $ p k * setprod (\i. ?g i $ p i) ?Uk) + (b k$ p k * setprod (\i. ?h i $ p i) ?Uk)" + by (metis th1 th2) also have "\ = setprod (\i. ?g i $ p i) (insert k ?Uk) + setprod (\i. ?h i $ p i) (insert k ?Uk)" unfolding setprod_insert[OF th3] by simp - finally have "setprod (\i. ?f i $ p i) ?U = setprod (\i. ?g i $ p i) ?U + setprod (\i. ?h i $ p i) ?U" unfolding kU[symmetric] . - then show "of_int (sign p) * setprod (\i. ?f i $ p i) ?U = of_int (sign p) * setprod (\i. ?g i $ p i) ?U + of_int (sign p) * setprod (\i. ?h i $ p i) ?U" + finally have "setprod (\i. ?f i $ p i) ?U = + setprod (\i. ?g i $ p i) ?U + setprod (\i. ?h i $ p i) ?U" unfolding kU[symmetric] . + then show "of_int (sign p) * setprod (\i. ?f i $ p i) ?U = + of_int (sign p) * setprod (\i. ?g i $ p i) ?U + of_int (sign p) * setprod (\i. ?h i $ p i) ?U" by (simp add: field_simps) qed 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)" - -unfolding det_def vec_lambda_beta setsum_right_distrib + c * det((\ i. if i = k then a i else b i)::'a::comm_ring_1^'n^'n)" + unfolding det_def vec_lambda_beta setsum_right_distrib proof (rule setsum_cong2) let ?U = "UNIV :: 'n set" let ?pU = "{p. p permutes ?U}" let ?f = "(\i. if i = k then c*s a i else b i)::'n \ 'a::comm_ring_1^'n" let ?g = "(\ i. if i = k then a i else b i)::'n \ 'a::comm_ring_1^'n" - fix p assume p: "p \ ?pU" + fix p + assume p: "p \ ?pU" let ?Uk = "?U - {k}" from p have pU: "p permutes ?U" by blast have kU: "?U = insert k ?Uk" by blast - {fix j assume j: "j \ ?Uk" - from j have "?f j $ p j = ?g j $ p j" by simp} + { + fix j + assume j: "j \ ?Uk" + from j have "?f j $ p j = ?g j $ p j" by simp + } then have th1: "setprod (\i. ?f i $ p i) ?Uk = setprod (\i. ?g i $ p i) ?Uk" apply - - apply (rule setprod_cong, simp_all) + apply (rule setprod_cong) + apply simp_all done have th3: "finite ?Uk" "k \ ?Uk" by auto have "setprod (\i. ?f i $ p i) ?U = setprod (\i. ?f i $ p i) (insert k ?Uk)" @@ -379,29 +447,34 @@ also have "\ = ?f k $ p k * setprod (\i. ?f i $ p i) ?Uk" apply (rule setprod_insert) apply simp - by blast - also have "\ = (c*s a k) $ p k * setprod (\i. ?f i $ p i) ?Uk" by (simp add: field_simps) + apply blast + done + also have "\ = (c*s a k) $ p k * setprod (\i. ?f i $ p i) ?Uk" + by (simp add: field_simps) also have "\ = c* (a k $ p k * setprod (\i. ?g i $ p i) ?Uk)" unfolding th1 by (simp add: mult_ac) also have "\ = c* (setprod (\i. ?g i $ p i) (insert k ?Uk))" - unfolding setprod_insert[OF th3] by simp - finally have "setprod (\i. ?f i $ p i) ?U = c* (setprod (\i. ?g i $ p i) ?U)" unfolding kU[symmetric] . - then show "of_int (sign p) * setprod (\i. ?f i $ p i) ?U = c * (of_int (sign p) * setprod (\i. ?g i $ p i) ?U)" + unfolding setprod_insert[OF th3] by simp + finally have "setprod (\i. ?f i $ p i) ?U = c* (setprod (\i. ?g i $ p i) ?U)" + unfolding kU[symmetric] . + then show "of_int (sign p) * setprod (\i. ?f i $ p i) ?U = + c * (of_int (sign p) * setprod (\i. ?g i $ p i) ?U)" by (simp add: field_simps) qed 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] -apply (simp) - unfolding vector_smult_lzero . + using det_row_mul[of k 0 "\i. 1" b] + apply simp + apply (simp only: vector_smult_lzero) + done lemma det_row_operation: fixes A :: "'a::linordered_idom^'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- +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" @@ -415,30 +488,38 @@ fixes A :: "real^'n^'n" assumes x: "x \ span {row j A |j. j \ i}" shows "det (\ k. if k = i then row i A + x else row k A) = det A" -proof- +proof - let ?U = "UNIV :: 'n set" let ?S = "{row j A |j. j \ i}" let ?d = "\x. det (\ k. if k = i then x else row k A)" let ?P = "\x. ?d (row i A + x) = det A" - {fix k - - have "(if k = i then row i A + 0 else row k A) = row k A" by simp} + { + fix k + have "(if k = i then row i A + 0 else row k A) = row k A" by simp + } then have P0: "?P 0" apply - apply (rule cong[of det, OF refl]) - by (vector row_def) + apply (vector row_def) + done moreover - {fix c z y assume zS: "z \ ?S" and Py: "?P y" + { + fix c z y + assume zS: "z \ ?S" and Py: "?P y" from zS obtain j where j: "z = row j A" "i \ j" by blast let ?w = "row i A + y" have th0: "row i A + (c*s z + y) = ?w + c*s z" by vector have thz: "?d z = 0" apply (rule det_identical_rows[OF j(2)]) - using j by (vector row_def) - have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" unfolding th0 .. - then have "?P (c*s z + y)" unfolding thz Py det_row_mul[of i] det_row_add[of i] - by simp } - + using j + apply (vector row_def) + done + have "?d (row i A + (c*s z + y)) = ?d (?w + c*s z)" + unfolding th0 .. + then have "?P (c*s z + y)" + unfolding thz Py det_row_mul[of i] det_row_add[of i] + by simp + } ultimately show ?thesis apply - apply (rule span_induct_alt[of ?P ?S, OF P0, folded scalar_mult_eq_scaleR]) @@ -456,53 +537,68 @@ fixes A:: "real^'n^'n" assumes d: "dependent (rows A)" shows "det A = 0" -proof- +proof - let ?U = "UNIV :: 'n set" from d obtain i where i: "row i A \ span (rows A - {row i A})" unfolding dependent_def rows_def by blast - {fix j k assume jk: "j \ k" - and c: "row j A = row k A" - from det_identical_rows[OF jk c] have ?thesis .} + { + fix j k + assume jk: "j \ k" and c: "row j A = row k A" + from det_identical_rows[OF jk c] have ?thesis . + } moreover - {assume H: "\ i j. i \ j \ row i A \ row j A" + { + assume H: "\ i j. i \ j \ row i A \ row j A" have th0: "- row i A \ span {row j A|j. j \ i}" apply (rule span_neg) apply (rule set_rev_mp) apply (rule i) apply (rule span_mono) - using H i by (auto simp add: rows_def) + using H i + apply (auto simp add: rows_def) + done from det_row_span[OF th0] have "det A = det (\ k. if k = i then 0 *s 1 else row k A)" unfolding right_minus vector_smult_lzero .. with det_row_mul[of i "0::real" "\i. 1"] - have "det A = 0" by simp} + have "det A = 0" by simp + } ultimately show ?thesis by blast qed -lemma det_dependent_columns: assumes d: "dependent(columns (A::real^'n^'n))" shows "det A = 0" -by (metis d det_dependent_rows rows_transpose det_transpose) +lemma det_dependent_columns: + assumes d: "dependent (columns (A::real^'n^'n))" + shows "det A = 0" + by (metis d det_dependent_rows rows_transpose det_transpose) (* ------------------------------------------------------------------------- *) (* Multilinearity and the multiplication formula. *) (* ------------------------------------------------------------------------- *) lemma Cart_lambda_cong: "(\x. f x = g x) \ (vec_lambda f::'a^'n) = (vec_lambda g :: 'a^'n)" - apply (rule iffD1[OF vec_lambda_unique]) by vector + by (rule iffD1[OF vec_lambda_unique]) vector lemma det_linear_row_setsum: assumes fS: "finite S" - shows "det ((\ i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) = setsum (\j. det ((\ i. if i = k then a i j else c i)::'a^'n^'n)) S" -proof(induct rule: finite_induct[OF fS]) - case 1 thus ?case apply simp unfolding setsum_empty det_row_0[of k] .. + shows "det ((\ i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) = + setsum (\j. det ((\ i. if i = k then a i j else c i)::'a^'n^'n)) S" +proof (induct rule: finite_induct[OF fS]) + case 1 + then show ?case + apply simp + unfolding setsum_empty det_row_0[of k] + apply rule + done next case (2 x F) - then show ?case by (simp add: det_row_add cong del: if_weak_cong) + then show ?case + by (simp add: det_row_add cong del: if_weak_cong) qed 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) +proof (induct k) case 0 have th: "{f. \i. f i = i} = {id}" by auto show ?case by (auto simp add: th) @@ -526,13 +622,14 @@ lemma det_linear_rows_setsum_lemma: assumes fS: "finite S" and fT: "finite T" shows "det((\ i. if i \ T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n) = - setsum (\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(induct T arbitrary: a c set: finite) + setsum (\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 (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 - from "empty.prems" show ?case unfolding th0 by simp + have th0: "\x y. (\ i. if i \ {} then x i else y i) = (\ i. y i)" + by vector + from empty.prems show ?case unfolding th0 by simp next case (insert z T a c) let ?F = "\T. {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" @@ -540,42 +637,48 @@ let ?k = "\h. (h(z),(\i. if i = z then i else h i))" let ?s = "\ k a c f. det((\ i. if i \ T then a i (f i) else c i)::'a^'n^'n)" let ?c = "\i. if i = z then a i j else c i" - have thif: "\a b c d. (if a \ b then c else d) = (if a then c else if b then c else d)" by simp + have thif: "\a b c d. (if a \ b then c else d) = (if a then c else if b then c else d)" + by simp have thif2: "\a b c d e. (if a then b else if c then d else e) = - (if c then (if a then b else d) else (if a then b else e))" by simp - from `z \ T` have nz: "\i. i \ T \ i = z \ False" by auto + (if c then (if a then b else d) else (if a then b else e))" + by simp + from `z \ T` have nz: "\i. i \ T \ i = z \ False" + by auto have "det (\ i. if i \ insert z T then setsum (a i) S else c i) = - det (\ i. if i = z then setsum (a i) S - else if i \ T then setsum (a i) S else c i)" + det (\ i. if i = z then setsum (a i) S else if i \ T then setsum (a i) S else c i)" unfolding insert_iff thif .. - also have "\ = (\j\S. det (\ i. if i \ T then setsum (a i) S - else if i = z then a i j else c i))" + also have "\ = (\j\S. det (\ i. if i \ T then setsum (a i) S else if i = z then a i j else c i))" unfolding det_linear_row_setsum[OF fS] apply (subst thif2) - using nz by (simp cong del: if_weak_cong cong add: if_cong) + using nz + apply (simp cong del: if_weak_cong cong add: if_cong) + done finally have tha: "det (\ i. if i \ insert z T then setsum (a i) S else c i) = (\(j, f)\S \ ?F T. det (\ i. if i \ T then a i (f i) else if i = z then a i j else c i))" - unfolding insert.hyps unfolding setsum_cartesian_product by blast + unfolding insert.hyps unfolding setsum_cartesian_product by blast show ?case unfolding tha - apply(rule setsum_eq_general_reverses[where h= "?h" and k= "?k"], + apply (rule setsum_eq_general_reverses[where h= "?h" and k= "?k"], blast intro: finite_cartesian_product fS finite, blast intro: finite_cartesian_product fS finite) using `z \ T` apply auto apply (rule cong[OF refl[of det]]) - by vector + apply vector + done qed lemma det_linear_rows_setsum: assumes fS: "finite (S::'n::finite set)" - shows "det (\ i. setsum (a i) S) = setsum (\f. det (\ i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. \i. f i \ S}" -proof- - have th0: "\x y. ((\ i. if i \ (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\ i. x i)" by vector - - from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] show ?thesis by simp + shows "det (\ i. setsum (a i) S) = + setsum (\f. det (\ i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. \i. f i \ S}" +proof - + have th0: "\x y. ((\ i. if i \ (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\ i. x i)" + by vector + from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] + show ?thesis by simp qed lemma matrix_mul_setsum_alt: @@ -585,75 +688,93 @@ lemma det_rows_mul: "det((\ i. c i *s a i)::'a::comm_ring_1^'n^'n) = - setprod (\i. c i) (UNIV:: 'n set) * det((\ i. a i)::'a^'n^'n)" + setprod (\i. c i) (UNIV:: 'n set) * det((\ i. a i)::'a^'n^'n)" proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2) let ?U = "UNIV :: 'n set" let ?PU = "{p. p permutes ?U}" - fix p assume pU: "p \ ?PU" + fix p + assume pU: "p \ ?PU" let ?s = "of_int (sign p)" - from pU have p: "p permutes ?U" by blast + from pU have p: "p permutes ?U" + by blast have "setprod (\i. c i * a i $ p i) ?U = setprod c ?U * setprod (\i. a i $ p i) ?U" unfolding setprod_timesf .. then show "?s * (\xa\?U. c xa * a xa $ p xa) = - setprod c ?U * (?s* (\xa\?U. a xa $ p xa))" by (simp add: field_simps) + setprod c ?U * (?s* (\xa\?U. a xa $ p xa))" by (simp add: field_simps) qed lemma det_mul: fixes A B :: "'a::linordered_idom^'n^'n" shows "det (A ** B) = det A * det B" -proof- +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}" have fU: "finite ?U" by simp have fF: "finite ?F" by (rule finite) - {fix p assume p: "p permutes ?U" - + { + fix p + assume p: "p permutes ?U" have "p \ ?F" unfolding mem_Collect_eq permutes_in_image[OF p] - using p[unfolded permutes_def] by simp} + using p[unfolded permutes_def] by simp + } then have PUF: "?PU \ ?F" by blast - {fix f assume fPU: "f \ ?F - ?PU" + { + fix f + assume fPU: "f \ ?F - ?PU" have fUU: "f ` ?U \ ?U" using fPU by auto - from fPU have f: "\i \ ?U. f i \ ?U" - "\i. i \ ?U \ f i = i" "\(\y. \!x. f x = y)" unfolding permutes_def - by auto + from fPU have f: "\i \ ?U. f i \ ?U" "\i. i \ ?U \ f i = i" "\(\y. \!x. f x = y)" + unfolding permutes_def by auto let ?A = "(\ i. A$i$f i *s B$f i) :: 'a^'n^'n" let ?B = "(\ i. B$f i) :: 'a^'n^'n" - {assume fni: "\ inj_on f ?U" + { + assume fni: "\ inj_on f ?U" then obtain i j where ij: "f i = f j" "i \ j" unfolding inj_on_def by blast from ij have rth: "row i ?B = row j ?B" by (vector row_def) from det_identical_rows[OF ij(2) rth] have "det (\ i. A$i$f i *s B$f i) = 0" - unfolding det_rows_mul by simp} + unfolding det_rows_mul by simp + } moreover - {assume fi: "inj_on f ?U" + { + assume fi: "inj_on f ?U" from f fi have fith: "\i j. f i = f j \ i = j" unfolding inj_on_def by metis note fs = fi[unfolded surjective_iff_injective_gen[OF fU fU refl fUU, symmetric]] - {fix y + { + fix y from fs f have "\x. f x = y" by blast then obtain x where x: "f x = y" by blast - {fix z assume z: "f z = y" from fith x z have "z = x" by metis} - with x have "\!x. f x = y" by blast} - with f(3) have "det (\ i. A$i$f i *s B$f i) = 0" by blast} - ultimately have "det (\ i. A$i$f i *s B$f i) = 0" by blast} - hence zth: "\ f\ ?F - ?PU. det (\ i. A$i$f i *s B$f i) = 0" by simp - {fix p assume pU: "p \ ?PU" + { + fix z + assume z: "f z = y" + from fith x z have "z = x" by metis + } + with x have "\!x. f x = y" by blast + } + with f(3) have "det (\ i. A$i$f i *s B$f i) = 0" by blast + } + ultimately have "det (\ i. A$i$f i *s B$f i) = 0" by blast + } + hence zth: "\ f\ ?F - ?PU. det (\ i. A$i$f i *s B$f i) = 0" + by simp + { + fix p + assume pU: "p \ ?PU" from pU have p: "p permutes ?U" by blast let ?s = "\p. of_int (sign p)" - let ?f = "\q. ?s p * (\i\ ?U. A $ i $ p i) * - (?s q * (\i\ ?U. B $ i $ q i))" + let ?f = "\q. ?s p * (\i\ ?U. A $ i $ p i) * (?s q * (\i\ ?U. B $ i $ q i))" have "(setsum (\q. ?s q * - (\i\ ?U. (\ i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) = - (setsum (\q. ?s p * (\i\ ?U. A $ i $ p i) * - (?s q * (\i\ ?U. B $ i $ q i))) ?PU)" + (\i\ ?U. (\ i. A $ i $ p i *s B $ p i :: 'a^'n^'n) $ i $ q i)) ?PU) = + (setsum (\q. ?s p * (\i\ ?U. A $ i $ p i) * (?s q * (\i\ ?U. B $ i $ q i))) ?PU)" unfolding sum_permutations_compose_right[OF permutes_inv[OF p], of ?f] proof(rule setsum_cong2) - fix q assume qU: "q \ ?PU" + fix q + assume qU: "q \ ?PU" hence q: "q permutes ?U" by blast from p q have pp: "permutation p" and pq: "permutation q" unfolding permutation_permutes by auto @@ -666,11 +787,15 @@ by (simp add: th00 mult_ac sign_idempotent sign_compose) have th001: "setprod (\i. B$i$ q (inv p i)) ?U = setprod ((\i. B$i$ q (inv p i)) o p) ?U" by (rule setprod_permute[OF p]) - have thp: "setprod (\i. (\ i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = setprod (\i. A$i$p i) ?U * setprod (\i. B$i$ q (inv p i)) ?U" + have thp: "setprod (\i. (\ i. A$i$p i *s B$p i :: 'a^'n^'n) $i $ q i) ?U = + setprod (\i. A$i$p i) ?U * setprod (\i. B$i$ q (inv p i)) ?U" unfolding th001 setprod_timesf[symmetric] o_def permutes_inverses[OF p] apply (rule setprod_cong[OF refl]) - using permutes_in_image[OF q] by vector - show "?s q * setprod (\i. (((\ i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = ?s p * (setprod (\i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\i. B$i$(q o inv p) i) ?U)" + using permutes_in_image[OF q] + apply vector + done + show "?s q * setprod (\i. (((\ i. A$i$p i *s B$p i) :: 'a^'n^'n)$i$q i)) ?U = + ?s p * (setprod (\i. A$i$p i) ?U) * (?s (q o inv p) * setprod (\i. B$i$(q o inv p) i) ?U)" using ths thp pp pq permutation_inverse[OF pp] sign_inverse[OF pp] by (simp add: sign_nz th00 field_simps sign_idempotent sign_compose) qed @@ -703,22 +828,24 @@ lemma invertible_det_nz: fixes A::"real ^'n^'n" shows "invertible A \ det A \ 0" -proof- - {assume "invertible A" +proof - + { + assume "invertible A" then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1" unfolding invertible_righ_inverse by blast hence "det (A ** B) = det (mat 1 :: real ^'n^'n)" by simp - hence "det A \ 0" - apply (simp add: det_mul det_I) by algebra } + hence "det A \ 0" by (simp add: det_mul det_I) algebra + } moreover - {assume H: "\ invertible A" + { + assume H: "\ invertible A" let ?U = "UNIV :: 'n set" have fU: "finite ?U" by simp from H obtain c i where c: "setsum (\i. c i *s row i A) ?U = 0" and iU: "i \ ?U" and ci: "c i \ 0" unfolding invertible_righ_inverse unfolding matrix_right_invertible_independent_rows by blast - have stupid: "\(a::real^'n) b. a + b = 0 \ -a = b" + have *: "\(a::real^'n) b. a + b = 0 \ -a = b" apply (drule_tac f="op + (- a)" in cong[OF refl]) apply (simp only: ab_left_minus add_assoc[symmetric]) apply simp @@ -729,7 +856,7 @@ apply - apply (rule vector_mul_lcancel_imp[OF ci]) apply (auto simp add: field_simps) - unfolding stupid .. + unfolding * .. have thr: "- row i A \ span {row j A| j. j \ i}" unfolding thr0 apply (rule span_setsum) @@ -743,7 +870,8 @@ have thrb: "row i ?B = 0" using iU by (vector row_def) have "det A = 0" unfolding det_row_span[OF thr, symmetric] right_minus - unfolding det_zero_row[OF thrb] ..} + unfolding det_zero_row[OF thrb] .. + } ultimately show ?thesis by blast qed @@ -756,7 +884,7 @@ shows "det ((\ i. if i = k then setsum (\i. x$i *s row i A) (UNIV::'n set) else row i A)::real^'n^'n) = x$k * det A" (is "?lhs = ?rhs") -proof- +proof - let ?U = "UNIV :: 'n set" let ?Uk = "?U - {k}" have U: "?U = insert k ?Uk" by blast @@ -766,7 +894,8 @@ by (vector field_simps) have th001: "\f k . (\x. if x = k then f k else f x) = f" by auto have "(\ i. row i A) = A" by (vector row_def) - then have thd1: "det (\ i. row i A) = det A" by simp + then have thd1: "det (\ i. row i A) = det A" + by simp have thd0: "det (\ i. if i = k then row k A + (\i \ ?Uk. x $ i *s row i A) else row i A) = det A" apply (rule det_row_span) apply (rule span_setsum[OF fUk]) @@ -784,37 +913,44 @@ unfolding thd0 unfolding det_row_mul unfolding th001[of k "\i. row i A"] - unfolding thd1 by (simp add: field_simps) + unfolding thd1 + apply (simp add: field_simps) + done qed lemma cramer_lemma: fixes A :: "real^'n^'n" shows "det((\ i j. if j = k then (A *v x)$i else A$i$j):: real^'n^'n) = x$k * det A" -proof- +proof - let ?U = "UNIV :: 'n set" - have stupid: "\c. setsum (\i. c i *s row i (transpose A)) ?U = setsum (\i. c i *s column i A) ?U" + have *: "\c. setsum (\i. c i *s row i (transpose A)) ?U = setsum (\i. c i *s column i A) ?U" by (auto simp add: row_transpose intro: setsum_cong2) show ?thesis unfolding matrix_mult_vsum - unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric] - unfolding stupid[of "\i. x$i"] - apply (subst det_transpose[symmetric]) - apply (rule cong[OF refl[of det]]) by (vector transpose_def column_def row_def) + unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric] + unfolding *[of "\i. x$i"] + apply (subst det_transpose[symmetric]) + apply (rule cong[OF refl[of det]]) + apply (vector transpose_def column_def row_def) + done qed lemma cramer: fixes A ::"real^'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- +proof - from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1" unfolding invertible_det_nz[symmetric] invertible_def by blast have "(A ** B) *v b = b" by (simp add: B matrix_vector_mul_lid) - hence "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc) + then have "A *v (B *v b) = b" by (simp add: matrix_vector_mul_assoc) then have xe: "\x. A*v x = b" by blast - {fix x assume x: "A *v x = b" - have "x = (\ k. det(\ i j. if j=k then b$i else A$i$j) / det A)" - unfolding x[symmetric] - using d0 by (simp add: vec_eq_iff cramer_lemma field_simps)} + { + fix x + assume x: "A *v x = b" + have "x = (\ k. det(\ i j. if j=k then b$i else A$i$j) / det A)" + unfolding x[symmetric] + using d0 by (simp add: vec_eq_iff cramer_lemma field_simps) + } with xe show ?thesis by auto qed @@ -824,16 +960,19 @@ definition "orthogonal_transformation f \ linear f \ (\v w. f v \ f w = v \ w)" -lemma orthogonal_transformation: "orthogonal_transformation f \ linear f \ (\(v::real ^_). norm (f v) = norm v)" +lemma orthogonal_transformation: + "orthogonal_transformation f \ linear f \ (\(v::real ^_). norm (f v) = norm v)" unfolding orthogonal_transformation_def apply auto apply (erule_tac x=v in allE)+ apply (simp add: norm_eq_sqrt_inner) - by (simp add: dot_norm linear_add[symmetric]) + apply (simp add: dot_norm linear_add[symmetric]) + done -definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \ transpose Q ** Q = mat 1 \ Q ** transpose Q = mat 1" +definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \ + transpose Q ** Q = mat 1 \ Q ** transpose Q = mat 1" -lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \ transpose Q ** Q = mat 1" +lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \ transpose Q ** Q = mat 1" by (metis matrix_left_right_inverse orthogonal_matrix_def) lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)" @@ -842,28 +981,31 @@ lemma orthogonal_matrix_mul: fixes A :: "real ^'n^'n" assumes oA : "orthogonal_matrix A" - and oB: "orthogonal_matrix B" + and oB: "orthogonal_matrix B" shows "orthogonal_matrix(A ** B)" using oA oB unfolding orthogonal_matrix matrix_transpose_mul apply (subst matrix_mul_assoc) apply (subst matrix_mul_assoc[symmetric]) - by (simp add: matrix_mul_rid) + apply (simp add: matrix_mul_rid) + done lemma orthogonal_transformation_matrix: fixes f:: "real^'n \ real^'n" shows "orthogonal_transformation f \ linear f \ orthogonal_matrix(matrix f)" (is "?lhs \ ?rhs") -proof- +proof - let ?mf = "matrix f" let ?ot = "orthogonal_transformation f" let ?U = "UNIV :: 'n set" have fU: "finite ?U" by simp let ?m1 = "mat 1 :: real ^'n^'n" - {assume ot: ?ot + { + assume ot: ?ot from ot have lf: "linear f" and fd: "\v w. f v \ f w = v \ w" unfolding orthogonal_transformation_def orthogonal_matrix by blast+ - {fix i j + { + fix i j let ?A = "transpose ?mf ** ?mf" have th0: "\b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)" "\b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)" @@ -871,16 +1013,22 @@ from fd[rule_format, of "axis i 1" "axis j 1", unfolded matrix_works[OF lf, symmetric] dot_matrix_vector_mul] have "?A$i$j = ?m1 $ i $ j" by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def - th0 setsum_delta[OF fU] mat_def axis_def) } - hence "orthogonal_matrix ?mf" unfolding orthogonal_matrix by vector - with lf have ?rhs by blast} + th0 setsum_delta[OF fU] mat_def axis_def) + } + then have "orthogonal_matrix ?mf" unfolding orthogonal_matrix + by vector + with lf have ?rhs by blast + } moreover - {assume lf: "linear f" and om: "orthogonal_matrix ?mf" + { + assume lf: "linear f" and om: "orthogonal_matrix ?mf" from lf om have ?lhs unfolding orthogonal_matrix_def norm_eq orthogonal_transformation unfolding matrix_works[OF lf, symmetric] apply (subst dot_matrix_vector_mul) - by (simp add: dot_matrix_product matrix_mul_lid)} + apply (simp add: dot_matrix_product matrix_mul_lid) + done + } ultimately show ?thesis by blast qed @@ -888,21 +1036,26 @@ fixes Q:: "'a::linordered_idom^'n^'n" assumes oQ: "orthogonal_matrix Q" shows "det Q = 1 \ det Q = - 1" -proof- - +proof - have th: "\x::'a. x = 1 \ x = - 1 \ x*x = 1" (is "\x::'a. ?ths x") - proof- + proof - fix x:: 'a - have th0: "x*x - 1 = (x - 1)*(x + 1)" by (simp add: field_simps) + have th0: "x*x - 1 = (x - 1)*(x + 1)" + by (simp add: field_simps) have th1: "\(x::'a) y. x = - y \ x + y = 0" - apply (subst eq_iff_diff_eq_0) by simp - have "x*x = 1 \ x*x - 1 = 0" by simp + apply (subst eq_iff_diff_eq_0) + apply simp + done + have "x * x = 1 \ x*x - 1 = 0" by simp also have "\ \ x = 1 \ x = - 1" unfolding th0 th1 by simp finally show "?ths x" .. qed - from oQ have "Q ** transpose Q = mat 1" by (metis orthogonal_matrix_def) - hence "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)" by simp - hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transpose) + from oQ have "Q ** transpose Q = mat 1" + by (metis orthogonal_matrix_def) + then have "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)" + by simp + then have "det Q * det Q = 1" + by (simp add: det_mul det_I det_transpose) then show ?thesis unfolding th . qed @@ -911,25 +1064,29 @@ (* ------------------------------------------------------------------------- *) lemma scaling_linear: fixes f :: "real ^'n \ real ^'n" - assumes f0: "f 0 = 0" and fd: "\x y. dist (f x) (f y) = c * dist x y" + assumes f0: "f 0 = 0" + and fd: "\x y. dist (f x) (f y) = c * dist x y" shows "linear f" -proof- - {fix v w - {fix x note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] } +proof - + { + fix v w + { + fix x + note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] + } note th0 = this have "f v \ f w = c\<^sup>2 * (v \ w)" unfolding dot_norm_neg dist_norm[symmetric] unfolding th0 fd[rule_format] by (simp add: power2_eq_square field_simps)} note fc = this show ?thesis - unfolding linear_def vector_eq[where 'a="real^'n"] scalar_mult_eq_scaleR + unfolding linear_def vector_eq[where 'a="real^'n"] scalar_mult_eq_scaleR by (simp add: inner_add fc field_simps) qed lemma isometry_linear: - "f (0:: real^'n) = (0:: real^'n) \ \x y. dist(f x) (f y) = dist x y - \ linear f" -by (rule scaling_linear[where c=1]) simp_all + "f (0:: real^'n) = (0:: real^'n) \ \x y. dist(f x) (f y) = dist x y \ linear f" + by (rule scaling_linear[where c=1]) simp_all (* ------------------------------------------------------------------------- *) (* Hence another formulation of orthogonal transformation. *) @@ -948,7 +1105,8 @@ apply clarify apply (erule_tac x=v in allE) apply (erule_tac x=0 in allE) - by (simp add: dist_norm) + apply (simp add: dist_norm) + done (* ------------------------------------------------------------------------- *) (* Can extend an isometry from unit sphere. *) @@ -957,15 +1115,19 @@ lemma isometry_sphere_extend: fixes f:: "real ^'n \ real ^'n" assumes f1: "\x. norm x = 1 \ norm (f x) = 1" - and fd1: "\ x y. norm x = 1 \ norm y = 1 \ dist (f x) (f y) = dist x y" + and fd1: "\ x y. norm x = 1 \ norm y = 1 \ dist (f x) (f y) = dist x y" shows "\g. orthogonal_transformation g \ (\x. norm x = 1 \ g x = f x)" -proof- - {fix x y x' y' x0 y0 x0' y0' :: "real ^'n" - assume H: "x = norm x *\<^sub>R x0" "y = norm y *\<^sub>R y0" - "x' = norm x *\<^sub>R x0'" "y' = norm y *\<^sub>R y0'" - "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1" - "norm(x0' - y0') = norm(x0 - y0)" - hence *:"x0 \ y0 = x0' \ y0' + y0' \ x0' - y0 \ x0 " by(simp add: norm_eq norm_eq_1 inner_add inner_diff) +proof - + { + fix x y x' y' x0 y0 x0' y0' :: "real ^'n" + assume H: + "x = norm x *\<^sub>R x0" + "y = norm y *\<^sub>R y0" + "x' = norm x *\<^sub>R x0'" "y' = norm y *\<^sub>R y0'" + "norm x0 = 1" "norm x0' = 1" "norm y0 = 1" "norm y0' = 1" + "norm(x0' - y0') = norm(x0 - y0)" + hence *: "x0 \ y0 = x0' \ y0' + y0' \ x0' - y0 \ x0 " + by (simp add: norm_eq norm_eq_1 inner_add inner_diff) have "norm(x' - y') = norm(x - y)" apply (subst H(1)) apply (subst H(2)) @@ -974,48 +1136,71 @@ using H(5-9) apply (simp add: norm_eq norm_eq_1) apply (simp add: inner_diff scalar_mult_eq_scaleR) unfolding * - by (simp add: field_simps) } + apply (simp add: field_simps) + done + } note th0 = this let ?g = "\x. if x = 0 then 0 else norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)" - {fix x:: "real ^'n" assume nx: "norm x = 1" - have "?g x = f x" using nx by auto} - hence thfg: "\x. norm x = 1 \ ?g x = f x" by blast + { + fix x:: "real ^'n" + assume nx: "norm x = 1" + have "?g x = f x" using nx by auto + } + then have thfg: "\x. norm x = 1 \ ?g x = f x" + by blast have g0: "?g 0 = 0" by simp - {fix x y :: "real ^'n" - {assume "x = 0" "y = 0" - then have "dist (?g x) (?g y) = dist x y" by simp } + { + fix x y :: "real ^'n" + { + assume "x = 0" "y = 0" + then have "dist (?g x) (?g y) = dist x y" by simp + } moreover - {assume "x = 0" "y \ 0" + { + assume "x = 0" "y \ 0" then have "dist (?g x) (?g y) = dist x y" apply (simp add: dist_norm) apply (rule f1[rule_format]) - by(simp add: field_simps)} + apply (simp add: field_simps) + done + } moreover - {assume "x \ 0" "y = 0" + { + assume "x \ 0" "y = 0" then have "dist (?g x) (?g y) = dist x y" apply (simp add: dist_norm) apply (rule f1[rule_format]) - by(simp add: field_simps)} + apply (simp add: field_simps) + done + } moreover - {assume z: "x \ 0" "y \ 0" - have th00: "x = norm x *\<^sub>R (inverse (norm x) *\<^sub>R x)" "y = norm y *\<^sub>R (inverse (norm y) *\<^sub>R y)" "norm x *\<^sub>R f ((inverse (norm x) *\<^sub>R x)) = norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)" + { + assume z: "x \ 0" "y \ 0" + have th00: + "x = norm x *\<^sub>R (inverse (norm x) *\<^sub>R x)" + "y = norm y *\<^sub>R (inverse (norm y) *\<^sub>R y)" + "norm x *\<^sub>R f ((inverse (norm x) *\<^sub>R x)) = norm x *\<^sub>R f (inverse (norm x) *\<^sub>R x)" "norm y *\<^sub>R f (inverse (norm y) *\<^sub>R y) = norm y *\<^sub>R f (inverse (norm y) *\<^sub>R y)" "norm (inverse (norm x) *\<^sub>R x) = 1" "norm (f (inverse (norm x) *\<^sub>R x)) = 1" "norm (inverse (norm y) *\<^sub>R y) = 1" "norm (f (inverse (norm y) *\<^sub>R y)) = 1" "norm (f (inverse (norm x) *\<^sub>R x) - f (inverse (norm y) *\<^sub>R y)) = - norm (inverse (norm x) *\<^sub>R x - inverse (norm y) *\<^sub>R y)" + norm (inverse (norm x) *\<^sub>R x - inverse (norm y) *\<^sub>R y)" using z by (auto simp add: field_simps intro: f1[rule_format] fd1[rule_format, unfolded dist_norm]) from z th0[OF th00] have "dist (?g x) (?g y) = dist x y" - by (simp add: dist_norm)} - ultimately have "dist (?g x) (?g y) = dist x y" by blast} + by (simp add: dist_norm) + } + ultimately have "dist (?g x) (?g y) = dist x y" by blast + } note thd = this show ?thesis apply (rule exI[where x= ?g]) unfolding orthogonal_transformation_isometry - using g0 thfg thd by metis + using g0 thfg thd + apply metis + done qed (* ------------------------------------------------------------------------- *) @@ -1029,14 +1214,17 @@ fixes Q :: "'a::linordered_idom^'n^'n" shows " orthogonal_matrix Q \ rotation_matrix Q \ rotoinversion_matrix Q" by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix) + (* ------------------------------------------------------------------------- *) (* Explicit formulas for low dimensions. *) (* ------------------------------------------------------------------------- *) -lemma setprod_1: "setprod f {(1::nat)..1} = f 1" by simp +lemma setprod_1: "setprod f {(1::nat)..1} = f 1" + by simp lemma setprod_2: "setprod f {(1::nat)..2} = f 1 * f 2" by (simp add: eval_nat_numeral setprod_numseg mult_commute) + lemma setprod_3: "setprod f {(1::nat)..3} = f 1 * f 2 * f 3" by (simp add: eval_nat_numeral setprod_numseg mult_commute) @@ -1044,33 +1232,33 @@ by (simp add: det_def sign_id) lemma det_2: "det (A::'a::comm_ring_1^2^2) = A$1$1 * A$2$2 - A$1$2 * A$2$1" -proof- +proof - have f12: "finite {2::2}" "1 \ {2::2}" by auto show ?thesis - unfolding det_def UNIV_2 - unfolding setsum_over_permutations_insert[OF f12] - unfolding permutes_sing - by (simp add: sign_swap_id sign_id swap_id_eq) + unfolding det_def UNIV_2 + unfolding setsum_over_permutations_insert[OF f12] + unfolding permutes_sing + by (simp add: sign_swap_id sign_id swap_id_eq) qed -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 + - A$1$3 * A$2$1 * A$3$2 - - A$1$1 * A$2$3 * A$3$2 - - A$1$2 * A$2$1 * A$3$3 - - A$1$3 * A$2$2 * A$3$1" -proof- +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 + + A$1$3 * A$2$1 * A$3$2 - + A$1$1 * A$2$3 * A$3$2 - + A$1$2 * A$2$1 * A$3$3 - + A$1$3 * A$2$2 * A$3$1" +proof - have f123: "finite {2::3, 3}" "1 \ {2::3, 3}" by auto have f23: "finite {3::3}" "2 \ {3::3}" by auto show ?thesis - unfolding det_def UNIV_3 - unfolding setsum_over_permutations_insert[OF f123] - unfolding setsum_over_permutations_insert[OF f23] - - unfolding permutes_sing - by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) + unfolding det_def UNIV_3 + unfolding setsum_over_permutations_insert[OF f123] + unfolding setsum_over_permutations_insert[OF f23] + unfolding permutes_sing + by (simp add: sign_swap_id permutation_swap_id sign_compose sign_id swap_id_eq) qed end diff -r 4766fbe322b5 -r 220f306f5c4e src/HOL/Multivariate_Analysis/Operator_Norm.thy --- a/src/HOL/Multivariate_Analysis/Operator_Norm.thy Wed Aug 28 22:50:23 2013 +0200 +++ b/src/HOL/Multivariate_Analysis/Operator_Norm.thy Wed Aug 28 23:41:21 2013 +0200 @@ -11,72 +11,83 @@ definition "onorm f = Sup {norm (f x)| x. norm x = 1}" lemma norm_bound_generalize: - fixes f:: "'a::euclidean_space \ 'b::euclidean_space" + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes lf: "linear f" - shows "(\x. norm x = 1 \ norm (f x) \ b) \ (\x. norm (f x) \ b * norm x)" (is "?lhs \ ?rhs") -proof- - {assume H: ?rhs - {fix x :: "'a" assume x: "norm x = 1" - from H[rule_format, of x] x have "norm (f x) \ b" by simp} - then have ?lhs by blast } + shows "(\x. norm x = 1 \ norm (f x) \ b) \ (\x. norm (f x) \ b * norm x)" + (is "?lhs \ ?rhs") +proof + assume H: ?rhs + { + fix x :: "'a" + assume x: "norm x = 1" + from H[rule_format, of x] x have "norm (f x) \ b" by simp + } + then show ?lhs by blast +next + assume H: ?lhs + have bp: "b \ 0" + apply - + apply (rule order_trans [OF norm_ge_zero]) + apply (rule H[rule_format, of "SOME x::'a. x \ Basis"]) + apply (auto intro: SOME_Basis norm_Basis) + done + { + fix x :: "'a" + { + assume "x = 0" + then have "norm (f x) \ b * norm x" + by (simp add: linear_0[OF lf] bp) + } + moreover + { + assume x0: "x \ 0" + then have n0: "norm x \ 0" by (metis norm_eq_zero) + let ?c = "1/ norm x" + have "norm (?c *\<^sub>R x) = 1" using x0 by (simp add: n0) + with H have "norm (f (?c *\<^sub>R x)) \ b" by blast + then have "?c * norm (f x) \ b" + by (simp add: linear_cmul[OF lf]) + then have "norm (f x) \ b * norm x" + using n0 norm_ge_zero[of x] by (auto simp add: field_simps) + } + ultimately have "norm (f x) \ b * norm x" by blast + } + then show ?rhs by blast +qed - moreover - {assume H: ?lhs - have bp: "b \ 0" - apply - - apply(rule order_trans [OF norm_ge_zero]) - apply(rule H[rule_format, of "SOME x::'a. x \ Basis"]) - by (auto intro: SOME_Basis norm_Basis) - {fix x :: "'a" - {assume "x = 0" - then have "norm (f x) \ b * norm x" by (simp add: linear_0[OF lf] bp)} - moreover - {assume x0: "x \ 0" - hence n0: "norm x \ 0" by (metis norm_eq_zero) - let ?c = "1/ norm x" - have "norm (?c *\<^sub>R x) = 1" using x0 by (simp add: n0) - with H have "norm (f (?c *\<^sub>R x)) \ b" by blast - hence "?c * norm (f x) \ b" - by (simp add: linear_cmul[OF lf]) - hence "norm (f x) \ b * norm x" - using n0 norm_ge_zero[of x] by (auto simp add: field_simps)} - ultimately have "norm (f x) \ b * norm x" by blast} - then have ?rhs by blast} - ultimately show ?thesis by blast -qed - lemma onorm: fixes f:: "'a::euclidean_space \ 'b::euclidean_space" assumes lf: "linear f" - shows "norm (f x) <= onorm f * norm x" - and "\x. norm (f x) <= b * norm x \ onorm f <= b" -proof- - { - let ?S = "{norm (f x) |x. norm x = 1}" - have "norm (f (SOME i. i \ Basis)) \ ?S" - by (auto intro!: exI[of _ "SOME i. i \ Basis"] norm_Basis SOME_Basis) - hence Se: "?S \ {}" by auto - from linear_bounded[OF lf] have b: "\ b. ?S *<= b" - unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def) - { from isLub_cSup[OF Se b, unfolded onorm_def[symmetric]] - show "norm (f x) <= onorm f * norm x" - apply - - apply (rule spec[where x = x]) - unfolding norm_bound_generalize[OF lf, symmetric] - by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} - { - show "\x. norm (f x) <= b * norm x \ onorm f <= b" - using isLub_cSup[OF Se b, unfolded onorm_def[symmetric]] - unfolding norm_bound_generalize[OF lf, symmetric] - by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def)} - } + shows "norm (f x) \ onorm f * norm x" + and "\x. norm (f x) \ b * norm x \ onorm f \ b" +proof - + let ?S = "{norm (f x) |x. norm x = 1}" + have "norm (f (SOME i. i \ Basis)) \ ?S" + by (auto intro!: exI[of _ "SOME i. i \ Basis"] norm_Basis SOME_Basis) + then have Se: "?S \ {}" by auto + from linear_bounded[OF lf] have b: "\ b. ?S *<= b" + unfolding norm_bound_generalize[OF lf, symmetric] by (auto simp add: setle_def) + from isLub_cSup[OF Se b, unfolded onorm_def[symmetric]] + show "norm (f x) <= onorm f * norm x" + apply - + apply (rule spec[where x = x]) + unfolding norm_bound_generalize[OF lf, symmetric] + apply (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def) + done + show "\x. norm (f x) <= b * norm x \ onorm f <= b" + using isLub_cSup[OF Se b, unfolded onorm_def[symmetric]] + unfolding norm_bound_generalize[OF lf, symmetric] + by (auto simp add: isLub_def isUb_def leastP_def setge_def setle_def) qed -lemma onorm_pos_le: assumes lf: "linear (f::'n::euclidean_space \ 'm::euclidean_space)" shows "0 <= onorm f" - using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "SOME i. i \ Basis"]] +lemma onorm_pos_le: + assumes lf: "linear (f::'n::euclidean_space \ 'm::euclidean_space)" + shows "0 \ onorm f" + using order_trans[OF norm_ge_zero onorm(1)[OF lf, of "SOME i. i \ Basis"]] by (simp add: SOME_Basis) -lemma onorm_eq_0: assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" +lemma onorm_eq_0: + assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" shows "onorm f = 0 \ (\x. f x = 0)" using onorm[OF lf] apply (auto simp add: onorm_pos_le) @@ -87,47 +98,53 @@ done lemma onorm_const: "onorm(\x::'a::euclidean_space. (y::'b::euclidean_space)) = norm y" -proof- +proof - let ?f = "\x::'a. (y::'b)" have th: "{norm (?f x)| x. norm x = 1} = {norm y}" by (auto simp: SOME_Basis intro!: exI[of _ "SOME i. i \ Basis"]) show ?thesis unfolding onorm_def th - apply (rule cSup_unique) by (simp_all add: setle_def) + apply (rule cSup_unique) + apply (simp_all add: setle_def) + done qed -lemma onorm_pos_lt: assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" +lemma onorm_pos_lt: + assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" shows "0 < onorm f \ ~(\x. f x = 0)" unfolding onorm_eq_0[OF lf, symmetric] using onorm_pos_le[OF lf] by arith lemma onorm_compose: assumes lf: "linear (f::'n::euclidean_space \ 'm::euclidean_space)" - and lg: "linear (g::'k::euclidean_space \ 'n::euclidean_space)" - shows "onorm (f o g) <= onorm f * onorm g" - apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format]) - unfolding o_def - apply (subst mult_assoc) - apply (rule order_trans) - apply (rule onorm(1)[OF lf]) - apply (rule mult_left_mono) - apply (rule onorm(1)[OF lg]) - apply (rule onorm_pos_le[OF lf]) - done + and lg: "linear (g::'k::euclidean_space \ 'n::euclidean_space)" + shows "onorm (f o g) \ onorm f * onorm g" + apply (rule onorm(2)[OF linear_compose[OF lg lf], rule_format]) + unfolding o_def + apply (subst mult_assoc) + apply (rule order_trans) + apply (rule onorm(1)[OF lf]) + apply (rule mult_left_mono) + apply (rule onorm(1)[OF lg]) + apply (rule onorm_pos_le[OF lf]) + done -lemma onorm_neg_lemma: assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" +lemma onorm_neg_lemma: + assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" shows "onorm (\x. - f x) \ onorm f" using onorm[OF linear_compose_neg[OF lf]] onorm[OF lf] unfolding norm_minus_cancel by metis -lemma onorm_neg: assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" +lemma onorm_neg: + assumes lf: "linear (f::'a::euclidean_space \ 'b::euclidean_space)" shows "onorm (\x. - f x) = onorm f" using onorm_neg_lemma[OF lf] onorm_neg_lemma[OF linear_compose_neg[OF lf]] by simp lemma onorm_triangle: - assumes lf: "linear (f::'n::euclidean_space \ 'm::euclidean_space)" and lg: "linear g" - shows "onorm (\x. f x + g x) <= onorm f + onorm g" + assumes lf: "linear (f::'n::euclidean_space \ 'm::euclidean_space)" + and lg: "linear g" + shows "onorm (\x. f x + g x) \ onorm f + onorm g" apply(rule onorm(2)[OF linear_compose_add[OF lf lg], rule_format]) apply (rule order_trans) apply (rule norm_triangle_ineq) @@ -137,17 +154,20 @@ apply (rule onorm(1)[OF lg]) done -lemma onorm_triangle_le: "linear (f::'n::euclidean_space \ 'm::euclidean_space) \ linear g \ onorm(f) + onorm(g) <= e - \ onorm(\x. f x + g x) <= e" +lemma onorm_triangle_le: + "linear (f::'n::euclidean_space \ 'm::euclidean_space) \ + linear g \ onorm f + onorm g \ e \ onorm (\x. f x + g x) \ e" apply (rule order_trans) apply (rule onorm_triangle) apply assumption+ done -lemma onorm_triangle_lt: "linear (f::'n::euclidean_space \ 'm::euclidean_space) \ linear g \ onorm(f) + onorm(g) < e - ==> onorm(\x. f x + g x) < e" +lemma onorm_triangle_lt: + "linear (f::'n::euclidean_space \ 'm::euclidean_space) \ linear g \ + onorm f + onorm g < e \ onorm(\x. f x + g x) < e" apply (rule order_le_less_trans) apply (rule onorm_triangle) - by assumption+ + apply assumption+ + done end