src/HOL/Analysis/Cartesian_Space.thy
author wenzelm
Sat, 20 Feb 2021 13:42:37 +0100
changeset 73255 7e2a9a8c2b85
parent 71633 07bec530f02e
child 73466 ee1c4962671c
permissions -rw-r--r--
provide naproche-755224402e36;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
68189
6163c90694ef tuned headers;
wenzelm
parents: 68074
diff changeset
     1
(*  Title:      HOL/Analysis/Cartesian_Space.thy
6163c90694ef tuned headers;
wenzelm
parents: 68074
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
6163c90694ef tuned headers;
wenzelm
parents: 68074
diff changeset
     3
    Author:     Jose Divasón <jose.divasonm at unirioja.es>
6163c90694ef tuned headers;
wenzelm
parents: 68074
diff changeset
     4
    Author:     Jesús Aransay <jesus-maria.aransay at unirioja.es>
6163c90694ef tuned headers;
wenzelm
parents: 68074
diff changeset
     5
    Author:     Johannes Hölzl, VU Amsterdam
6163c90694ef tuned headers;
wenzelm
parents: 68074
diff changeset
     6
    Author:     Fabian Immler, TUM
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
     7
*)
68189
6163c90694ef tuned headers;
wenzelm
parents: 68074
diff changeset
     8
69665
60110f6d0b4e tuned headers
nipkow
parents: 69064
diff changeset
     9
section "Linear Algebra on Finite Cartesian Products"
60110f6d0b4e tuned headers
nipkow
parents: 69064
diff changeset
    10
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    11
theory Cartesian_Space
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    12
  imports
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    13
    Finite_Cartesian_Product Linear_Algebra
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    14
begin
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    15
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
    16
subsection\<^marker>\<open>tag unimportant\<close> \<open>Type @{typ \<open>'a ^ 'n\<close>} and fields as vector spaces\<close> (*much of the following
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    17
is really basic linear algebra, check for overlap? rename subsection?  *)
69667
82bb6225588b tuned headers
nipkow
parents: 69666
diff changeset
    18
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    19
definition "cart_basis = {axis i 1 | i. i\<in>UNIV}"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    20
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    21
lemma finite_cart_basis: "finite (cart_basis)" unfolding cart_basis_def
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    22
  using finite_Atleast_Atmost_nat by fastforce
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    23
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    24
lemma card_cart_basis: "card (cart_basis::('a::zero_neq_one^'i) set) = CARD('i)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    25
  unfolding cart_basis_def Setcompr_eq_image
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    26
  by (rule card_image) (auto simp: inj_on_def axis_eq_axis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    27
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    28
interpretation vec: vector_space "(*s) "
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    29
  by unfold_locales (vector algebra_simps)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    30
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    31
lemma independent_cart_basis:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    32
  "vec.independent (cart_basis)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    33
proof (rule vec.independent_if_scalars_zero)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    34
  show "finite (cart_basis)" using finite_cart_basis .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    35
  fix f::"('a, 'b) vec \<Rightarrow> 'a" and x::"('a, 'b) vec"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    36
  assume eq_0: "(\<Sum>x\<in>cart_basis. f x *s x) = 0" and x_in: "x \<in> cart_basis"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    37
  obtain i where x: "x = axis i 1" using x_in unfolding cart_basis_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    38
  have sum_eq_0: "(\<Sum>x\<in>(cart_basis) - {x}. f x * (x $ i)) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    39
  proof (rule sum.neutral, rule ballI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    40
    fix xa assume xa: "xa \<in> cart_basis - {x}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    41
    obtain a where a: "xa = axis a 1" and a_not_i: "a \<noteq> i"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    42
      using xa x unfolding cart_basis_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    43
    have "xa $ i = 0" unfolding a axis_def using a_not_i by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    44
    thus "f xa * xa $ i = 0" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    45
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    46
  have "0 = (\<Sum>x\<in>cart_basis. f x *s x) $ i" using eq_0 by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    47
  also have "... = (\<Sum>x\<in>cart_basis. (f x *s x) $ i)" unfolding sum_component ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    48
  also have "... = (\<Sum>x\<in>cart_basis. f x * (x $ i))" unfolding vector_smult_component ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    49
  also have "... = f x * (x $ i) + (\<Sum>x\<in>(cart_basis) - {x}. f x * (x $ i))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    50
    by (rule sum.remove[OF finite_cart_basis x_in])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    51
  also have "... =  f x * (x $ i)" unfolding sum_eq_0 by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    52
  also have "... = f x" unfolding x axis_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    53
  finally show "f x = 0" ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    54
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    55
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
    56
lemma span_cart_basis:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    57
  "vec.span (cart_basis) = UNIV"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    58
proof (auto)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    59
  fix x::"('a, 'b) vec"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    60
  let ?f="\<lambda>v. x $ (THE i. v = axis i 1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    61
  show "x \<in> vec.span (cart_basis)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    62
    apply (unfold vec.span_finite[OF finite_cart_basis])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    63
    apply (rule image_eqI[of _ _ ?f])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    64
     apply (subst  vec_eq_iff)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    65
     apply clarify
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    66
  proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    67
    fix i::'b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    68
    let ?w = "axis i (1::'a)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    69
    have the_eq_i: "(THE a. ?w = axis a 1) = i"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    70
      by (rule the_equality, auto simp: axis_eq_axis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    71
    have sum_eq_0: "(\<Sum>v\<in>(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i) = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    72
    proof (rule sum.neutral, rule ballI)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    73
      fix xa::"('a, 'b) vec"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    74
      assume xa: "xa \<in> cart_basis - {?w}"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    75
      obtain j where j: "xa = axis j 1" and i_not_j: "i \<noteq> j" using xa unfolding cart_basis_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    76
      have the_eq_j: "(THE i. xa = axis i 1) = j"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    77
      proof (rule the_equality)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    78
        show "xa = axis j 1" using j .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    79
        show "\<And>i. xa = axis i 1 \<Longrightarrow> i = j" by (metis axis_eq_axis j zero_neq_one)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    80
      qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    81
      show "x $ (THE i. xa = axis i 1) * xa $ i = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    82
        apply (subst (2) j)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    83
        unfolding the_eq_j unfolding axis_def using i_not_j by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    84
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    85
    have "(\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) *s v) $ i =
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    86
  (\<Sum>v\<in>cart_basis. (x $ (THE i. v = axis i 1) *s v) $ i)" unfolding sum_component ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    87
    also have "... = (\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) * v $ i)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    88
      unfolding vector_smult_component ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    89
    also have "... = x $ (THE a. ?w = axis a 1) * ?w $ i + (\<Sum>v\<in>(cart_basis) - {?w}. x $ (THE i. v = axis i 1) * v $ i)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    90
      by (rule sum.remove[OF finite_cart_basis], auto simp add: cart_basis_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    91
    also have "... = x $ (THE a. ?w = axis a 1) * ?w $ i" unfolding sum_eq_0 by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    92
    also have "... = x $ i" unfolding the_eq_i unfolding axis_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    93
    finally show "x $ i = (\<Sum>v\<in>cart_basis. x $ (THE i. v = axis i 1) *s v) $ i" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    94
  qed simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    95
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    96
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    97
(*Some interpretations:*)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
    98
interpretation vec: finite_dimensional_vector_space "(*s)" "cart_basis"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    99
  by (unfold_locales, auto simp add: finite_cart_basis independent_cart_basis span_cart_basis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   100
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   101
lemma matrix_vector_mul_linear_gen[intro, simp]:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   102
  "Vector_Spaces.linear (*s) (*s) ((*v) A)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   103
  by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   104
    (vector matrix_vector_mult_def sum.distrib algebra_simps)+
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   105
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   106
lemma span_vec_eq: "vec.span X = span X"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   107
  and dim_vec_eq: "vec.dim X = dim X"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   108
  and dependent_vec_eq: "vec.dependent X = dependent X"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   109
  and subspace_vec_eq: "vec.subspace X = subspace X"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   110
  for X::"(real^'n) set"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   111
  unfolding span_raw_def dim_raw_def dependent_raw_def subspace_raw_def
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   112
  by (auto simp: scalar_mult_eq_scaleR)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   113
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   114
lemma linear_componentwise:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   115
  fixes f:: "'a::field ^'m \<Rightarrow> 'a ^ 'n"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   116
  assumes lf: "Vector_Spaces.linear (*s) (*s) f"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   117
  shows "(f x)$j = sum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   118
proof -
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   119
  interpret lf: Vector_Spaces.linear "(*s)" "(*s)" f
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   120
    using lf .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   121
  let ?M = "(UNIV :: 'm set)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   122
  let ?N = "(UNIV :: 'n set)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   123
  have fM: "finite ?M" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   124
  have "?rhs = (sum (\<lambda>i. (x$i) *s (f (axis i 1))) ?M)$j"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   125
    unfolding sum_component by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   126
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   127
    unfolding lf.sum[symmetric] lf.scale[symmetric]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   128
    unfolding basis_expansion by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   129
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   130
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   131
interpretation vec: Vector_Spaces.linear "(*s)" "(*s)" "(*v) A"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   132
  using matrix_vector_mul_linear_gen.
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   133
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   134
interpretation vec: finite_dimensional_vector_space_pair "(*s)" cart_basis "(*s)" cart_basis ..
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   135
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   136
lemma matrix_works:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   137
  assumes lf: "Vector_Spaces.linear (*s) (*s) f"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   138
  shows "matrix f *v x = f (x::'a::field ^ 'n)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   139
  apply (simp add: matrix_def matrix_vector_mult_def vec_eq_iff mult.commute)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   140
  apply clarify
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   141
  apply (rule linear_componentwise[OF lf, symmetric])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   142
  done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   143
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   144
lemma matrix_of_matrix_vector_mul[simp]: "matrix(\<lambda>x. A *v (x :: 'a::field ^ 'n)) = A"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   145
  by (simp add: matrix_eq matrix_works)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   146
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   147
lemma matrix_compose_gen:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   148
  assumes lf: "Vector_Spaces.linear (*s) (*s) (f::'a::{field}^'n \<Rightarrow> 'a^'m)"
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   149
    and lg: "Vector_Spaces.linear (*s) (*s) (g::'a^'m \<Rightarrow> 'a^_)"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   150
  shows "matrix (g o f) = matrix g ** matrix f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   151
  using lf lg Vector_Spaces.linear_compose[OF lf lg] matrix_works[OF Vector_Spaces.linear_compose[OF lf lg]]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   152
  by (simp add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   153
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   154
lemma matrix_compose:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   155
  assumes "linear (f::real^'n \<Rightarrow> real^'m)" "linear (g::real^'m \<Rightarrow> real^_)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   156
  shows "matrix (g o f) = matrix g ** matrix f"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   157
  using matrix_compose_gen[of f g] assms
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   158
  by (simp add: linear_def scalar_mult_eq_scaleR)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   159
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   160
lemma left_invertible_transpose:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   161
  "(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   162
  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   163
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   164
lemma right_invertible_transpose:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   165
  "(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   166
  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   167
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   168
lemma linear_matrix_vector_mul_eq:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   169
  "Vector_Spaces.linear (*s) (*s) f \<longleftrightarrow> linear (f :: real^'n \<Rightarrow> real ^'m)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   170
  by (simp add: scalar_mult_eq_scaleR linear_def)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   171
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   172
lemma matrix_vector_mul[simp]:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   173
  "Vector_Spaces.linear (*s) (*s) g \<Longrightarrow> (\<lambda>y. matrix g *v y) = g"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   174
  "linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   175
  "bounded_linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   176
  for f :: "real^'n \<Rightarrow> real ^'m"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   177
  by (simp_all add: ext matrix_works linear_matrix_vector_mul_eq linear_linear)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   178
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   179
lemma matrix_left_invertible_injective:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   180
  fixes A :: "'a::field^'n^'m"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   181
  shows "(\<exists>B. B ** A = mat 1) \<longleftrightarrow> inj ((*v) A)"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   182
proof safe
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   183
  fix B
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   184
  assume B: "B ** A = mat 1"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   185
  show "inj ((*v) A)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   186
    unfolding inj_on_def
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   187
      by (metis B matrix_vector_mul_assoc matrix_vector_mul_lid)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   188
next
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   189
  assume "inj ((*v) A)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   190
  from vec.linear_injective_left_inverse[OF matrix_vector_mul_linear_gen this]
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   191
  obtain g where "Vector_Spaces.linear (*s) (*s) g" and g: "g \<circ> (*v) A = id"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   192
    by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   193
  have "matrix g ** A = mat 1"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   194
    by (metis matrix_vector_mul_linear_gen \<open>Vector_Spaces.linear (*s) (*s) g\<close> g matrix_compose_gen
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   195
        matrix_eq matrix_id_mat_1 matrix_vector_mul(1))
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   196
  then show "\<exists>B. B ** A = mat 1"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   197
    by metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   198
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   199
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   200
lemma matrix_left_invertible_ker:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   201
  "(\<exists>B. (B::'a::{field} ^'m^'n) ** (A::'a::{field}^'n^'m) = mat 1) \<longleftrightarrow> (\<forall>x. A *v x = 0 \<longrightarrow> x = 0)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   202
  unfolding matrix_left_invertible_injective
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   203
  using vec.inj_on_iff_eq_0[OF vec.subspace_UNIV, of A]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   204
  by (simp add: inj_on_def)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   205
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   206
lemma matrix_right_invertible_surjective:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   207
  "(\<exists>B. (A::'a::field^'n^'m) ** (B::'a::field^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   208
proof -
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   209
  { fix B :: "'a ^'m^'n"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   210
    assume AB: "A ** B = mat 1"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   211
    { fix x :: "'a ^ 'm"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   212
      have "A *v (B *v x) = x"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   213
        by (simp add: matrix_vector_mul_assoc AB) }
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   214
    hence "surj ((*v) A)" unfolding surj_def by metis }
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   215
  moreover
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   216
  { assume sf: "surj ((*v) A)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   217
    from vec.linear_surjective_right_inverse[OF _ this]
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   218
    obtain g:: "'a ^'m \<Rightarrow> 'a ^'n" where g: "Vector_Spaces.linear (*s) (*s) g" "(*v) A \<circ> g = id"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   219
      by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   220
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   221
    have "A ** (matrix g) = mat 1"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   222
      unfolding matrix_eq  matrix_vector_mul_lid
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   223
        matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   224
      using g(2) unfolding o_def fun_eq_iff id_def
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   225
      .
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   226
    hence "\<exists>B. A ** (B::'a^'m^'n) = mat 1" by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   227
  }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   228
  ultimately show ?thesis unfolding surj_def by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   229
qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   230
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   231
lemma matrix_left_invertible_independent_columns:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   232
  fixes A :: "'a::{field}^'n^'m"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   233
  shows "(\<exists>(B::'a ^'m^'n). B ** A = mat 1) \<longleftrightarrow>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   234
      (\<forall>c. sum (\<lambda>i. c i *s column i A) (UNIV :: 'n set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   235
    (is "?lhs \<longleftrightarrow> ?rhs")
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   236
proof -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   237
  let ?U = "UNIV :: 'n set"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   238
  { assume k: "\<forall>x. A *v x = 0 \<longrightarrow> x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   239
    { fix c i
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   240
      assume c: "sum (\<lambda>i. c i *s column i A) ?U = 0" and i: "i \<in> ?U"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   241
      let ?x = "\<chi> i. c i"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   242
      have th0:"A *v ?x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   243
        using c
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   244
        by (vector matrix_mult_sum)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   245
      from k[rule_format, OF th0] i
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   246
      have "c i = 0" by (vector vec_eq_iff)}
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   247
    hence ?rhs by blast }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   248
  moreover
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   249
  { assume H: ?rhs
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   250
    { fix x assume x: "A *v x = 0"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   251
      let ?c = "\<lambda>i. ((x$i ):: 'a)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   252
      from H[rule_format, of ?c, unfolded matrix_mult_sum[symmetric], OF x]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   253
      have "x = 0" by vector }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   254
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   255
  ultimately show ?thesis unfolding matrix_left_invertible_ker by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   256
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   257
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   258
lemma matrix_right_invertible_independent_rows:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   259
  fixes A :: "'a::{field}^'n^'m"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   260
  shows "(\<exists>(B::'a^'m^'n). A ** B = mat 1) \<longleftrightarrow>
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   261
    (\<forall>c. sum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   262
  unfolding left_invertible_transpose[symmetric]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   263
    matrix_left_invertible_independent_columns
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   264
  by (simp add:)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   265
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   266
lemma matrix_right_invertible_span_columns:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   267
  "(\<exists>(B::'a::field ^'n^'m). (A::'a ^'m^'n) ** B = mat 1) \<longleftrightarrow>
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   268
    vec.span (columns A) = UNIV" (is "?lhs = ?rhs")
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   269
proof -
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   270
  let ?U = "UNIV :: 'm set"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   271
  have fU: "finite ?U" by simp
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   272
  have lhseq: "?lhs \<longleftrightarrow> (\<forall>y. \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y)"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   273
    unfolding matrix_right_invertible_surjective matrix_mult_sum surj_def
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   274
    by (simp add: eq_commute)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   275
  have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> vec.span (columns A))" by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   276
  { assume h: ?lhs
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   277
    { fix x:: "'a ^'n"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   278
      from h[unfolded lhseq, rule_format, of x] obtain y :: "'a ^'m"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   279
        where y: "sum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   280
      have "x \<in> vec.span (columns A)"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   281
        unfolding y[symmetric] scalar_mult_eq_scaleR
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   282
      proof (rule vec.span_sum [OF vec.span_scale])
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   283
        show "column i A \<in> vec.span (columns A)" for i
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   284
          using columns_def vec.span_superset by auto
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   285
      qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   286
    }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   287
    then have ?rhs unfolding rhseq by blast }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   288
  moreover
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   289
  { assume h:?rhs
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   290
    let ?P = "\<lambda>(y::'a ^'n). \<exists>(x::'a^'m). sum (\<lambda>i. (x$i) *s column i A) ?U = y"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   291
    { fix y
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   292
      have "y \<in> vec.span (columns A)"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   293
        unfolding h by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   294
      then have "?P y"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   295
      proof (induction rule: vec.span_induct_alt)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   296
        case base
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   297
        then show ?case
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   298
          by (metis (full_types) matrix_mult_sum matrix_vector_mult_0_right)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   299
      next
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   300
        case (step c y1 y2)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   301
        from step obtain i where i: "i \<in> ?U" "y1 = column i A"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   302
          unfolding columns_def by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   303
        obtain x:: "'a ^'m" where x: "sum (\<lambda>i. (x$i) *s column i A) ?U = y2"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   304
          using step by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   305
        let ?x = "(\<chi> j. if j = i then c + (x$i) else (x$j))::'a^'m"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   306
        show ?case
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   307
        proof (rule exI[where x= "?x"], vector, auto simp add: i x[symmetric] if_distrib distrib_left if_distribR cong del: if_weak_cong)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   308
          fix j
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   309
          have th: "\<forall>xa \<in> ?U. (if xa = i then (c + (x$i)) * ((column xa A)$j)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   310
              else (x$xa) * ((column xa A$j))) = (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   311
            using i(1) by (simp add: field_simps)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   312
          have "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   313
              else (x$xa) * ((column xa A$j))) ?U = sum (\<lambda>xa. (if xa = i then c * ((column i A)$j) else 0) + ((x$xa) * ((column xa A)$j))) ?U"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   314
            by (rule sum.cong[OF refl]) (use th in blast)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   315
          also have "\<dots> = sum (\<lambda>xa. if xa = i then c * ((column i A)$j) else 0) ?U + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   316
            by (simp add: sum.distrib)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   317
          also have "\<dots> = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   318
            unfolding sum.delta[OF fU]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   319
            using i(1) by simp
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   320
          finally show "sum (\<lambda>xa. if xa = i then (c + (x$i)) * ((column xa A)$j)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   321
            else (x$xa) * ((column xa A$j))) ?U = c * ((column i A)$j) + sum (\<lambda>xa. ((x$xa) * ((column xa A)$j))) ?U" .
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   322
        qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   323
      qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   324
    }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   325
    then have ?lhs unfolding lhseq ..
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   326
  }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   327
  ultimately show ?thesis by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   328
qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   329
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   330
lemma matrix_left_invertible_span_rows_gen:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   331
  "(\<exists>(B::'a^'m^'n). B ** (A::'a::field^'n^'m) = mat 1) \<longleftrightarrow> vec.span (rows A) = UNIV"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   332
  unfolding right_invertible_transpose[symmetric]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   333
  unfolding columns_transpose[symmetric]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   334
  unfolding matrix_right_invertible_span_columns
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   335
  ..
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   336
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   337
lemma matrix_left_invertible_span_rows:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   338
  "(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   339
  using matrix_left_invertible_span_rows_gen[of A] by (simp add: span_vec_eq)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   340
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   341
lemma matrix_left_right_inverse:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   342
  fixes A A' :: "'a::{field}^'n^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   343
  shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   344
proof -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   345
  { fix A A' :: "'a ^'n^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   346
    assume AA': "A ** A' = mat 1"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   347
    have sA: "surj ((*v) A)"
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   348
      using AA' matrix_right_invertible_surjective by auto
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   349
    from vec.linear_surjective_isomorphism[OF matrix_vector_mul_linear_gen sA]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   350
    obtain f' :: "'a ^'n \<Rightarrow> 'a ^'n"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   351
      where f': "Vector_Spaces.linear (*s) (*s) f'" "\<forall>x. f' (A *v x) = x" "\<forall>x. A *v f' x = x" by blast
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   352
    have th: "matrix f' ** A = mat 1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   353
      by (simp add: matrix_eq matrix_works[OF f'(1)]
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   354
          matrix_vector_mul_assoc[symmetric] f'(2)[rule_format])
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   355
    hence "(matrix f' ** A) ** A' = mat 1 ** A'" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   356
    hence "matrix f' = A'"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   357
      by (simp add: matrix_mul_assoc[symmetric] AA')
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   358
    hence "matrix f' ** A = A' ** A" by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   359
    hence "A' ** A = mat 1" by (simp add: th)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   360
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   361
  then show ?thesis by blast
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   362
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   363
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   364
lemma invertible_left_inverse:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   365
  fixes A :: "'a::{field}^'n^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   366
  shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). B ** A = mat 1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   367
  by (metis invertible_def matrix_left_right_inverse)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   368
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   369
lemma invertible_right_inverse:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   370
  fixes A :: "'a::{field}^'n^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   371
  shows "invertible A \<longleftrightarrow> (\<exists>(B::'a^'n^'n). A** B = mat 1)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   372
  by (metis invertible_def matrix_left_right_inverse)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   373
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   374
lemma invertible_mult:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   375
  assumes inv_A: "invertible A"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   376
  and inv_B: "invertible B"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   377
  shows "invertible (A**B)"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   378
proof -
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   379
  obtain A' where AA': "A ** A' = mat 1" and A'A: "A' ** A = mat 1" 
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   380
    using inv_A unfolding invertible_def by blast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   381
  obtain B' where BB': "B ** B' = mat 1" and B'B: "B' ** B = mat 1" 
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   382
    using inv_B unfolding invertible_def by blast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   383
  show ?thesis
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   384
  proof (unfold invertible_def, rule exI[of _ "B'**A'"], rule conjI)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   385
    have "A ** B ** (B' ** A') = A ** (B ** (B' ** A'))" 
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   386
      using matrix_mul_assoc[of A B "(B' ** A')", symmetric] .
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   387
    also have "... = A ** (B ** B' ** A')" unfolding matrix_mul_assoc[of B "B'" "A'"] ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   388
    also have "... = A ** (mat 1 ** A')" unfolding BB' ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   389
    also have "... = A ** A'" unfolding matrix_mul_lid ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   390
    also have "... = mat 1" unfolding AA' ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   391
    finally show "A ** B ** (B' ** A') = mat (1::'a)" .    
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   392
    have "B' ** A' ** (A ** B) = B' ** (A' ** (A ** B))" using matrix_mul_assoc[of B' A' "(A ** B)", symmetric] .
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   393
    also have "... =  B' ** (A' ** A ** B)" unfolding matrix_mul_assoc[of A' A B] ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   394
    also have "... =  B' ** (mat 1 ** B)" unfolding A'A ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   395
    also have "... = B' ** B"  unfolding matrix_mul_lid ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   396
    also have "... = mat 1" unfolding B'B ..
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   397
    finally show "B' ** A' ** (A ** B) = mat 1" .
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   398
  qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   399
qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   400
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   401
lemma transpose_invertible:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   402
  fixes A :: "real^'n^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   403
  assumes "invertible A"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   404
  shows "invertible (transpose A)"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   405
  by (meson assms invertible_def matrix_left_right_inverse right_invertible_transpose)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   406
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   407
lemma vector_matrix_mul_assoc:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   408
  fixes v :: "('a::comm_semiring_1)^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   409
  shows "(v v* M) v* N = v v* (M ** N)"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   410
proof -
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   411
  from matrix_vector_mul_assoc
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   412
  have "transpose N *v (transpose M *v v) = (transpose N ** transpose M) *v v" by fast
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   413
  thus "(v v* M) v* N = v v* (M ** N)"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   414
    by (simp add: matrix_transpose_mul [symmetric])
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   415
qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   416
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   417
lemma matrix_scaleR_vector_ac:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   418
  fixes A :: "real^('m::finite)^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   419
  shows "A *v (k *\<^sub>R v) = k *\<^sub>R A *v v"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   420
  by (metis matrix_vector_mult_scaleR transpose_scalar vector_scaleR_matrix_ac vector_transpose_matrix)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   421
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   422
lemma scaleR_matrix_vector_assoc:
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   423
  fixes A :: "real^('m::finite)^'n"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   424
  shows "k *\<^sub>R (A *v v) = k *\<^sub>R A *v v"
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   425
  by (metis matrix_scaleR_vector_ac matrix_vector_mult_scaleR)
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   426
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   427
(*Finally, some interesting theorems and interpretations that don't appear in any file of the
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   428
  library.*)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   429
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   430
locale linear_first_finite_dimensional_vector_space =
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   431
  l?: Vector_Spaces.linear scaleB scaleC f +
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   432
  B?: finite_dimensional_vector_space scaleB BasisB
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   433
  for scaleB :: "('a::field => 'b::ab_group_add => 'b)" (infixr "*b" 75)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   434
  and scaleC :: "('a => 'c::ab_group_add => 'c)" (infixr "*c" 75)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   435
  and BasisB :: "('b set)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   436
  and f :: "('b=>'c)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   437
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   438
lemma vec_dim_card: "vec.dim (UNIV::('a::{field}^'n) set) = CARD ('n)"
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   439
proof -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   440
  let ?f="\<lambda>i::'n. axis i (1::'a)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   441
  have "vec.dim (UNIV::('a::{field}^'n) set) = card (cart_basis::('a^'n) set)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   442
    unfolding vec.dim_UNIV ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   443
  also have "... = card ({i. i\<in> UNIV}::('n) set)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   444
    proof (rule bij_betw_same_card[of ?f, symmetric], unfold bij_betw_def, auto)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   445
      show "inj (\<lambda>i::'n. axis i (1::'a))"  by (simp add: inj_on_def axis_eq_axis)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   446
      fix i::'n
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   447
      show "axis i 1 \<in> cart_basis" unfolding cart_basis_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   448
      fix x::"'a^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   449
      assume "x \<in> cart_basis"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   450
      thus "x \<in> range (\<lambda>i. axis i 1)" unfolding cart_basis_def by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   451
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   452
  also have "... = CARD('n)" by auto
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   453
  finally show ?thesis .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   454
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   455
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   456
interpretation vector_space_over_itself: vector_space "(*) :: 'a::field \<Rightarrow> 'a \<Rightarrow> 'a"
69667
82bb6225588b tuned headers
nipkow
parents: 69666
diff changeset
   457
by unfold_locales (simp_all add: algebra_simps)
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   458
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   459
lemmas [simp del] = vector_space_over_itself.scale_scale
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   460
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   461
interpretation vector_space_over_itself: finite_dimensional_vector_space
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   462
  "(*) :: 'a::field => 'a => 'a" "{1}"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   463
  by unfold_locales (auto simp: vector_space_over_itself.span_singleton)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   464
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   465
lemma dimension_eq_1[code_unfold]: "vector_space_over_itself.dimension TYPE('a::field)= 1"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   466
  unfolding vector_space_over_itself.dimension_def by simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   467
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   468
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   469
lemma dim_subset_UNIV_cart_gen:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   470
  fixes S :: "('a::field^'n) set"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   471
  shows "vec.dim S \<le> CARD('n)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   472
  by (metis vec.dim_eq_full vec.dim_subset_UNIV vec.span_UNIV vec_dim_card)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   473
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   474
lemma dim_subset_UNIV_cart:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   475
  fixes S :: "(real^'n) set"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   476
  shows "dim S \<le> CARD('n)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   477
  using dim_subset_UNIV_cart_gen[of S] by (simp add: dim_vec_eq)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   478
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   479
text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close>
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   480
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   481
lemma matrix_mult_dot: "A *v x = (\<chi> i. inner (A$i) x)"
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   482
  by (simp add: matrix_vector_mult_def inner_vec_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   483
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   484
lemma adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   485
  apply (rule adjoint_unique)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   486
  apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   487
    sum_distrib_right sum_distrib_left)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   488
  apply (subst sum.swap)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   489
  apply (simp add:  ac_simps)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   490
  done
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   491
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   492
lemma matrix_adjoint: assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)"
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   493
  shows "matrix(adjoint f) = transpose(matrix f)"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   494
proof -
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   495
  have "matrix(adjoint f) = matrix(adjoint ((*v) (matrix f)))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   496
    by (simp add: lf)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   497
  also have "\<dots> = transpose(matrix f)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   498
    unfolding adjoint_matrix matrix_of_matrix_vector_mul
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   499
    apply rule
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   500
    done
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   501
  finally show ?thesis .
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   502
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   503
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   504
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
   505
subsection\<open> Rank of a matrix\<close>
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   506
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   507
text\<open>Equivalence of row and column rank is taken from George Mackiw's paper, Mathematics Magazine 1995, p. 285.\<close>
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   508
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   509
lemma matrix_vector_mult_in_columnspace_gen:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   510
  fixes A :: "'a::field^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   511
  shows "(A *v x) \<in> vec.span(columns A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   512
  apply (simp add: matrix_vector_column columns_def transpose_def column_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   513
  apply (intro vec.span_sum vec.span_scale)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   514
  apply (force intro: vec.span_base)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   515
  done
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   516
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   517
lemma matrix_vector_mult_in_columnspace:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   518
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   519
  shows "(A *v x) \<in> span(columns A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   520
  using matrix_vector_mult_in_columnspace_gen[of A x] by (simp add: span_vec_eq)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   521
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   522
lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   523
  by (simp add: subspace_def orthogonal_clauses)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   524
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   525
lemma orthogonal_nullspace_rowspace:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   526
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   527
  assumes 0: "A *v x = 0" and y: "y \<in> span(rows A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   528
  shows "orthogonal x y"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   529
  using y
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   530
proof (induction rule: span_induct)
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   531
  case base
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   532
  then show ?case
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   533
    by (simp add: subspace_orthogonal_to_vector)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   534
next
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   535
  case (step v)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   536
  then obtain i where "v = row i A"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   537
    by (auto simp: rows_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   538
  with 0 show ?case
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   539
    unfolding orthogonal_def inner_vec_def matrix_vector_mult_def row_def
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   540
    by (simp add: mult.commute) (metis (no_types) vec_lambda_beta zero_index)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   541
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   542
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   543
lemma nullspace_inter_rowspace:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   544
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   545
  shows "A *v x = 0 \<and> x \<in> span(rows A) \<longleftrightarrow> x = 0"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   546
  using orthogonal_nullspace_rowspace orthogonal_self span_zero matrix_vector_mult_0_right
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   547
  by blast
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   548
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   549
lemma matrix_vector_mul_injective_on_rowspace:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   550
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   551
  shows "\<lbrakk>A *v x = A *v y; x \<in> span(rows A); y \<in> span(rows A)\<rbrakk> \<Longrightarrow> x = y"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   552
  using nullspace_inter_rowspace [of A "x-y"]
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   553
  by (metis diff_eq_diff_eq diff_self matrix_vector_mult_diff_distrib span_diff)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   554
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   555
definition\<^marker>\<open>tag important\<close> rank :: "'a::field^'n^'m=>nat"
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   556
  where row_rank_def_gen: "rank A \<equiv> vec.dim(rows A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   557
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   558
lemma row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m"
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   559
  by (auto simp: row_rank_def_gen dim_vec_eq)
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   560
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   561
lemma dim_rows_le_dim_columns:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   562
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   563
  shows "dim(rows A) \<le> dim(columns A)"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   564
proof -
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   565
  have "dim (span (rows A)) \<le> dim (span (columns A))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   566
  proof -
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   567
    obtain B where "independent B" "span(rows A) \<subseteq> span B"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   568
              and B: "B \<subseteq> span(rows A)""card B = dim (span(rows A))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   569
      using basis_exists [of "span(rows A)"] by metis
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   570
    with span_subspace have eq: "span B = span(rows A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   571
      by auto
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   572
    then have inj: "inj_on ((*v) A) (span B)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   573
      by (simp add: inj_on_def matrix_vector_mul_injective_on_rowspace)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   574
    then have ind: "independent ((*v) A ` B)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   575
      by (rule linear_independent_injective_image [OF Finite_Cartesian_Product.matrix_vector_mul_linear \<open>independent B\<close>])
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   576
    have "dim (span (rows A)) \<le> card ((*v) A ` B)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   577
      unfolding B(2)[symmetric]
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   578
      using inj
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   579
      by (auto simp: card_image inj_on_subset span_superset)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   580
    also have "\<dots> \<le> dim (span (columns A))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   581
      using _ ind
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   582
      by (rule independent_card_le_dim) (auto intro!: matrix_vector_mult_in_columnspace)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   583
    finally show ?thesis .
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   584
  qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   585
  then show ?thesis
71044
nipkow
parents: 70136
diff changeset
   586
    by (simp)
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   587
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   588
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   589
lemma column_rank_def:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   590
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   591
  shows "rank A = dim(columns A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   592
  unfolding row_rank_def
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   593
  by (metis columns_transpose dim_rows_le_dim_columns le_antisym rows_transpose)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   594
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   595
lemma rank_transpose:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   596
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   597
  shows "rank(transpose A) = rank A"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   598
  by (metis column_rank_def row_rank_def rows_transpose)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   599
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   600
lemma matrix_vector_mult_basis:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   601
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   602
  shows "A *v (axis k 1) = column k A"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   603
  by (simp add: cart_eq_inner_axis column_def matrix_mult_dot)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   604
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   605
lemma columns_image_basis:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   606
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   607
  shows "columns A = (*v) A ` (range (\<lambda>i. axis i 1))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   608
  by (force simp: columns_def matrix_vector_mult_basis [symmetric])
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   609
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   610
lemma rank_dim_range:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   611
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   612
  shows "rank A = dim(range (\<lambda>x. A *v x))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   613
  unfolding column_rank_def
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   614
proof (rule span_eq_dim)
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   615
  have "span (columns A) \<subseteq> span (range ((*v) A))" (is "?l \<subseteq> ?r")
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   616
    by (simp add: columns_image_basis image_subsetI span_mono)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   617
  then show "?l = ?r"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   618
    by (metis (no_types, lifting) image_subset_iff matrix_vector_mult_in_columnspace
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   619
        span_eq span_span)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   620
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   621
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   622
lemma rank_bound:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   623
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   624
  shows "rank A \<le> min CARD('m) (CARD('n))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   625
  by (metis (mono_tags, lifting) dim_subset_UNIV_cart min.bounded_iff
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   626
      column_rank_def row_rank_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   627
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   628
lemma full_rank_injective:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   629
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   630
  shows "rank A = CARD('n) \<longleftrightarrow> inj ((*v) A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   631
  by (simp add: matrix_left_invertible_injective [symmetric] matrix_left_invertible_span_rows row_rank_def
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   632
      dim_eq_full [symmetric] card_cart_basis vec.dimension_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   633
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   634
lemma full_rank_surjective:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   635
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   636
  shows "rank A = CARD('m) \<longleftrightarrow> surj ((*v) A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   637
  by (simp add: matrix_right_invertible_surjective [symmetric] left_invertible_transpose [symmetric]
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   638
                matrix_left_invertible_injective full_rank_injective [symmetric] rank_transpose)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   639
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   640
lemma rank_I: "rank(mat 1::real^'n^'n) = CARD('n)"
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   641
  by (simp add: full_rank_injective inj_on_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   642
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   643
lemma less_rank_noninjective:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   644
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   645
  shows "rank A < CARD('n) \<longleftrightarrow> \<not> inj ((*v) A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   646
using less_le rank_bound by (auto simp: full_rank_injective [symmetric])
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   647
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   648
lemma matrix_nonfull_linear_equations_eq:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   649
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   650
  shows "(\<exists>x. (x \<noteq> 0) \<and> A *v x = 0) \<longleftrightarrow> rank A \<noteq> CARD('n)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   651
  by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   652
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   653
lemma rank_eq_0: "rank A = 0 \<longleftrightarrow> A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0"
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   654
  for A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   655
  by (auto simp: rank_dim_range matrix_eq)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   656
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   657
lemma rank_mul_le_right:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   658
  fixes A :: "real^'n^'m" and B :: "real^'p^'n"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   659
  shows "rank(A ** B) \<le> rank B"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   660
proof -
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   661
  have "rank(A ** B) \<le> dim ((*v) A ` range ((*v) B))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   662
    by (auto simp: rank_dim_range image_comp o_def matrix_vector_mul_assoc)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   663
  also have "\<dots> \<le> rank B"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   664
    by (simp add: rank_dim_range dim_image_le)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   665
  finally show ?thesis .
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   666
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   667
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   668
lemma rank_mul_le_left:
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   669
  fixes A :: "real^'n^'m" and B :: "real^'p^'n"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   670
  shows "rank(A ** B) \<le> rank A"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   671
  by (metis matrix_transpose_mul rank_mul_le_right rank_transpose)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   672
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   673
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   674
subsection\<^marker>\<open>tag unimportant\<close> \<open>Lemmas for working on \<open>real^1/2/3/4\<close>\<close>
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   675
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   676
lemma exhaust_2:
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   677
  fixes x :: 2
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   678
  shows "x = 1 \<or> x = 2"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   679
proof (induct x)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   680
  case (of_int z)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   681
  then have "0 \<le> z" and "z < 2" by simp_all
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   682
  then have "z = 0 | z = 1" by arith
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   683
  then show ?case by auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   684
qed
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   685
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   686
lemma forall_2: "(\<forall>i::2. P i) \<longleftrightarrow> P 1 \<and> P 2"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   687
  by (metis exhaust_2)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   688
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   689
lemma exhaust_3:
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   690
  fixes x :: 3
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   691
  shows "x = 1 \<or> x = 2 \<or> x = 3"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   692
proof (induct x)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   693
  case (of_int z)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   694
  then have "0 \<le> z" and "z < 3" by simp_all
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   695
  then have "z = 0 \<or> z = 1 \<or> z = 2" by arith
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   696
  then show ?case by auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   697
qed
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   698
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   699
lemma forall_3: "(\<forall>i::3. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   700
  by (metis exhaust_3)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   701
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   702
lemma exhaust_4:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   703
  fixes x :: 4
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   704
  shows "x = 1 \<or> x = 2 \<or> x = 3 \<or> x = 4"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   705
proof (induct x)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   706
  case (of_int z)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   707
  then have "0 \<le> z" and "z < 4" by simp_all
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   708
  then have "z = 0 \<or> z = 1 \<or> z = 2 \<or> z = 3" by arith
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   709
  then show ?case by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   710
qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   711
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   712
lemma forall_4: "(\<forall>i::4. P i) \<longleftrightarrow> P 1 \<and> P 2 \<and> P 3 \<and> P 4"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   713
  by (metis exhaust_4)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   714
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   715
lemma UNIV_1 [simp]: "UNIV = {1::1}"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   716
  by (auto simp add: num1_eq_iff)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   717
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   718
lemma UNIV_2: "UNIV = {1::2, 2::2}"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   719
  using exhaust_2 by auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   720
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   721
lemma UNIV_3: "UNIV = {1::3, 2::3, 3::3}"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   722
  using exhaust_3 by auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   723
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   724
lemma UNIV_4: "UNIV = {1::4, 2::4, 3::4, 4::4}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   725
  using exhaust_4 by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   726
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   727
lemma sum_1: "sum f (UNIV::1 set) = f 1"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   728
  unfolding UNIV_1 by simp
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   729
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   730
lemma sum_2: "sum f (UNIV::2 set) = f 1 + f 2"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   731
  unfolding UNIV_2 by simp
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   732
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   733
lemma sum_3: "sum f (UNIV::3 set) = f 1 + f 2 + f 3"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   734
  unfolding UNIV_3 by (simp add: ac_simps)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   735
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   736
lemma sum_4: "sum f (UNIV::4 set) = f 1 + f 2 + f 3 + f 4"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   737
  unfolding UNIV_4 by (simp add: ac_simps)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69723
diff changeset
   738
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   739
subsection\<^marker>\<open>tag unimportant\<close>\<open>The collapse of the general concepts to dimension one\<close>
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   740
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   741
lemma vector_one: "(x::'a ^1) = (\<chi> i. (x$1))"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   742
  by (simp add: vec_eq_iff)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   743
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   744
lemma forall_one: "(\<forall>(x::'a ^1). P x) \<longleftrightarrow> (\<forall>x. P(\<chi> i. x))"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   745
  apply auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   746
  apply (erule_tac x= "x$1" in allE)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   747
  apply (simp only: vector_one[symmetric])
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   748
  done
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   749
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   750
lemma norm_vector_1: "norm (x :: _^1) = norm (x$1)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   751
  by (simp add: norm_vec_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   752
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   753
lemma dist_vector_1:
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   754
  fixes x :: "'a::real_normed_vector^1"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   755
  shows "dist x y = dist (x$1) (y$1)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   756
  by (simp add: dist_norm norm_vector_1)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   757
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   758
lemma norm_real: "norm(x::real ^ 1) = \<bar>x$1\<bar>"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   759
  by (simp add: norm_vector_1)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   760
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   761
lemma dist_real: "dist(x::real ^ 1) y = \<bar>(x$1) - (y$1)\<bar>"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   762
  by (auto simp add: norm_real dist_norm)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   763
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   764
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   765
subsection\<^marker>\<open>tag unimportant\<close>\<open>Routine results connecting the types \<^typ>\<open>real^1\<close> and \<^typ>\<open>real\<close>\<close>
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   766
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   767
lemma vector_one_nth [simp]:
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   768
  fixes x :: "'a^1" shows "vec (x $ 1) = x"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   769
  by (metis vec_def vector_one)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   770
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   771
lemma tendsto_at_within_vector_1:
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   772
  fixes S :: "'a :: metric_space set"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   773
  assumes "(f \<longlongrightarrow> fx) (at x within S)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   774
  shows "((\<lambda>y::'a^1. \<chi> i. f (y $ 1)) \<longlongrightarrow> (vec fx::'a^1)) (at (vec x) within vec ` S)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   775
proof (rule topological_tendstoI)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   776
  fix T :: "('a^1) set"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   777
  assume "open T" "vec fx \<in> T"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   778
  have "\<forall>\<^sub>F x in at x within S. f x \<in> (\<lambda>x. x $ 1) ` T"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   779
    using \<open>open T\<close> \<open>vec fx \<in> T\<close> assms open_image_vec_nth tendsto_def by fastforce
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   780
  then show "\<forall>\<^sub>F x::'a^1 in at (vec x) within vec ` S. (\<chi> i. f (x $ 1)) \<in> T"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   781
    unfolding eventually_at dist_norm [symmetric]
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   782
    by (rule ex_forward)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   783
       (use \<open>open T\<close> in 
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   784
         \<open>fastforce simp: dist_norm dist_vec_def L2_set_def image_iff vector_one open_vec_def\<close>)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   785
qed
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   786
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   787
lemma has_derivative_vector_1:
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   788
  assumes der_g: "(g has_derivative (\<lambda>x. x * g' a)) (at a within S)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   789
  shows "((\<lambda>x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' a))
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   790
         (at ((vec a)::real^1) within vec ` S)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   791
    using der_g
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   792
    apply (auto simp: Deriv.has_derivative_within bounded_linear_scaleR_right norm_vector_1)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   793
    apply (drule tendsto_at_within_vector_1, vector)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   794
    apply (auto simp: algebra_simps eventually_at tendsto_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   795
    done
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   796
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   797
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   798
subsection\<^marker>\<open>tag unimportant\<close>\<open>Explicit vector construction from lists\<close>
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   799
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   800
definition "vector l = (\<chi> i. foldr (\<lambda>x f n. fun_upd (f (n+1)) n x) l (\<lambda>n x. 0) 1 i)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   801
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   802
lemma vector_1 [simp]: "(vector[x]) $1 = x"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   803
  unfolding vector_def by simp
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   804
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   805
lemma vector_2 [simp]: "(vector[x,y]) $1 = x" "(vector[x,y] :: 'a^2)$2 = (y::'a::zero)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   806
  unfolding vector_def by simp_all
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   807
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   808
lemma vector_3 [simp]:
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   809
 "(vector [x,y,z] ::('a::zero)^3)$1 = x"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   810
 "(vector [x,y,z] ::('a::zero)^3)$2 = y"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   811
 "(vector [x,y,z] ::('a::zero)^3)$3 = z"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   812
  unfolding vector_def by simp_all
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   813
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   814
lemma forall_vector_1: "(\<forall>v::'a::zero^1. P v) \<longleftrightarrow> (\<forall>x. P(vector[x]))"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   815
  by (metis vector_1 vector_one)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   816
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   817
lemma forall_vector_2: "(\<forall>v::'a::zero^2. P v) \<longleftrightarrow> (\<forall>x y. P(vector[x, y]))"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   818
  apply auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   819
  apply (erule_tac x="v$1" in allE)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   820
  apply (erule_tac x="v$2" in allE)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   821
  apply (subgoal_tac "vector [v$1, v$2] = v")
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   822
  apply simp
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   823
  apply (vector vector_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   824
  apply (simp add: forall_2)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   825
  done
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   826
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   827
lemma forall_vector_3: "(\<forall>v::'a::zero^3. P v) \<longleftrightarrow> (\<forall>x y z. P(vector[x, y, z]))"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   828
  apply auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   829
  apply (erule_tac x="v$1" in allE)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   830
  apply (erule_tac x="v$2" in allE)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   831
  apply (erule_tac x="v$3" in allE)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   832
  apply (subgoal_tac "vector [v$1, v$2, v$3] = v")
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   833
  apply simp
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   834
  apply (vector vector_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   835
  apply (simp add: forall_3)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   836
  done
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   837
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   838
subsection\<^marker>\<open>tag unimportant\<close> \<open>lambda skolemization on cartesian products\<close>
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   839
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   840
lemma lambda_skolem: "(\<forall>i. \<exists>x. P i x) \<longleftrightarrow>
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   841
   (\<exists>x::'a ^ 'n. \<forall>i. P i (x $ i))" (is "?lhs \<longleftrightarrow> ?rhs")
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   842
proof -
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   843
  let ?S = "(UNIV :: 'n set)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   844
  { assume H: "?rhs"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   845
    then have ?lhs by auto }
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   846
  moreover
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   847
  { assume H: "?lhs"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   848
    then obtain f where f:"\<forall>i. P i (f i)" unfolding choice_iff by metis
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   849
    let ?x = "(\<chi> i. (f i)) :: 'a ^ 'n"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   850
    { fix i
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   851
      from f have "P i (f i)" by metis
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   852
      then have "P i (?x $ i)" by auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   853
    }
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   854
    hence "\<forall>i. P i (?x$i)" by metis
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   855
    hence ?rhs by metis }
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   856
  ultimately show ?thesis by metis
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   857
qed
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   858
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   859
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   860
text \<open>The same result in terms of square matrices.\<close>
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   861
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   862
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   863
text \<open>Considering an n-element vector as an n-by-1 or 1-by-n matrix.\<close>
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   864
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   865
definition "rowvector v = (\<chi> i j. (v$j))"
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   866
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   867
definition "columnvector v = (\<chi> i j. (v$i))"
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   868
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   869
lemma transpose_columnvector: "transpose(columnvector v) = rowvector v"
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   870
  by (simp add: transpose_def rowvector_def columnvector_def vec_eq_iff)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   871
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   872
lemma transpose_rowvector: "transpose(rowvector v) = columnvector v"
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   873
  by (simp add: transpose_def columnvector_def rowvector_def vec_eq_iff)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   874
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   875
lemma dot_rowvector_columnvector: "columnvector (A *v v) = A ** columnvector v"
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   876
  by (vector columnvector_def matrix_matrix_mult_def matrix_vector_mult_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   877
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   878
lemma dot_matrix_product:
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   879
  "(x::real^'n) \<bullet> y = (((rowvector x ::real^'n^1) ** (columnvector y :: real^1^'n))$1)$1"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   880
  by (vector matrix_matrix_mult_def rowvector_def columnvector_def inner_vec_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   881
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   882
lemma dot_matrix_vector_mul:
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   883
  fixes A B :: "real ^'n ^'n" and x y :: "real ^'n"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   884
  shows "(A *v x) \<bullet> (B *v y) =
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   885
      (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   886
  unfolding dot_matrix_product transpose_columnvector[symmetric]
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   887
    dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc ..
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   888
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   889
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   890
lemma dim_substandard_cart: "vec.dim {x::'a::field^'n. \<forall>i. i \<notin> d \<longrightarrow> x$i = 0} = card d"
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   891
  (is "vec.dim ?A = _")
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   892
proof (rule vec.dim_unique)
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   893
  let ?B = "((\<lambda>x. axis x 1) ` d)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   894
  have subset_basis: "?B \<subseteq> cart_basis"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   895
    by (auto simp: cart_basis_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   896
  show "?B \<subseteq> ?A"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   897
    by (auto simp: axis_def)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   898
  show "vec.independent ((\<lambda>x. axis x 1) ` d)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   899
    using subset_basis
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   900
    by (rule vec.independent_mono[OF vec.independent_Basis])
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   901
  have "x \<in> vec.span ?B" if "\<forall>i. i \<notin> d \<longrightarrow> x $ i = 0" for x::"'a^'n"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   902
  proof -
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   903
    have "finite ?B"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   904
      using subset_basis finite_cart_basis
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   905
      by (rule finite_subset)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   906
    have "x = (\<Sum>i\<in>UNIV. x $ i *s axis i 1)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   907
      by (rule basis_expansion[symmetric])
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   908
    also have "\<dots> = (\<Sum>i\<in>d. (x $ i) *s axis i 1)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   909
      by (rule sum.mono_neutral_cong_right) (auto simp: that)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   910
    also have "\<dots> \<in> vec.span ?B"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   911
      by (simp add: vec.span_sum vec.span_clauses)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   912
    finally show "x \<in> vec.span ?B" .
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   913
  qed
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   914
  then show "?A \<subseteq> vec.span ?B" by auto
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   915
qed (simp add: card_image inj_on_def axis_eq_axis)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   916
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   917
lemma affinity_inverses:
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   918
  assumes m0: "m \<noteq> (0::'a::field)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   919
  shows "(\<lambda>x. m *s x + c) \<circ> (\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) = id"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   920
  "(\<lambda>x. inverse(m) *s x + (-(inverse(m) *s c))) \<circ> (\<lambda>x. m *s x + c) = id"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   921
  using m0
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   922
  by (auto simp add: fun_eq_iff vector_add_ldistrib diff_conv_add_uminus simp del: add_uminus_conv_diff)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   923
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   924
lemma vector_affinity_eq:
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   925
  assumes m0: "(m::'a::field) \<noteq> 0"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   926
  shows "m *s x + c = y \<longleftrightarrow> x = inverse m *s y + -(inverse m *s c)"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   927
proof
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   928
  assume h: "m *s x + c = y"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   929
  hence "m *s x = y - c" by (simp add: field_simps)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   930
  hence "inverse m *s (m *s x) = inverse m *s (y - c)" by simp
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   931
  then show "x = inverse m *s y + - (inverse m *s c)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   932
    using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   933
next
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   934
  assume h: "x = inverse m *s y + - (inverse m *s c)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   935
  show "m *s x + c = y" unfolding h
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   936
    using m0 by (simp add: vector_smult_assoc vector_ssub_ldistrib)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   937
qed
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   938
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   939
lemma vector_eq_affinity:
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   940
    "(m::'a::field) \<noteq> 0 ==> (y = m *s x + c \<longleftrightarrow> inverse(m) *s y + -(inverse(m) *s c) = x)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   941
  using vector_affinity_eq[where m=m and x=x and y=y and c=c]
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   942
  by metis
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   943
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   944
lemma vector_cart:
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   945
  fixes f :: "real^'n \<Rightarrow> real"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   946
  shows "(\<chi> i. f (axis i 1)) = (\<Sum>i\<in>Basis. f i *\<^sub>R i)"
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   947
  unfolding euclidean_eq_iff[where 'a="real^'n"]
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   948
  by simp (simp add: Basis_vec_def inner_axis)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   949
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   950
lemma const_vector_cart:"((\<chi> i. d)::real^'n) = (\<Sum>i\<in>Basis. d *\<^sub>R i)"
69669
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   951
  by (rule vector_cart)
de2f0a24b0f0 Reorg of material
nipkow
parents: 69667
diff changeset
   952
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   953
subsection\<^marker>\<open>tag unimportant\<close> \<open>Explicit formulas for low dimensions\<close>
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   954
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   955
lemma  prod_neutral_const: "prod f {(1::nat)..1} = f 1"
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   956
  by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   957
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   958
lemma  prod_2: "prod f {(1::nat)..2} = f 1 * f 2"
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   959
  by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   960
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   961
lemma  prod_3: "prod f {(1::nat)..3} = f 1 * f 2 * f 3"
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   962
  by (simp add: eval_nat_numeral atLeastAtMostSuc_conv mult.commute)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   963
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   964
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
   965
subsection  \<open>Orthogonality of a matrix\<close>
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   966
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   967
definition\<^marker>\<open>tag important\<close>  "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow>
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   968
  transpose Q ** Q = mat 1 \<and> Q ** transpose Q = mat 1"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   969
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   970
lemma  orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n) \<longleftrightarrow> transpose Q ** Q = mat 1"
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   971
  by (metis matrix_left_right_inverse orthogonal_matrix_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   972
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   973
lemma  orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)"
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   974
  by (simp add: orthogonal_matrix_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   975
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   976
proposition  orthogonal_matrix_mul:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   977
  fixes A :: "real ^'n^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   978
  assumes  "orthogonal_matrix A" "orthogonal_matrix B"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   979
  shows "orthogonal_matrix(A ** B)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   980
  using assms
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   981
  by (simp add: orthogonal_matrix matrix_transpose_mul matrix_left_right_inverse matrix_mul_assoc)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   982
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   983
proposition  orthogonal_transformation_matrix:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   984
  fixes f:: "real^'n \<Rightarrow> real^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   985
  shows "orthogonal_transformation f \<longleftrightarrow> linear f \<and> orthogonal_matrix(matrix f)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   986
  (is "?lhs \<longleftrightarrow> ?rhs")
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
   987
proof -
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   988
  let ?mf = "matrix f"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   989
  let ?ot = "orthogonal_transformation f"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   990
  let ?U = "UNIV :: 'n set"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   991
  have fU: "finite ?U" by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   992
  let ?m1 = "mat 1 :: real ^'n^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   993
  {
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   994
    assume ot: ?ot
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   995
    from ot have lf: "Vector_Spaces.linear (*s) (*s) f" and fd: "\<And>v w. f v \<bullet> f w = v \<bullet> w"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   996
      unfolding orthogonal_transformation_def orthogonal_matrix linear_def scalar_mult_eq_scaleR
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   997
      by blast+
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   998
    {
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
   999
      fix i j
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1000
      let ?A = "transpose ?mf ** ?mf"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1001
      have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1002
        "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1003
        by simp_all
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1004
      from fd[of "axis i 1" "axis j 1",
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1005
        simplified matrix_works[OF lf, symmetric] dot_matrix_vector_mul]
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1006
      have "?A$i$j = ?m1 $ i $ j"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1007
        by (simp add: inner_vec_def matrix_matrix_mult_def columnvector_def rowvector_def
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1008
            th0 sum.delta[OF fU] mat_def axis_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1009
    }
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1010
    then have "orthogonal_matrix ?mf"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1011
      unfolding orthogonal_matrix
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1012
      by vector
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1013
    with lf have ?rhs
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1014
      unfolding linear_def scalar_mult_eq_scaleR
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1015
      by blast
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1016
  }
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1017
  moreover
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1018
  {
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1019
    assume lf: "Vector_Spaces.linear (*s) (*s) f" and om: "orthogonal_matrix ?mf"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1020
    from lf om have ?lhs
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1021
      unfolding orthogonal_matrix_def norm_eq orthogonal_transformation
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1022
      apply (simp only: matrix_works[OF lf, symmetric] dot_matrix_vector_mul)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1023
      apply (simp add: dot_matrix_product linear_def scalar_mult_eq_scaleR)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1024
      done
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1025
  }
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1026
  ultimately show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1027
    by (auto simp: linear_def scalar_mult_eq_scaleR)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1028
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1029
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1030
71044
nipkow
parents: 70136
diff changeset
  1031
subsection \<open>Finding an Orthogonal Matrix\<close>
nipkow
parents: 70136
diff changeset
  1032
nipkow
parents: 70136
diff changeset
  1033
text \<open>We can find an orthogonal matrix taking any unit vector to any other.\<close>
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1034
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1035
lemma  orthogonal_matrix_transpose [simp]:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1036
     "orthogonal_matrix(transpose A) \<longleftrightarrow> orthogonal_matrix A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1037
  by (auto simp: orthogonal_matrix_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1038
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1039
lemma  orthogonal_matrix_orthonormal_columns:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1040
  fixes A :: "real^'n^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1041
  shows "orthogonal_matrix A \<longleftrightarrow>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1042
          (\<forall>i. norm(column i A) = 1) \<and>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1043
          (\<forall>i j. i \<noteq> j \<longrightarrow> orthogonal (column i A) (column j A))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1044
  by (auto simp: orthogonal_matrix matrix_mult_transpose_dot_column vec_eq_iff mat_def norm_eq_1 orthogonal_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1045
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1046
lemma  orthogonal_matrix_orthonormal_rows:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1047
  fixes A :: "real^'n^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1048
  shows "orthogonal_matrix A \<longleftrightarrow>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1049
          (\<forall>i. norm(row i A) = 1) \<and>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1050
          (\<forall>i j. i \<noteq> j \<longrightarrow> orthogonal (row i A) (row j A))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1051
  using orthogonal_matrix_orthonormal_columns [of "transpose A"] by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1052
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1053
proposition  orthogonal_matrix_exists_basis:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1054
  fixes a :: "real^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1055
  assumes "norm a = 1"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1056
  obtains A where "orthogonal_matrix A" "A *v (axis k 1) = a"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1057
proof -
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1058
  obtain S where "a \<in> S" "pairwise orthogonal S" and noS: "\<And>x. x \<in> S \<Longrightarrow> norm x = 1"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1059
   and "independent S" "card S = CARD('n)" "span S = UNIV"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1060
    using vector_in_orthonormal_basis assms by force
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1061
  then obtain f0 where "bij_betw f0 (UNIV::'n set) S"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1062
    by (metis finite_class.finite_UNIV finite_same_card_bij finiteI_independent)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1063
  then obtain f where f: "bij_betw f (UNIV::'n set) S" and a: "a = f k"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1064
    using bij_swap_iff [of k "inv f0 a" f0]
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1065
    by (metis UNIV_I \<open>a \<in> S\<close> bij_betw_inv_into_right bij_betw_swap_iff swap_apply(1))
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1066
  show thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1067
  proof
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1068
    have [simp]: "\<And>i. norm (f i) = 1"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1069
      using bij_betwE [OF \<open>bij_betw f UNIV S\<close>] by (blast intro: noS)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1070
    have [simp]: "\<And>i j. i \<noteq> j \<Longrightarrow> orthogonal (f i) (f j)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1071
      using \<open>pairwise orthogonal S\<close> \<open>bij_betw f UNIV S\<close>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1072
      by (auto simp: pairwise_def bij_betw_def inj_on_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1073
    show "orthogonal_matrix (\<chi> i j. f j $ i)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1074
      by (simp add: orthogonal_matrix_orthonormal_columns column_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1075
    show "(\<chi> i j. f j $ i) *v axis k 1 = a"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1076
      by (simp add: matrix_vector_mult_def axis_def a if_distrib cong: if_cong)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1077
  qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1078
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1079
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1080
lemma  orthogonal_transformation_exists_1:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1081
  fixes a b :: "real^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1082
  assumes "norm a = 1" "norm b = 1"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1083
  obtains f where "orthogonal_transformation f" "f a = b"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1084
proof -
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1085
  obtain k::'n where True
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1086
    by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1087
  obtain A B where AB: "orthogonal_matrix A" "orthogonal_matrix B" and eq: "A *v (axis k 1) = a" "B *v (axis k 1) = b"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1088
    using orthogonal_matrix_exists_basis assms by metis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1089
  let ?f = "\<lambda>x. (B ** transpose A) *v x"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1090
  show thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1091
  proof
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1092
    show "orthogonal_transformation ?f"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1093
      by (subst orthogonal_transformation_matrix)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1094
        (auto simp: AB orthogonal_matrix_mul)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1095
  next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1096
    show "?f a = b"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1097
      using \<open>orthogonal_matrix A\<close> unfolding orthogonal_matrix_def
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1098
      by (metis eq matrix_mul_rid matrix_vector_mul_assoc)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1099
  qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1100
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1101
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1102
proposition  orthogonal_transformation_exists:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1103
  fixes a b :: "real^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1104
  assumes "norm a = norm b"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1105
  obtains f where "orthogonal_transformation f" "f a = b"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1106
proof (cases "a = 0 \<or> b = 0")
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1107
  case True
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1108
  with assms show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1109
    using that by force
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1110
next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1111
  case False
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1112
  then obtain f where f: "orthogonal_transformation f" and eq: "f (a /\<^sub>R norm a) = (b /\<^sub>R norm b)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1113
    by (auto intro: orthogonal_transformation_exists_1 [of "a /\<^sub>R norm a" "b /\<^sub>R norm b"])
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1114
  show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1115
  proof
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1116
    interpret linear f
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1117
      using f by (simp add: orthogonal_transformation_linear)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1118
    have "f a /\<^sub>R norm a = f (a /\<^sub>R norm a)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1119
      by (simp add: scale)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1120
    also have "\<dots> = b /\<^sub>R norm a"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1121
      by (simp add: eq assms [symmetric])
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1122
    finally show "f a = b"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1123
      using False by auto
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1124
  qed (use f in auto)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1125
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1126
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1127
71044
nipkow
parents: 70136
diff changeset
  1128
subsection  \<open>Scaling and isometry\<close>
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1129
71044
nipkow
parents: 70136
diff changeset
  1130
proposition scaling_linear:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1131
  fixes f :: "'a::real_inner \<Rightarrow> 'a::real_inner"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1132
  assumes f0: "f 0 = 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1133
    and fd: "\<forall>x y. dist (f x) (f y) = c * dist x y"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1134
  shows "linear f"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1135
proof -
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1136
  {
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1137
    fix v w
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1138
    have "norm (f x) = c * norm x" for x
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1139
      by (metis dist_0_norm f0 fd)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1140
    then have "f v \<bullet> f w = c\<^sup>2 * (v \<bullet> w)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1141
      unfolding dot_norm_neg dist_norm[symmetric]
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1142
      by (simp add: fd power2_eq_square field_simps)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1143
  }
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1144
  then show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1145
    unfolding linear_iff vector_eq[where 'a="'a"] scalar_mult_eq_scaleR
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1146
    by (simp add: inner_add field_simps)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1147
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1148
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1149
lemma  isometry_linear:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1150
  "f (0::'a::real_inner) = (0::'a) \<Longrightarrow> \<forall>x y. dist(f x) (f y) = dist x y \<Longrightarrow> linear f"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1151
  by (rule scaling_linear[where c=1]) simp_all
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1152
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1153
text \<open>Hence another formulation of orthogonal transformation\<close>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1154
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1155
proposition  orthogonal_transformation_isometry:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1156
  "orthogonal_transformation f \<longleftrightarrow> f(0::'a::real_inner) = (0::'a) \<and> (\<forall>x y. dist(f x) (f y) = dist x y)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1157
  unfolding orthogonal_transformation
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1158
  apply (auto simp: linear_0 isometry_linear)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1159
   apply (metis (no_types, hide_lams) dist_norm linear_diff)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1160
  by (metis dist_0_norm)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1161
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1162
71044
nipkow
parents: 70136
diff changeset
  1163
text \<open>Can extend an isometry from unit sphere:\<close>
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1164
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1165
lemma  isometry_sphere_extend:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1166
  fixes f:: "'a::real_inner \<Rightarrow> 'a"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1167
  assumes f1: "\<And>x. norm x = 1 \<Longrightarrow> norm (f x) = 1"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1168
    and fd1: "\<And>x y. \<lbrakk>norm x = 1; norm y = 1\<rbrakk> \<Longrightarrow> dist (f x) (f y) = dist x y"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1169
  shows "\<exists>g. orthogonal_transformation g \<and> (\<forall>x. norm x = 1 \<longrightarrow> g x = f x)"
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1170
proof -
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1171
  {
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1172
    fix x y x' y' u v u' v' :: "'a"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1173
    assume H: "x = norm x *\<^sub>R u" "y = norm y *\<^sub>R v"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1174
              "x' = norm x *\<^sub>R u'" "y' = norm y *\<^sub>R v'"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1175
      and J: "norm u = 1" "norm u' = 1" "norm v = 1" "norm v' = 1" "norm(u' - v') = norm(u - v)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1176
    then have *: "u \<bullet> v = u' \<bullet> v' + v' \<bullet> u' - v \<bullet> u "
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1177
      by (simp add: norm_eq norm_eq_1 inner_add inner_diff)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1178
    have "norm (norm x *\<^sub>R u' - norm y *\<^sub>R v') = norm (norm x *\<^sub>R u - norm y *\<^sub>R v)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1179
      using J by (simp add: norm_eq norm_eq_1 inner_diff * field_simps)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1180
    then have "norm(x' - y') = norm(x - y)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1181
      using H by metis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1182
  }
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1183
  note norm_eq = this
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1184
  let ?g = "\<lambda>x. if x = 0 then 0 else norm x *\<^sub>R f (x /\<^sub>R norm x)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1185
  have thfg: "?g x = f x" if "norm x = 1" for x
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1186
    using that by auto
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1187
  have thd: "dist (?g x) (?g y) = dist x y" for x y
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1188
  proof (cases "x=0 \<or> y=0")
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1189
    case False
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1190
    show "dist (?g x) (?g y) = dist x y"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1191
      unfolding dist_norm
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1192
    proof (rule norm_eq)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1193
      show "x = norm x *\<^sub>R (x /\<^sub>R norm x)" "y = norm y *\<^sub>R (y /\<^sub>R norm y)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1194
           "norm (f (x /\<^sub>R norm x)) = 1" "norm (f (y /\<^sub>R norm y)) = 1"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1195
        using False f1 by auto
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1196
    qed (use False in \<open>auto simp: field_simps intro: f1 fd1[unfolded dist_norm]\<close>)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1197
  qed (auto simp: f1)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1198
  show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1199
    unfolding orthogonal_transformation_isometry
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1200
    by (rule exI[where x= ?g]) (metis thfg thd)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1201
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1202
69683
8b3458ca0762 subsection is always %important
immler
parents: 69679
diff changeset
  1203
subsection\<open>Induction on matrix row operations\<close>
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1204
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1205
lemma induct_matrix_row_operations:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1206
  fixes P :: "real^'n^'n \<Rightarrow> bool"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1207
  assumes zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1208
    and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1209
    and swap_cols: "\<And>A m n. \<lbrakk>P A; m \<noteq> n\<rbrakk> \<Longrightarrow> P(\<chi> i j. A $ i $ Fun.swap m n id j)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1210
    and row_op: "\<And>A m n c. \<lbrakk>P A; m \<noteq> n\<rbrakk>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1211
                   \<Longrightarrow> P(\<chi> i. if i = m then row m A + c *\<^sub>R row n A else row i A)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1212
  shows "P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1213
proof -
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1214
  have "P A" if "(\<And>i j. \<lbrakk>j \<in> -K;  i \<noteq> j\<rbrakk> \<Longrightarrow> A$i$j = 0)" for A K
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1215
  proof -
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1216
    have "finite K"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1217
      by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1218
    then show ?thesis using that
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1219
    proof (induction arbitrary: A rule: finite_induct)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1220
      case empty
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1221
      with diagonal show ?case
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1222
        by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1223
    next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1224
      case (insert k K)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1225
      note insertK = insert
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1226
      have "P A" if kk: "A$k$k \<noteq> 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1227
        and 0: "\<And>i j. \<lbrakk>j \<in> - insert k K; i \<noteq> j\<rbrakk> \<Longrightarrow> A$i$j = 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1228
               "\<And>i. \<lbrakk>i \<in> -L; i \<noteq> k\<rbrakk> \<Longrightarrow> A$i$k = 0" for A L
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1229
      proof -
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1230
        have "finite L"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1231
          by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1232
        then show ?thesis using 0 kk
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1233
        proof (induction arbitrary: A rule: finite_induct)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1234
          case (empty B)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1235
          show ?case
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1236
          proof (rule insertK)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1237
            fix i j
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1238
            assume "i \<in> - K" "j \<noteq> i"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1239
            show "B $ j $ i = 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1240
              using \<open>j \<noteq> i\<close> \<open>i \<in> - K\<close> empty
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1241
              by (metis ComplD ComplI Compl_eq_Diff_UNIV Diff_empty UNIV_I insert_iff)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1242
          qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1243
        next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1244
          case (insert l L B)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1245
          show ?case
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1246
          proof (cases "k = l")
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1247
            case True
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1248
            with insert show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1249
              by auto
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1250
          next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1251
            case False
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1252
            let ?C = "\<chi> i. if i = l then row l B - (B $ l $ k / B $ k $ k) *\<^sub>R row k B else row i B"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1253
            have 1: "\<lbrakk>j \<in> - insert k K; i \<noteq> j\<rbrakk> \<Longrightarrow> ?C $ i $ j = 0" for j i
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1254
              by (auto simp: insert.prems(1) row_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1255
            have 2: "?C $ i $ k = 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1256
              if "i \<in> - L" "i \<noteq> k" for i
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1257
            proof (cases "i=l")
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1258
              case True
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1259
              with that insert.prems show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1260
                by (simp add: row_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1261
            next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1262
              case False
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1263
              with that show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1264
                by (simp add: insert.prems(2) row_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1265
            qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1266
            have 3: "?C $ k $ k \<noteq> 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1267
              by (auto simp: insert.prems row_def \<open>k \<noteq> l\<close>)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1268
            have PC: "P ?C"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1269
              using insert.IH [OF 1 2 3] by auto
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1270
            have eqB: "(\<chi> i. if i = l then row l ?C + (B $ l $ k / B $ k $ k) *\<^sub>R row k ?C else row i ?C) = B"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1271
              using \<open>k \<noteq> l\<close> by (simp add: vec_eq_iff row_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1272
            show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1273
              using row_op [OF PC, of l k, where c = "B$l$k / B$k$k"] eqB \<open>k \<noteq> l\<close>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1274
              by (simp add: cong: if_cong)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1275
          qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1276
        qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1277
      qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1278
      then have nonzero_hyp: "P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1279
        if kk: "A$k$k \<noteq> 0" and zeroes: "\<And>i j. j \<in> - insert k K \<and> i\<noteq>j \<Longrightarrow> A$i$j = 0" for A
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1280
        by (auto simp: intro!: kk zeroes)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1281
      show ?case
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1282
      proof (cases "row k A = 0")
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1283
        case True
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1284
        with zero_row show ?thesis by auto
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1285
      next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1286
        case False
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1287
        then obtain l where l: "A$k$l \<noteq> 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1288
          by (auto simp: row_def zero_vec_def vec_eq_iff)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1289
        show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1290
        proof (cases "k = l")
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1291
          case True
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1292
          with l nonzero_hyp insert.prems show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1293
            by blast
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1294
        next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1295
          case False
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1296
          have *: "A $ i $ Fun.swap k l id j = 0" if "j \<noteq> k" "j \<notin> K" "i \<noteq> j" for i j
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1297
            using False l insert.prems that
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1298
            by (auto simp: swap_def insert split: if_split_asm)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1299
          have "P (\<chi> i j. (\<chi> i j. A $ i $ Fun.swap k l id j) $ i $ Fun.swap k l id j)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1300
            by (rule swap_cols [OF nonzero_hyp False]) (auto simp: l *)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1301
          moreover
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1302
          have "(\<chi> i j. (\<chi> i j. A $ i $ Fun.swap k l id j) $ i $ Fun.swap k l id j) = A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1303
            by (vector Fun.swap_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1304
          ultimately show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1305
            by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1306
        qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1307
      qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1308
    qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1309
  qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1310
  then show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1311
    by blast
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1312
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1313
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1314
lemma induct_matrix_elementary:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1315
  fixes P :: "real^'n^'n \<Rightarrow> bool"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1316
  assumes mult: "\<And>A B. \<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P(A ** B)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1317
    and zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1318
    and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1319
    and swap1: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. mat 1 $ i $ Fun.swap m n id j)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1320
    and idplus: "\<And>m n c. m \<noteq> n \<Longrightarrow> P(\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1321
  shows "P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1322
proof -
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1323
  have swap: "P (\<chi> i j. A $ i $ Fun.swap m n id j)"  (is "P ?C")
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1324
    if "P A" "m \<noteq> n" for A m n
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1325
  proof -
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1326
    have "A ** (\<chi> i j. mat 1 $ i $ Fun.swap m n id j) = ?C"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1327
      by (simp add: matrix_matrix_mult_def mat_def vec_eq_iff if_distrib sum.delta_remove)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1328
    then show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1329
      using mult swap1 that by metis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1330
  qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1331
  have row: "P (\<chi> i. if i = m then row m A + c *\<^sub>R row n A else row i A)"  (is "P ?C")
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1332
    if "P A" "m \<noteq> n" for A m n c
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1333
  proof -
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1334
    let ?B = "\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1335
    have "?B ** A = ?C"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1336
      using \<open>m \<noteq> n\<close> unfolding matrix_matrix_mult_def row_def of_bool_def
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1337
      by (auto simp: vec_eq_iff if_distrib [of "\<lambda>x. x * y" for y] sum.remove cong: if_cong)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1338
    then show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1339
      by (rule subst) (auto simp: that mult idplus)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1340
  qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1341
  show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1342
    by (rule induct_matrix_row_operations [OF zero_row diagonal swap row])
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1343
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1344
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1345
lemma induct_matrix_elementary_alt:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1346
  fixes P :: "real^'n^'n \<Rightarrow> bool"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1347
  assumes mult: "\<And>A B. \<lbrakk>P A; P B\<rbrakk> \<Longrightarrow> P(A ** B)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1348
    and zero_row: "\<And>A i. row i A = 0 \<Longrightarrow> P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1349
    and diagonal: "\<And>A. (\<And>i j. i \<noteq> j \<Longrightarrow> A$i$j = 0) \<Longrightarrow> P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1350
    and swap1: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. mat 1 $ i $ Fun.swap m n id j)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1351
    and idplus: "\<And>m n. m \<noteq> n \<Longrightarrow> P(\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1352
  shows "P A"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1353
proof -
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1354
  have *: "P (\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1355
    if "m \<noteq> n" for m n c
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1356
  proof (cases "c = 0")
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1357
    case True
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1358
    with diagonal show ?thesis by auto
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1359
  next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1360
    case False
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1361
    then have eq: "(\<chi> i j. if i = m \<and> j = n then c else of_bool (i = j)) =
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1362
                      (\<chi> i j. if i = j then (if j = n then inverse c else 1) else 0) **
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1363
                      (\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)) **
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1364
                      (\<chi> i j. if i = j then if j = n then c else 1 else 0)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1365
      using \<open>m \<noteq> n\<close>
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1366
      apply (simp add: matrix_matrix_mult_def vec_eq_iff of_bool_def if_distrib [of "\<lambda>x. y * x" for y] cong: if_cong)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1367
      apply (simp add: if_if_eq_conj sum.neutral conj_commute cong: conj_cong)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1368
      done
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1369
    show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1370
      apply (subst eq)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1371
      apply (intro mult idplus that)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1372
       apply (auto intro: diagonal)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1373
      done
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1374
  qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1375
  show ?thesis
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1376
    by (rule induct_matrix_elementary) (auto intro: assms *)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1377
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1378
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1379
lemma matrix_vector_mult_matrix_matrix_mult_compose:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1380
  "(*v) (A ** B) = (*v) A \<circ> (*v) B"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1381
  by (auto simp: matrix_vector_mul_assoc)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1382
69723
9b9f203e0ba3 tagged 2 theories ie Cartesian_Euclidean_Space Cartesian_Space
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69683
diff changeset
  1383
lemma induct_linear_elementary:
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1384
  fixes f :: "real^'n \<Rightarrow> real^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1385
  assumes "linear f"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1386
    and comp: "\<And>f g. \<lbrakk>linear f; linear g; P f; P g\<rbrakk> \<Longrightarrow> P(f \<circ> g)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1387
    and zeroes: "\<And>f i. \<lbrakk>linear f; \<And>x. (f x) $ i = 0\<rbrakk> \<Longrightarrow> P f"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1388
    and const: "\<And>c. P(\<lambda>x. \<chi> i. c i * x$i)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1389
    and swap: "\<And>m n::'n. m \<noteq> n \<Longrightarrow> P(\<lambda>x. \<chi> i. x $ Fun.swap m n id i)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1390
    and idplus: "\<And>m n::'n. m \<noteq> n \<Longrightarrow> P(\<lambda>x. \<chi> i. if i = m then x$m + x$n else x$i)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1391
  shows "P f"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1392
proof -
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1393
  have "P ((*v) A)" for A
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1394
  proof (rule induct_matrix_elementary_alt)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1395
    fix A B
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1396
    assume "P ((*v) A)" and "P ((*v) B)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1397
    then show "P ((*v) (A ** B))"
71044
nipkow
parents: 70136
diff changeset
  1398
      by (auto simp add: matrix_vector_mult_matrix_matrix_mult_compose intro!: comp)
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1399
  next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1400
    fix A :: "real^'n^'n" and i
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1401
    assume "row i A = 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1402
    show "P ((*v) A)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1403
      using matrix_vector_mul_linear
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1404
      by (rule zeroes[where i=i])
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1405
        (metis \<open>row i A = 0\<close> inner_zero_left matrix_vector_mul_component row_def vec_lambda_eta)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1406
  next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1407
    fix A :: "real^'n^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1408
    assume 0: "\<And>i j. i \<noteq> j \<Longrightarrow> A $ i $ j = 0"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1409
    have "A $ i $ i * x $ i = (\<Sum>j\<in>UNIV. A $ i $ j * x $ j)" for x and i :: "'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1410
      by (simp add: 0 comm_monoid_add_class.sum.remove [where x=i])
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1411
    then have "(\<lambda>x. \<chi> i. A $ i $ i * x $ i) = ((*v) A)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1412
      by (auto simp: 0 matrix_vector_mult_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1413
    then show "P ((*v) A)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1414
      using const [of "\<lambda>i. A $ i $ i"] by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1415
  next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1416
    fix m n :: "'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1417
    assume "m \<noteq> n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1418
    have eq: "(\<Sum>j\<in>UNIV. if i = Fun.swap m n id j then x $ j else 0) =
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1419
              (\<Sum>j\<in>UNIV. if j = Fun.swap m n id i then x $ j else 0)"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1420
      for i and x :: "real^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1421
      unfolding swap_def by (rule sum.cong) auto
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1422
    have "(\<lambda>x::real^'n. \<chi> i. x $ Fun.swap m n id i) = ((*v) (\<chi> i j. if i = Fun.swap m n id j then 1 else 0))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1423
      by (auto simp: mat_def matrix_vector_mult_def eq if_distrib [of "\<lambda>x. x * y" for y] cong: if_cong)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1424
    with swap [OF \<open>m \<noteq> n\<close>] show "P ((*v) (\<chi> i j. mat 1 $ i $ Fun.swap m n id j))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1425
      by (simp add: mat_def matrix_vector_mult_def)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1426
  next
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1427
    fix m n :: "'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1428
    assume "m \<noteq> n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1429
    then have "x $ m + x $ n = (\<Sum>j\<in>UNIV. of_bool (j = n \<or> m = j) * x $ j)" for x :: "real^'n"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1430
      by (auto simp: of_bool_def if_distrib [of "\<lambda>x. x * y" for y] sum.remove cong: if_cong)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1431
    then have "(\<lambda>x::real^'n. \<chi> i. if i = m then x $ m + x $ n else x $ i) =
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1432
               ((*v) (\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1433
      unfolding matrix_vector_mult_def of_bool_def
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1434
      by (auto simp: vec_eq_iff if_distrib [of "\<lambda>x. x * y" for y] cong: if_cong)
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1435
    then show "P ((*v) (\<chi> i j. of_bool (i = m \<and> j = n \<or> i = j)))"
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1436
      using idplus [OF \<open>m \<noteq> n\<close>] by simp
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1437
  qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1438
  then show ?thesis
71633
07bec530f02e cleaned proofs
nipkow
parents: 71044
diff changeset
  1439
    by (metis \<open>linear f\<close> matrix_vector_mul(2))
69675
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1440
qed
880ab0f27ddf Reorg, in particular Determinants as well as some linear algebra from Starlike and Change_Of_Vars
immler
parents: 69669
diff changeset
  1441
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
  1442
end