src/HOL/Multivariate_Analysis/Determinants.thy
changeset 34291 4e896680897e
parent 34289 c9c14c72d035
child 35028 108662d50512
equal deleted inserted replaced
34290:1edf0f223c6e 34291:4e896680897e
    65   shows "setprod f S \<le> 1"
    65   shows "setprod f S \<le> 1"
    66 using setprod_le[OF fS f] unfolding setprod_1 .
    66 using setprod_le[OF fS f] unfolding setprod_1 .
    67 
    67 
    68 subsection{* Trace *}
    68 subsection{* Trace *}
    69 
    69 
    70 definition trace :: "'a::semiring_1^'n^'n::finite \<Rightarrow> 'a" where
    70 definition trace :: "'a::semiring_1^'n^'n \<Rightarrow> 'a" where
    71   "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
    71   "trace A = setsum (\<lambda>i. ((A$i)$i)) (UNIV::'n set)"
    72 
    72 
    73 lemma trace_0: "trace(mat 0) = 0"
    73 lemma trace_0: "trace(mat 0) = 0"
    74   by (simp add: trace_def mat_def)
    74   by (simp add: trace_def mat_def)
    75 
    75 
    76 lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n::finite) = of_nat(CARD('n))"
    76 lemma trace_I: "trace(mat 1 :: 'a::semiring_1^'n^'n) = of_nat(CARD('n))"
    77   by (simp add: trace_def mat_def)
    77   by (simp add: trace_def mat_def)
    78 
    78 
    79 lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n::finite) + B) = trace A + trace B"
    79 lemma trace_add: "trace ((A::'a::comm_semiring_1^'n^'n) + B) = trace A + trace B"
    80   by (simp add: trace_def setsum_addf)
    80   by (simp add: trace_def setsum_addf)
    81 
    81 
    82 lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n::finite) - B) = trace A - trace B"
    82 lemma trace_sub: "trace ((A::'a::comm_ring_1^'n^'n) - B) = trace A - trace B"
    83   by (simp add: trace_def setsum_subtractf)
    83   by (simp add: trace_def setsum_subtractf)
    84 
    84 
    85 lemma trace_mul_sym:"trace ((A::'a::comm_semiring_1^'n^'n::finite) ** B) = trace (B**A)"
    85 lemma trace_mul_sym:"trace ((A::'a::comm_semiring_1^'n^'n) ** B) = trace (B**A)"
    86   apply (simp add: trace_def matrix_matrix_mult_def)
    86   apply (simp add: trace_def matrix_matrix_mult_def)
    87   apply (subst setsum_commute)
    87   apply (subst setsum_commute)
    88   by (simp add: mult_commute)
    88   by (simp add: mult_commute)
    89 
    89 
    90 (* ------------------------------------------------------------------------- *)
    90 (* ------------------------------------------------------------------------- *)
    91 (* Definition of determinant.                                                *)
    91 (* Definition of determinant.                                                *)
    92 (* ------------------------------------------------------------------------- *)
    92 (* ------------------------------------------------------------------------- *)
    93 
    93 
    94 definition det:: "'a::comm_ring_1^'n^'n::finite \<Rightarrow> 'a" where
    94 definition det:: "'a::comm_ring_1^'n^'n \<Rightarrow> 'a" where
    95   "det A = setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)) {p. p permutes (UNIV :: 'n set)}"
    95   "det A = setsum (\<lambda>p. of_int (sign p) * setprod (\<lambda>i. A$i$p i) (UNIV :: 'n set)) {p. p permutes (UNIV :: 'n set)}"
    96 
    96 
    97 (* ------------------------------------------------------------------------- *)
    97 (* ------------------------------------------------------------------------- *)
    98 (* A few general lemmas we need below.                                       *)
    98 (* A few general lemmas we need below.                                       *)
    99 (* ------------------------------------------------------------------------- *)
    99 (* ------------------------------------------------------------------------- *)
   119 
   119 
   120 (* ------------------------------------------------------------------------- *)
   120 (* ------------------------------------------------------------------------- *)
   121 (* Basic determinant properties.                                             *)
   121 (* Basic determinant properties.                                             *)
   122 (* ------------------------------------------------------------------------- *)
   122 (* ------------------------------------------------------------------------- *)
   123 
   123 
   124 lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n::finite)"
   124 lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n)"
   125 proof-
   125 proof-
   126   let ?di = "\<lambda>A i j. A$i$j"
   126   let ?di = "\<lambda>A i j. A$i$j"
   127   let ?U = "(UNIV :: 'n set)"
   127   let ?U = "(UNIV :: 'n set)"
   128   have fU: "finite ?U" by simp
   128   have fU: "finite ?U" by simp
   129   {fix p assume p: "p \<in> {p. p permutes ?U}"
   129   {fix p assume p: "p \<in> {p. p permutes ?U}"
   149   then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse)
   149   then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse)
   150   apply (rule setsum_cong2) by blast
   150   apply (rule setsum_cong2) by blast
   151 qed
   151 qed
   152 
   152 
   153 lemma det_lowerdiagonal:
   153 lemma det_lowerdiagonal:
   154   fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}"
   154   fixes A :: "'a::comm_ring_1^('n::{finite,wellorder})^('n::{finite,wellorder})"
   155   assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0"
   155   assumes ld: "\<And>i j. i < j \<Longrightarrow> A$i$j = 0"
   156   shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
   156   shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
   157 proof-
   157 proof-
   158   let ?U = "UNIV:: 'n set"
   158   let ?U = "UNIV:: 'n set"
   159   let ?PU = "{p. p permutes ?U}"
   159   let ?PU = "{p. p permutes ?U}"
   171   from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
   171   from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
   172     unfolding det_def by (simp add: sign_id)
   172     unfolding det_def by (simp add: sign_id)
   173 qed
   173 qed
   174 
   174 
   175 lemma det_upperdiagonal:
   175 lemma det_upperdiagonal:
   176   fixes A :: "'a::comm_ring_1^'n^'n::{finite,wellorder}"
   176   fixes A :: "'a::comm_ring_1^'n::{finite,wellorder}^'n::{finite,wellorder}"
   177   assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0"
   177   assumes ld: "\<And>i j. i > j \<Longrightarrow> A$i$j = 0"
   178   shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
   178   shows "det A = setprod (\<lambda>i. A$i$i) (UNIV:: 'n set)"
   179 proof-
   179 proof-
   180   let ?U = "UNIV:: 'n set"
   180   let ?U = "UNIV:: 'n set"
   181   let ?PU = "{p. p permutes ?U}"
   181   let ?PU = "{p. p permutes ?U}"
   193   from   setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
   193   from   setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
   194     unfolding det_def by (simp add: sign_id)
   194     unfolding det_def by (simp add: sign_id)
   195 qed
   195 qed
   196 
   196 
   197 lemma det_diagonal:
   197 lemma det_diagonal:
   198   fixes A :: "'a::comm_ring_1^'n^'n::finite"
   198   fixes A :: "'a::comm_ring_1^'n^'n"
   199   assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0"
   199   assumes ld: "\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0"
   200   shows "det A = setprod (\<lambda>i. A$i$i) (UNIV::'n set)"
   200   shows "det A = setprod (\<lambda>i. A$i$i) (UNIV::'n set)"
   201 proof-
   201 proof-
   202   let ?U = "UNIV:: 'n set"
   202   let ?U = "UNIV:: 'n set"
   203   let ?PU = "{p. p permutes ?U}"
   203   let ?PU = "{p. p permutes ?U}"
   213   then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"  by blast
   213   then have p0: "\<forall>p \<in> ?PU - {id}. ?pp p = 0"  by blast
   214   from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
   214   from setsum_mono_zero_cong_left[OF fPU id0 p0] show ?thesis
   215     unfolding det_def by (simp add: sign_id)
   215     unfolding det_def by (simp add: sign_id)
   216 qed
   216 qed
   217 
   217 
   218 lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n::finite) = 1"
   218 lemma det_I: "det (mat 1 :: 'a::comm_ring_1^'n^'n) = 1"
   219 proof-
   219 proof-
   220   let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n"
   220   let ?A = "mat 1 :: 'a::comm_ring_1^'n^'n"
   221   let ?U = "UNIV :: 'n set"
   221   let ?U = "UNIV :: 'n set"
   222   let ?f = "\<lambda>i j. ?A$i$j"
   222   let ?f = "\<lambda>i j. ?A$i$j"
   223   {fix i assume i: "i \<in> ?U"
   223   {fix i assume i: "i \<in> ?U"
   230     by blast
   230     by blast
   231   also have "\<dots> = 1" unfolding th setprod_1 ..
   231   also have "\<dots> = 1" unfolding th setprod_1 ..
   232   finally show ?thesis .
   232   finally show ?thesis .
   233 qed
   233 qed
   234 
   234 
   235 lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n::finite) = 0"
   235 lemma det_0: "det (mat 0 :: 'a::comm_ring_1^'n^'n) = 0"
   236   by (simp add: det_def setprod_zero)
   236   by (simp add: det_def setprod_zero)
   237 
   237 
   238 lemma det_permute_rows:
   238 lemma det_permute_rows:
   239   fixes A :: "'a::comm_ring_1^'n^'n::finite"
   239   fixes A :: "'a::comm_ring_1^'n^'n"
   240   assumes p: "p permutes (UNIV :: 'n::finite set)"
   240   assumes p: "p permutes (UNIV :: 'n::finite set)"
   241   shows "det(\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
   241   shows "det(\<chi> i. A$p i :: 'a^'n^'n) = of_int (sign p) * det A"
   242   apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric])
   242   apply (simp add: det_def setsum_right_distrib mult_assoc[symmetric])
   243   apply (subst sum_permutations_compose_right[OF p])
   243   apply (subst sum_permutations_compose_right[OF p])
   244 proof(rule setsum_cong2)
   244 proof(rule setsum_cong2)
   260   show "of_int (sign (q o p)) * setprod (\<lambda>i. A$ p i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U"
   260   show "of_int (sign (q o p)) * setprod (\<lambda>i. A$ p i$ (q o p) i) ?U = of_int (sign p) * of_int (sign q) * setprod (\<lambda>i. A$i$q i) ?U"
   261     by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult)
   261     by (simp only: thp sign_compose[OF qp pp] mult_commute of_int_mult)
   262 qed
   262 qed
   263 
   263 
   264 lemma det_permute_columns:
   264 lemma det_permute_columns:
   265   fixes A :: "'a::comm_ring_1^'n^'n::finite"
   265   fixes A :: "'a::comm_ring_1^'n^'n"
   266   assumes p: "p permutes (UNIV :: 'n set)"
   266   assumes p: "p permutes (UNIV :: 'n set)"
   267   shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A"
   267   shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A"
   268 proof-
   268 proof-
   269   let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n"
   269   let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n"
   270   let ?At = "transp A"
   270   let ?At = "transp A"
   275     by (simp add: transp_def Cart_eq)
   275     by (simp add: transp_def Cart_eq)
   276   ultimately show ?thesis by simp
   276   ultimately show ?thesis by simp
   277 qed
   277 qed
   278 
   278 
   279 lemma det_identical_rows:
   279 lemma det_identical_rows:
   280   fixes A :: "'a::ordered_idom^'n^'n::finite"
   280   fixes A :: "'a::ordered_idom^'n^'n"
   281   assumes ij: "i \<noteq> j"
   281   assumes ij: "i \<noteq> j"
   282   and r: "row i A = row j A"
   282   and r: "row i A = row j A"
   283   shows "det A = 0"
   283   shows "det A = 0"
   284 proof-
   284 proof-
   285   have tha: "\<And>(a::'a) b. a = b ==> b = - a ==> a = 0"
   285   have tha: "\<And>(a::'a) b. a = b ==> b = - a ==> a = 0"
   293     by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1)
   293     by (simp add: det_permute_rows[OF permutes_swap_id] sign_swap_id ij th1)
   294   ultimately show "det A = 0" by (metis tha)
   294   ultimately show "det A = 0" by (metis tha)
   295 qed
   295 qed
   296 
   296 
   297 lemma det_identical_columns:
   297 lemma det_identical_columns:
   298   fixes A :: "'a::ordered_idom^'n^'n::finite"
   298   fixes A :: "'a::ordered_idom^'n^'n"
   299   assumes ij: "i \<noteq> j"
   299   assumes ij: "i \<noteq> j"
   300   and r: "column i A = column j A"
   300   and r: "column i A = column j A"
   301   shows "det A = 0"
   301   shows "det A = 0"
   302 apply (subst det_transp[symmetric])
   302 apply (subst det_transp[symmetric])
   303 apply (rule det_identical_rows[OF ij])
   303 apply (rule det_identical_rows[OF ij])
   304 by (metis row_transp r)
   304 by (metis row_transp r)
   305 
   305 
   306 lemma det_zero_row:
   306 lemma det_zero_row:
   307   fixes A :: "'a::{idom, ring_char_0}^'n^'n::finite"
   307   fixes A :: "'a::{idom, ring_char_0}^'n^'n"
   308   assumes r: "row i A = 0"
   308   assumes r: "row i A = 0"
   309   shows "det A = 0"
   309   shows "det A = 0"
   310 using r
   310 using r
   311 apply (simp add: row_def det_def Cart_eq)
   311 apply (simp add: row_def det_def Cart_eq)
   312 apply (rule setsum_0')
   312 apply (rule setsum_0')
   313 apply (auto simp: sign_nz)
   313 apply (auto simp: sign_nz)
   314 done
   314 done
   315 
   315 
   316 lemma det_zero_column:
   316 lemma det_zero_column:
   317   fixes A :: "'a::{idom,ring_char_0}^'n^'n::finite"
   317   fixes A :: "'a::{idom,ring_char_0}^'n^'n"
   318   assumes r: "column i A = 0"
   318   assumes r: "column i A = 0"
   319   shows "det A = 0"
   319   shows "det A = 0"
   320   apply (subst det_transp[symmetric])
   320   apply (subst det_transp[symmetric])
   321   apply (rule det_zero_row [of i])
   321   apply (rule det_zero_row [of i])
   322   by (metis row_transp r)
   322   by (metis row_transp r)
   405 using det_row_mul[of k 0 "\<lambda>i. 1" b]
   405 using det_row_mul[of k 0 "\<lambda>i. 1" b]
   406 apply (simp)
   406 apply (simp)
   407   unfolding vector_smult_lzero .
   407   unfolding vector_smult_lzero .
   408 
   408 
   409 lemma det_row_operation:
   409 lemma det_row_operation:
   410   fixes A :: "'a::ordered_idom^'n^'n::finite"
   410   fixes A :: "'a::ordered_idom^'n^'n"
   411   assumes ij: "i \<noteq> j"
   411   assumes ij: "i \<noteq> j"
   412   shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A"
   412   shows "det (\<chi> k. if k = i then row i A + c *s row j A else row k A) = det A"
   413 proof-
   413 proof-
   414   let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n"
   414   let ?Z = "(\<chi> k. if k = i then row j A else row k A) :: 'a ^'n^'n"
   415   have th: "row i ?Z = row j ?Z" by (vector row_def)
   415   have th: "row i ?Z = row j ?Z" by (vector row_def)
   419     unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2
   419     unfolding det_row_add [of i] det_row_mul[of i] det_identical_rows[OF ij th] th2
   420     by simp
   420     by simp
   421 qed
   421 qed
   422 
   422 
   423 lemma det_row_span:
   423 lemma det_row_span:
   424   fixes A :: "'a:: ordered_idom^'n^'n::finite"
   424   fixes A :: "'a:: ordered_idom^'n^'n"
   425   assumes x: "x \<in> span {row j A |j. j \<noteq> i}"
   425   assumes x: "x \<in> span {row j A |j. j \<noteq> i}"
   426   shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A"
   426   shows "det (\<chi> k. if k = i then row i A + x else row k A) = det A"
   427 proof-
   427 proof-
   428   let ?U = "UNIV :: 'n set"
   428   let ?U = "UNIV :: 'n set"
   429   let ?S = "{row j A |j. j \<noteq> i}"
   429   let ?S = "{row j A |j. j \<noteq> i}"
   460 (* May as well do this, though it's a bit unsatisfactory since it ignores    *)
   460 (* May as well do this, though it's a bit unsatisfactory since it ignores    *)
   461 (* exact duplicates by considering the rows/columns as a set.                *)
   461 (* exact duplicates by considering the rows/columns as a set.                *)
   462 (* ------------------------------------------------------------------------- *)
   462 (* ------------------------------------------------------------------------- *)
   463 
   463 
   464 lemma det_dependent_rows:
   464 lemma det_dependent_rows:
   465   fixes A:: "'a::ordered_idom^'n^'n::finite"
   465   fixes A:: "'a::ordered_idom^'n^'n"
   466   assumes d: "dependent (rows A)"
   466   assumes d: "dependent (rows A)"
   467   shows "det A = 0"
   467   shows "det A = 0"
   468 proof-
   468 proof-
   469   let ?U = "UNIV :: 'n set"
   469   let ?U = "UNIV :: 'n set"
   470   from d obtain i where i: "row i A \<in> span (rows A - {row i A})"
   470   from d obtain i where i: "row i A \<in> span (rows A - {row i A})"
   486     with det_row_mul[of i "0::'a" "\<lambda>i. 1"]
   486     with det_row_mul[of i "0::'a" "\<lambda>i. 1"]
   487     have "det A = 0" by simp}
   487     have "det A = 0" by simp}
   488   ultimately show ?thesis by blast
   488   ultimately show ?thesis by blast
   489 qed
   489 qed
   490 
   490 
   491 lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n::finite))" shows "det A = 0"
   491 lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::ordered_idom^'n^'n))" shows "det A = 0"
   492 by (metis d det_dependent_rows rows_transp det_transp)
   492 by (metis d det_dependent_rows rows_transp det_transp)
   493 
   493 
   494 (* ------------------------------------------------------------------------- *)
   494 (* ------------------------------------------------------------------------- *)
   495 (* Multilinearity and the multiplication formula.                            *)
   495 (* Multilinearity and the multiplication formula.                            *)
   496 (* ------------------------------------------------------------------------- *)
   496 (* ------------------------------------------------------------------------- *)
   497 
   497 
   498 lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n::finite)"
   498 lemma Cart_lambda_cong: "(\<And>x. f x = g x) \<Longrightarrow> (Cart_lambda f::'a^'n) = (Cart_lambda g :: 'a^'n)"
   499   apply (rule iffD1[OF Cart_lambda_unique]) by vector
   499   apply (rule iffD1[OF Cart_lambda_unique]) by vector
   500 
   500 
   501 lemma det_linear_row_setsum:
   501 lemma det_linear_row_setsum:
   502   assumes fS: "finite S"
   502   assumes fS: "finite S"
   503   shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n::finite) = setsum (\<lambda>j. det ((\<chi> i. if i = k then a  i j else c i)::'a^'n^'n)) S"
   503   shows "det ((\<chi> i. if i = k then setsum (a i) S else c i)::'a::comm_ring_1^'n^'n) = setsum (\<lambda>j. det ((\<chi> i. if i = k then a  i j else c i)::'a^'n^'n)) S"
   504 proof(induct rule: finite_induct[OF fS])
   504 proof(induct rule: finite_induct[OF fS])
   505   case 1 thus ?case apply simp  unfolding setsum_empty det_row_0[of k] ..
   505   case 1 thus ?case apply simp  unfolding setsum_empty det_row_0[of k] ..
   506 next
   506 next
   507   case (2 x F)
   507   case (2 x F)
   508   then  show ?case by (simp add: det_row_add cong del: if_weak_cong)
   508   then  show ?case by (simp add: det_row_add cong del: if_weak_cong)
   532 
   532 
   533 lemma eq_id_iff[simp]: "(\<forall>x. f x = x) = (f = id)" by (auto intro: ext)
   533 lemma eq_id_iff[simp]: "(\<forall>x. f x = x) = (f = id)" by (auto intro: ext)
   534 
   534 
   535 lemma det_linear_rows_setsum_lemma:
   535 lemma det_linear_rows_setsum_lemma:
   536   assumes fS: "finite S" and fT: "finite T"
   536   assumes fS: "finite S" and fT: "finite T"
   537   shows "det((\<chi> i. if i \<in> T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n::finite) =
   537   shows "det((\<chi> i. if i \<in> T then setsum (a i) S else c i):: 'a::comm_ring_1^'n^'n) =
   538              setsum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n))
   538              setsum (\<lambda>f. det((\<chi> i. if i \<in> T then a i (f i) else c i)::'a^'n^'n))
   539                  {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
   539                  {f. (\<forall>i \<in> T. f i \<in> S) \<and> (\<forall>i. i \<notin> T \<longrightarrow> f i = i)}"
   540 using fT
   540 using fT
   541 proof(induct T arbitrary: a c set: finite)
   541 proof(induct T arbitrary: a c set: finite)
   542   case empty
   542   case empty
   578     by vector
   578     by vector
   579 qed
   579 qed
   580 
   580 
   581 lemma det_linear_rows_setsum:
   581 lemma det_linear_rows_setsum:
   582   assumes fS: "finite (S::'n::finite set)"
   582   assumes fS: "finite (S::'n::finite set)"
   583   shows "det (\<chi> i. setsum (a i) S) = setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n::finite)) {f. \<forall>i. f i \<in> S}"
   583   shows "det (\<chi> i. setsum (a i) S) = setsum (\<lambda>f. det (\<chi> i. a i (f i) :: 'a::comm_ring_1 ^ 'n^'n)) {f. \<forall>i. f i \<in> S}"
   584 proof-
   584 proof-
   585   have th0: "\<And>x y. ((\<chi> i. if i \<in> (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" by vector
   585   have th0: "\<And>x y. ((\<chi> i. if i \<in> (UNIV:: 'n set) then x i else y i) :: 'a^'n^'n) = (\<chi> i. x i)" by vector
   586 
   586 
   587   from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] show ?thesis by simp
   587   from det_linear_rows_setsum_lemma[OF fS, of "UNIV :: 'n set" a, unfolded th0, OF finite] show ?thesis by simp
   588 qed
   588 qed
   589 
   589 
   590 lemma matrix_mul_setsum_alt:
   590 lemma matrix_mul_setsum_alt:
   591   fixes A B :: "'a::comm_ring_1^'n^'n::finite"
   591   fixes A B :: "'a::comm_ring_1^'n^'n"
   592   shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))"
   592   shows "A ** B = (\<chi> i. setsum (\<lambda>k. A$i$k *s B $ k) (UNIV :: 'n set))"
   593   by (vector matrix_matrix_mult_def setsum_component)
   593   by (vector matrix_matrix_mult_def setsum_component)
   594 
   594 
   595 lemma det_rows_mul:
   595 lemma det_rows_mul:
   596   "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n::finite) =
   596   "det((\<chi> i. c i *s a i)::'a::comm_ring_1^'n^'n) =
   597   setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
   597   setprod (\<lambda>i. c i) (UNIV:: 'n set) * det((\<chi> i. a i)::'a^'n^'n)"
   598 proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2)
   598 proof (simp add: det_def setsum_right_distrib cong add: setprod_cong, rule setsum_cong2)
   599   let ?U = "UNIV :: 'n set"
   599   let ?U = "UNIV :: 'n set"
   600   let ?PU = "{p. p permutes ?U}"
   600   let ?PU = "{p. p permutes ?U}"
   601   fix p assume pU: "p \<in> ?PU"
   601   fix p assume pU: "p \<in> ?PU"
   606   then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
   606   then show "?s * (\<Prod>xa\<in>?U. c xa * a xa $ p xa) =
   607         setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps)
   607         setprod c ?U * (?s* (\<Prod>xa\<in>?U. a xa $ p xa))" by (simp add: ring_simps)
   608 qed
   608 qed
   609 
   609 
   610 lemma det_mul:
   610 lemma det_mul:
   611   fixes A B :: "'a::ordered_idom^'n^'n::finite"
   611   fixes A B :: "'a::ordered_idom^'n^'n"
   612   shows "det (A ** B) = det A * det B"
   612   shows "det (A ** B) = det A * det B"
   613 proof-
   613 proof-
   614   let ?U = "UNIV :: 'n set"
   614   let ?U = "UNIV :: 'n set"
   615   let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
   615   let ?F = "{f. (\<forall>i\<in> ?U. f i \<in> ?U) \<and> (\<forall>i. i \<notin> ?U \<longrightarrow> f i = i)}"
   616   let ?PU = "{p. p permutes ?U}"
   616   let ?PU = "{p. p permutes ?U}"
   698 (* ------------------------------------------------------------------------- *)
   698 (* ------------------------------------------------------------------------- *)
   699 (* Relation to invertibility.                                                *)
   699 (* Relation to invertibility.                                                *)
   700 (* ------------------------------------------------------------------------- *)
   700 (* ------------------------------------------------------------------------- *)
   701 
   701 
   702 lemma invertible_left_inverse:
   702 lemma invertible_left_inverse:
   703   fixes A :: "real^'n^'n::finite"
   703   fixes A :: "real^'n^'n"
   704   shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). B** A = mat 1)"
   704   shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). B** A = mat 1)"
   705   by (metis invertible_def matrix_left_right_inverse)
   705   by (metis invertible_def matrix_left_right_inverse)
   706 
   706 
   707 lemma invertible_righ_inverse:
   707 lemma invertible_righ_inverse:
   708   fixes A :: "real^'n^'n::finite"
   708   fixes A :: "real^'n^'n"
   709   shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)"
   709   shows "invertible A \<longleftrightarrow> (\<exists>(B::real^'n^'n). A** B = mat 1)"
   710   by (metis invertible_def matrix_left_right_inverse)
   710   by (metis invertible_def matrix_left_right_inverse)
   711 
   711 
   712 lemma invertible_det_nz:
   712 lemma invertible_det_nz:
   713   fixes A::"real ^'n^'n::finite"
   713   fixes A::"real ^'n^'n"
   714   shows "invertible A \<longleftrightarrow> det A \<noteq> 0"
   714   shows "invertible A \<longleftrightarrow> det A \<noteq> 0"
   715 proof-
   715 proof-
   716   {assume "invertible A"
   716   {assume "invertible A"
   717     then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1"
   717     then obtain B :: "real ^'n^'n" where B: "A ** B = mat 1"
   718       unfolding invertible_righ_inverse by blast
   718       unfolding invertible_righ_inverse by blast
   759 (* ------------------------------------------------------------------------- *)
   759 (* ------------------------------------------------------------------------- *)
   760 (* Cramer's rule.                                                            *)
   760 (* Cramer's rule.                                                            *)
   761 (* ------------------------------------------------------------------------- *)
   761 (* ------------------------------------------------------------------------- *)
   762 
   762 
   763 lemma cramer_lemma_transp:
   763 lemma cramer_lemma_transp:
   764   fixes A:: "'a::ordered_idom^'n^'n::finite" and x :: "'a ^'n::finite"
   764   fixes A:: "'a::ordered_idom^'n^'n" and x :: "'a ^'n"
   765   shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
   765   shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
   766                            else row i A)::'a^'n^'n) = x$k * det A"
   766                            else row i A)::'a^'n^'n) = x$k * det A"
   767   (is "?lhs = ?rhs")
   767   (is "?lhs = ?rhs")
   768 proof-
   768 proof-
   769   let ?U = "UNIV :: 'n set"
   769   let ?U = "UNIV :: 'n set"
   795     unfolding th001[of k "\<lambda>i. row i A"]
   795     unfolding th001[of k "\<lambda>i. row i A"]
   796     unfolding thd1  by (simp add: ring_simps)
   796     unfolding thd1  by (simp add: ring_simps)
   797 qed
   797 qed
   798 
   798 
   799 lemma cramer_lemma:
   799 lemma cramer_lemma:
   800   fixes A :: "'a::ordered_idom ^'n^'n::finite"
   800   fixes A :: "'a::ordered_idom ^'n^'n"
   801   shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A"
   801   shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A"
   802 proof-
   802 proof-
   803   let ?U = "UNIV :: 'n set"
   803   let ?U = "UNIV :: 'n set"
   804   have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transp A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
   804   have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transp A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
   805     by (auto simp add: row_transp intro: setsum_cong2)
   805     by (auto simp add: row_transp intro: setsum_cong2)
   809   apply (subst det_transp[symmetric])
   809   apply (subst det_transp[symmetric])
   810   apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def)
   810   apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def)
   811 qed
   811 qed
   812 
   812 
   813 lemma cramer:
   813 lemma cramer:
   814   fixes A ::"real^'n^'n::finite"
   814   fixes A ::"real^'n^'n"
   815   assumes d0: "det A \<noteq> 0"
   815   assumes d0: "det A \<noteq> 0"
   816   shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
   816   shows "A *v x = b \<longleftrightarrow> x = (\<chi> k. det(\<chi> i j. if j=k then b$i else A$i$j :: real^'n^'n) / det A)"
   817 proof-
   817 proof-
   818   from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1"
   818   from d0 obtain B where B: "A ** B = mat 1" "B ** A = mat 1"
   819     unfolding invertible_det_nz[symmetric] invertible_def by blast
   819     unfolding invertible_det_nz[symmetric] invertible_def by blast
   838   apply auto
   838   apply auto
   839   apply (erule_tac x=v in allE)+
   839   apply (erule_tac x=v in allE)+
   840   apply (simp add: real_vector_norm_def)
   840   apply (simp add: real_vector_norm_def)
   841   by (simp add: dot_norm  linear_add[symmetric])
   841   by (simp add: dot_norm  linear_add[symmetric])
   842 
   842 
   843 definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n::finite) \<longleftrightarrow> transp Q ** Q = mat 1 \<and> Q ** transp Q = mat 1"
   843 definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> transp Q ** Q = mat 1 \<and> Q ** transp Q = mat 1"
   844 
   844 
   845 lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n::finite)  \<longleftrightarrow> transp Q ** Q = mat 1"
   845 lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n)  \<longleftrightarrow> transp Q ** Q = mat 1"
   846   by (metis matrix_left_right_inverse orthogonal_matrix_def)
   846   by (metis matrix_left_right_inverse orthogonal_matrix_def)
   847 
   847 
   848 lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n::finite)"
   848 lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)"
   849   by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid)
   849   by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid)
   850 
   850 
   851 lemma orthogonal_matrix_mul:
   851 lemma orthogonal_matrix_mul:
   852   fixes A :: "real ^'n^'n::finite"
   852   fixes A :: "real ^'n^'n"
   853   assumes oA : "orthogonal_matrix A"
   853   assumes oA : "orthogonal_matrix A"
   854   and oB: "orthogonal_matrix B"
   854   and oB: "orthogonal_matrix B"
   855   shows "orthogonal_matrix(A ** B)"
   855   shows "orthogonal_matrix(A ** B)"
   856   using oA oB
   856   using oA oB
   857   unfolding orthogonal_matrix matrix_transp_mul
   857   unfolding orthogonal_matrix matrix_transp_mul
   858   apply (subst matrix_mul_assoc)
   858   apply (subst matrix_mul_assoc)
   859   apply (subst matrix_mul_assoc[symmetric])
   859   apply (subst matrix_mul_assoc[symmetric])
   860   by (simp add: matrix_mul_rid)
   860   by (simp add: matrix_mul_rid)
   861 
   861 
   862 lemma orthogonal_transformation_matrix:
   862 lemma orthogonal_transformation_matrix:
   863   fixes f:: "real^'n \<Rightarrow> real^'n::finite"
   863   fixes f:: "real^'n \<Rightarrow> real^'n"
   864   shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)"
   864   shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)"
   865   (is "?lhs \<longleftrightarrow> ?rhs")
   865   (is "?lhs \<longleftrightarrow> ?rhs")
   866 proof-
   866 proof-
   867   let ?mf = "matrix f"
   867   let ?mf = "matrix f"
   868   let ?ot = "orthogonal_transformation f"
   868   let ?ot = "orthogonal_transformation f"
   891       by (simp add: dot_matrix_product matrix_mul_lid)}
   891       by (simp add: dot_matrix_product matrix_mul_lid)}
   892   ultimately show ?thesis by blast
   892   ultimately show ?thesis by blast
   893 qed
   893 qed
   894 
   894 
   895 lemma det_orthogonal_matrix:
   895 lemma det_orthogonal_matrix:
   896   fixes Q:: "'a::ordered_idom^'n^'n::finite"
   896   fixes Q:: "'a::ordered_idom^'n^'n"
   897   assumes oQ: "orthogonal_matrix Q"
   897   assumes oQ: "orthogonal_matrix Q"
   898   shows "det Q = 1 \<or> det Q = - 1"
   898   shows "det Q = 1 \<or> det Q = - 1"
   899 proof-
   899 proof-
   900 
   900 
   901   have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x")
   901   have th: "\<And>x::'a. x = 1 \<or> x = - 1 \<longleftrightarrow> x*x = 1" (is "\<And>x::'a. ?ths x")
   916 
   916 
   917 (* ------------------------------------------------------------------------- *)
   917 (* ------------------------------------------------------------------------- *)
   918 (* Linearity of scaling, and hence isometry, that preserves origin.          *)
   918 (* Linearity of scaling, and hence isometry, that preserves origin.          *)
   919 (* ------------------------------------------------------------------------- *)
   919 (* ------------------------------------------------------------------------- *)
   920 lemma scaling_linear:
   920 lemma scaling_linear:
   921   fixes f :: "real ^'n \<Rightarrow> real ^'n::finite"
   921   fixes f :: "real ^'n \<Rightarrow> real ^'n"
   922   assumes f0: "f 0 = 0" and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
   922   assumes f0: "f 0 = 0" and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
   923   shows "linear f"
   923   shows "linear f"
   924 proof-
   924 proof-
   925   {fix v w
   925   {fix v w
   926     {fix x note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] }
   926     {fix x note fd[rule_format, of x 0, unfolded dist_norm f0 diff_0_right] }
   932   show ?thesis unfolding linear_def vector_eq
   932   show ?thesis unfolding linear_def vector_eq
   933     by (simp add: dot_lmult dot_ladd dot_rmult dot_radd fc ring_simps)
   933     by (simp add: dot_lmult dot_ladd dot_rmult dot_radd fc ring_simps)
   934 qed
   934 qed
   935 
   935 
   936 lemma isometry_linear:
   936 lemma isometry_linear:
   937   "f (0:: real^'n) = (0:: real^'n::finite) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y
   937   "f (0:: real^'n) = (0:: real^'n) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y
   938         \<Longrightarrow> linear f"
   938         \<Longrightarrow> linear f"
   939 by (rule scaling_linear[where c=1]) simp_all
   939 by (rule scaling_linear[where c=1]) simp_all
   940 
   940 
   941 (* ------------------------------------------------------------------------- *)
   941 (* ------------------------------------------------------------------------- *)
   942 (* Hence another formulation of orthogonal transformation.                   *)
   942 (* Hence another formulation of orthogonal transformation.                   *)
   943 (* ------------------------------------------------------------------------- *)
   943 (* ------------------------------------------------------------------------- *)
   944 
   944 
   945 lemma orthogonal_transformation_isometry:
   945 lemma orthogonal_transformation_isometry:
   946   "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n::finite) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
   946   "orthogonal_transformation f \<longleftrightarrow> f(0::real^'n) = (0::real^'n) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
   947   unfolding orthogonal_transformation
   947   unfolding orthogonal_transformation
   948   apply (rule iffI)
   948   apply (rule iffI)
   949   apply clarify
   949   apply clarify
   950   apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_norm)
   950   apply (clarsimp simp add: linear_0 linear_sub[symmetric] dist_norm)
   951   apply (rule conjI)
   951   apply (rule conjI)
   960 (* ------------------------------------------------------------------------- *)
   960 (* ------------------------------------------------------------------------- *)
   961 (* Can extend an isometry from unit sphere.                                  *)
   961 (* Can extend an isometry from unit sphere.                                  *)
   962 (* ------------------------------------------------------------------------- *)
   962 (* ------------------------------------------------------------------------- *)
   963 
   963 
   964 lemma isometry_sphere_extend:
   964 lemma isometry_sphere_extend:
   965   fixes f:: "real ^'n \<Rightarrow> real ^'n::finite"
   965   fixes f:: "real ^'n \<Rightarrow> real ^'n"
   966   assumes f1: "\<forall>x. norm x = 1 \<longrightarrow> norm (f x) = 1"
   966   assumes f1: "\<forall>x. norm x = 1 \<longrightarrow> norm (f x) = 1"
   967   and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y"
   967   and fd1: "\<forall> x y. norm x = 1 \<longrightarrow> norm y = 1 \<longrightarrow> dist (f x) (f y) = dist x y"
   968   shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)"
   968   shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)"
   969 proof-
   969 proof-
   970   {fix x y x' y' x0 y0 x0' y0' :: "real ^'n"
   970   {fix x y x' y' x0 y0 x0' y0' :: "real ^'n"
  1032 
  1032 
  1033 definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
  1033 definition "rotation_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = 1"
  1034 definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
  1034 definition "rotoinversion_matrix Q \<longleftrightarrow> orthogonal_matrix Q \<and> det Q = - 1"
  1035 
  1035 
  1036 lemma orthogonal_rotation_or_rotoinversion:
  1036 lemma orthogonal_rotation_or_rotoinversion:
  1037   fixes Q :: "'a::ordered_idom^'n^'n::finite"
  1037   fixes Q :: "'a::ordered_idom^'n^'n"
  1038   shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
  1038   shows " orthogonal_matrix Q \<longleftrightarrow> rotation_matrix Q \<or> rotoinversion_matrix Q"
  1039   by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
  1039   by (metis rotoinversion_matrix_def rotation_matrix_def det_orthogonal_matrix)
  1040 (* ------------------------------------------------------------------------- *)
  1040 (* ------------------------------------------------------------------------- *)
  1041 (* Explicit formulas for low dimensions.                                     *)
  1041 (* Explicit formulas for low dimensions.                                     *)
  1042 (* ------------------------------------------------------------------------- *)
  1042 (* ------------------------------------------------------------------------- *)