src/HOL/Analysis/Cartesian_Space.thy
author nipkow
Wed, 16 Jan 2019 11:48:06 +0100
changeset 69666 d51e5e41fafe
parent 69665 60110f6d0b4e
child 69667 82bb6225588b
permissions -rw-r--r--
Reorg of material
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
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
    16
definition%unimportant "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
    17
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
    18
lemma%unimportant 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
    19
  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
    20
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
    21
lemma%unimportant 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
    22
  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
    23
  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
    24
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
    25
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
    26
  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
    27
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
    28
lemma%unimportant independent_cart_basis:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    29
  "vec.independent (cart_basis)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    30
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
    31
  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
    32
  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
    33
  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
    34
  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
    35
  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
    36
  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
    37
    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
    38
    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
    39
      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
    40
    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
    41
    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
    42
  qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    43
  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
    44
  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
    45
  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
    46
  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
    47
    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
    48
  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
    49
  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
    50
  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
    51
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    52
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
    53
lemma%unimportant span_cart_basis:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    54
  "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
    55
proof (auto)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    56
  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
    57
  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
    58
  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
    59
    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
    60
    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
    61
     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
    62
     apply clarify
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    63
  proof -
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    64
    fix i::'b
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    65
    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
    66
    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
    67
      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
    68
    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
    69
    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
    70
      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
    71
      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
    72
      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
    73
      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
    74
      proof (rule the_equality)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    75
        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
    76
        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
    77
      qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    78
      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
    79
        apply (subst (2) j)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    80
        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
    81
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    82
    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
    83
  (\<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
    84
    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
    85
      unfolding vector_smult_component ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    86
    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
    87
      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
    88
    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
    89
    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
    90
    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
    91
  qed simp
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    92
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    93
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
    94
(*Some interpretations:*)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
    95
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
    96
  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
    97
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
    98
lemma%unimportant 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
    99
  "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
   100
  by unfold_locales
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   101
    (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
   102
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   103
lemma%important span_vec_eq: "vec.span X = span X"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   104
  and dim_vec_eq: "vec.dim X = dim X"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   105
  and dependent_vec_eq: "vec.dependent X = dependent X"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   106
  and subspace_vec_eq: "vec.subspace X = subspace X"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   107
  for X::"(real^'n) set"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   108
  unfolding span_raw_def dim_raw_def dependent_raw_def subspace_raw_def
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   109
  by (auto simp: scalar_mult_eq_scaleR)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   110
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   111
lemma%important linear_componentwise:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   112
  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
   113
  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
   114
  shows "(f x)$j = sum (\<lambda>i. (x$i) * (f (axis i 1)$j)) (UNIV :: 'm set)" (is "?lhs = ?rhs")
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   115
proof%unimportant -
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   116
  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
   117
    using lf .
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   118
  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
   119
  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
   120
  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
   121
  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
   122
    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
   123
  then show ?thesis
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   124
    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
   125
    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
   126
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   127
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   128
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
   129
  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
   130
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   131
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
   132
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   133
lemma%unimportant matrix_works:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   134
  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
   135
  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
   136
  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
   137
  apply clarify
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   138
  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
   139
  done
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   140
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   141
lemma%unimportant 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
   142
  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
   143
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   144
lemma%unimportant matrix_compose_gen:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   145
  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
   146
    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
   147
  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
   148
  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
   149
  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
   150
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   151
lemma%unimportant matrix_compose:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   152
  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
   153
  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
   154
  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
   155
  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
   156
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   157
lemma%unimportant left_invertible_transpose:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   158
  "(\<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
   159
  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   160
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   161
lemma%unimportant right_invertible_transpose:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   162
  "(\<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
   163
  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   164
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   165
lemma%unimportant 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
   166
  "Vector_Spaces.linear (*s) (*s) f \<longleftrightarrow> linear (f :: real^'n \<Rightarrow> real ^'m)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   167
  by (simp add: scalar_mult_eq_scaleR linear_def)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   168
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   169
lemma%unimportant matrix_vector_mul[simp]:
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   170
  "Vector_Spaces.linear (*s) (*s) g \<Longrightarrow> (\<lambda>y. matrix g *v y) = g"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   171
  "linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   172
  "bounded_linear f \<Longrightarrow> (\<lambda>x. matrix f *v x) = f"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   173
  for f :: "real^'n \<Rightarrow> real ^'m"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   174
  by (simp_all add: ext matrix_works linear_matrix_vector_mul_eq linear_linear)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   175
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   176
lemma%important matrix_left_invertible_injective:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   177
  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
   178
  shows "(\<exists>B. B ** A = mat 1) \<longleftrightarrow> inj ((*v) A)"
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   179
proof%unimportant safe
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   180
  fix B
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   181
  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
   182
  show "inj ((*v) A)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   183
    unfolding inj_on_def
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   184
      by (metis B matrix_vector_mul_assoc matrix_vector_mul_lid)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   185
next
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   186
  assume "inj ((*v) A)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   187
  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
   188
  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
   189
    by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   190
  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
   191
    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
   192
        matrix_eq matrix_id_mat_1 matrix_vector_mul(1))
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   193
  then show "\<exists>B. B ** A = mat 1"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   194
    by metis
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   195
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   196
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   197
lemma%unimportant 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
   198
  "(\<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
   199
  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
   200
  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
   201
  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
   202
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   203
lemma%important matrix_right_invertible_surjective:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   204
  "(\<exists>B. (A::'a::field^'n^'m) ** (B::'a::field^'m^'n) = mat 1) \<longleftrightarrow> surj (\<lambda>x. A *v x)"
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   205
proof%unimportant -
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   206
  { fix B :: "'a ^'m^'n"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   207
    assume AB: "A ** B = mat 1"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   208
    { fix x :: "'a ^ 'm"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   209
      have "A *v (B *v x) = x"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   210
        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
   211
    hence "surj ((*v) A)" unfolding surj_def by metis }
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   212
  moreover
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   213
  { assume sf: "surj ((*v) A)"
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   214
    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
   215
    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
   216
      by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   217
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   218
    have "A ** (matrix g) = mat 1"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   219
      unfolding matrix_eq  matrix_vector_mul_lid
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   220
        matrix_vector_mul_assoc[symmetric] matrix_works[OF g(1)]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   221
      using g(2) unfolding o_def fun_eq_iff id_def
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   222
      .
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   223
    hence "\<exists>B. A ** (B::'a^'m^'n) = mat 1" by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   224
  }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   225
  ultimately show ?thesis unfolding surj_def by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   226
qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   227
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   228
lemma%important 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
   229
  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
   230
  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
   231
      (\<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
   232
    (is "?lhs \<longleftrightarrow> ?rhs")
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   233
proof%unimportant -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   234
  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
   235
  { 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
   236
    { fix c i
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   237
      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
   238
      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
   239
      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
   240
        using c
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   241
        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
   242
      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
   243
      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
   244
    hence ?rhs by blast }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   245
  moreover
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   246
  { assume H: ?rhs
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   247
    { 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
   248
      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
   249
      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
   250
      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
   251
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   252
  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
   253
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   254
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   255
lemma%unimportant 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
   256
  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
   257
  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
   258
    (\<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
   259
  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
   260
    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
   261
  by (simp add:)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   262
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   263
lemma%important matrix_right_invertible_span_columns:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   264
  "(\<exists>(B::'a::field ^'n^'m). (A::'a ^'m^'n) ** B = mat 1) \<longleftrightarrow>
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   265
    vec.span (columns A) = UNIV" (is "?lhs = ?rhs")
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   266
proof%unimportant -
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   267
  let ?U = "UNIV :: 'm set"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   268
  have fU: "finite ?U" by simp
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   269
  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
   270
    unfolding matrix_right_invertible_surjective matrix_mult_sum surj_def
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   271
    by (simp add: eq_commute)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   272
  have rhseq: "?rhs \<longleftrightarrow> (\<forall>x. x \<in> vec.span (columns A))" by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   273
  { assume h: ?lhs
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   274
    { fix x:: "'a ^'n"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   275
      from h[unfolded lhseq, rule_format, of x] obtain y :: "'a ^'m"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   276
        where y: "sum (\<lambda>i. (y$i) *s column i A) ?U = x" by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   277
      have "x \<in> vec.span (columns A)"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   278
        unfolding y[symmetric] scalar_mult_eq_scaleR
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   279
      proof (rule vec.span_sum [OF vec.span_scale])
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   280
        show "column i A \<in> vec.span (columns A)" for i
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   281
          using columns_def vec.span_superset by auto
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   282
      qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   283
    }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   284
    then have ?rhs unfolding rhseq by blast }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   285
  moreover
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   286
  { assume h:?rhs
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   287
    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
   288
    { fix y
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   289
      have "y \<in> vec.span (columns A)"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   290
        unfolding h by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   291
      then have "?P y"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   292
      proof (induction rule: vec.span_induct_alt)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   293
        case base
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   294
        then show ?case
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   295
          by (metis (full_types) matrix_mult_sum matrix_vector_mult_0_right)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   296
      next
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   297
        case (step c y1 y2)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   298
        from step obtain i where i: "i \<in> ?U" "y1 = column i A"
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   299
          unfolding columns_def by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   300
        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
   301
          using step by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   302
        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
   303
        show ?case
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   304
        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
   305
          fix j
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   306
          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
   307
              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
   308
            using i(1) by (simp add: field_simps)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   309
          have "sum (\<lambda>xa. 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))) ?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
   311
            by (rule sum.cong[OF refl]) (use th in blast)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   312
          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
   313
            by (simp add: sum.distrib)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   314
          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
   315
            unfolding sum.delta[OF fU]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   316
            using i(1) by simp
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   317
          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
   318
            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
   319
        qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   320
      qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   321
    }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   322
    then have ?lhs unfolding lhseq ..
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   323
  }
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   324
  ultimately show ?thesis by blast
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   325
qed
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   326
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   327
lemma%unimportant matrix_left_invertible_span_rows_gen:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   328
  "(\<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
   329
  unfolding right_invertible_transpose[symmetric]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   330
  unfolding columns_transpose[symmetric]
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   331
  unfolding matrix_right_invertible_span_columns
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   332
  ..
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   333
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   334
lemma%unimportant matrix_left_invertible_span_rows:
68074
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   335
  "(\<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
   336
  using matrix_left_invertible_span_rows_gen[of A] by (simp add: span_vec_eq)
8d50467f7555 fixed HOL-Analysis
immler
parents: 68073
diff changeset
   337
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   338
lemma%important 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
   339
  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
   340
  shows "A ** A' = mat 1 \<longleftrightarrow> A' ** A = mat 1"
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   341
proof%unimportant -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   342
  { 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
   343
    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
   344
    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
   345
      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
   346
    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
   347
    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
   348
      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
   349
    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
   350
      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
   351
          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
   352
    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
   353
    hence "matrix f' = A'"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   354
      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
   355
    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
   356
    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
   357
  }
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   358
  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
   359
qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   360
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   361
lemma%unimportant invertible_left_inverse:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   362
  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
   363
  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
   364
  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
   365
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   366
lemma%unimportant invertible_right_inverse:
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   367
  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
   368
  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
   369
  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
   370
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   371
lemma%important 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
   372
  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
   373
  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
   374
  shows "invertible (A**B)"
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   375
proof%unimportant -
68073
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   376
  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
   377
    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
   378
  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
   379
    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
   380
  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
   381
  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
   382
    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
   383
      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
   384
    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
   385
    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
   386
    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
   387
    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
   388
    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
   389
    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
   390
    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
   391
    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
   392
    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
   393
    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
   394
    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
   395
  qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   396
qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   397
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   398
lemma%unimportant 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
   399
  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
   400
  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
   401
  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
   402
  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
   403
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   404
lemma%important 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
   405
  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
   406
  shows "(v v* M) v* N = v v* (M ** N)"
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   407
proof%unimportant -
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
  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
   409
  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
   410
  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
   411
    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
   412
qed
fad29d2a17a5 merged; resolved conflicts manually (esp. lemmas that have been moved from Linear_Algebra and Cartesian_Euclidean_Space)
immler
parents: 68072
diff changeset
   413
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   414
lemma%unimportant 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
   415
  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
   416
  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
   417
  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
   418
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   419
lemma%unimportant 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
   420
  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
   421
  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
   422
  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
   423
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   424
(*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
   425
  library.*)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   426
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   427
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
   428
  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
   429
  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
   430
  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
   431
  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
   432
  and BasisB :: "('b set)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   433
  and f :: "('b=>'c)"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   434
68833
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   435
lemma%important vec_dim_card: "vec.dim (UNIV::('a::{field}^'n) set) = CARD ('n)"
fde093888c16 tagged 21 theories in the Analysis library for the manual
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68189
diff changeset
   436
proof%unimportant -
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   437
  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
   438
  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
   439
    unfolding vec.dim_UNIV ..
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   440
  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
   441
    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
   442
      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
   443
      fix i::'n
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   444
      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
   445
      fix x::"'a^'n"
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   446
      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
   447
      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
   448
    qed
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   449
  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
   450
  finally show ?thesis .
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
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   453
interpretation vector_space_over_itself: vector_space "(*) :: 'a::field => 'a => 'a"
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   454
  by unfold_locales (simp_all add: algebra_simps)
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   455
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   456
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
   457
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   458
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
   459
  "(*) :: '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
   460
  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
   461
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   462
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
   463
  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
   464
69666
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   465
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   466
lemma%unimportant dim_subset_UNIV_cart_gen:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   467
  fixes S :: "('a::field^'n) set"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   468
  shows "vec.dim S \<le> CARD('n)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   469
  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
   470
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   471
lemma%unimportant dim_subset_UNIV_cart:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   472
  fixes S :: "(real^'n) set"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   473
  shows "dim S \<le> CARD('n)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   474
  using dim_subset_UNIV_cart_gen[of S] by (simp add: dim_vec_eq)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   475
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   476
text\<open>Two sometimes fruitful ways of looking at matrix-vector multiplication.\<close>
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   477
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   478
lemma%important matrix_mult_dot: "A *v x = (\<chi> i. inner (A$i) x)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   479
  by (simp add: matrix_vector_mult_def inner_vec_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   480
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   481
lemma%unimportant adjoint_matrix: "adjoint(\<lambda>x. (A::real^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   482
  apply (rule adjoint_unique)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   483
  apply (simp add: transpose_def inner_vec_def matrix_vector_mult_def
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   484
    sum_distrib_right sum_distrib_left)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   485
  apply (subst sum.swap)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   486
  apply (simp add:  ac_simps)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   487
  done
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   488
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   489
lemma%important matrix_adjoint: assumes lf: "linear (f :: real^'n \<Rightarrow> real ^'m)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   490
  shows "matrix(adjoint f) = transpose(matrix f)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   491
proof%unimportant -
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   492
  have "matrix(adjoint f) = matrix(adjoint ((*v) (matrix f)))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   493
    by (simp add: lf)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   494
  also have "\<dots> = transpose(matrix f)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   495
    unfolding adjoint_matrix matrix_of_matrix_vector_mul
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   496
    apply rule
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   497
    done
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   498
  finally show ?thesis .
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   499
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   500
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   501
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   502
subsection%important\<open> Rank of a matrix\<close>
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   503
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   504
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
   505
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   506
lemma%unimportant matrix_vector_mult_in_columnspace_gen:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   507
  fixes A :: "'a::field^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   508
  shows "(A *v x) \<in> vec.span(columns A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   509
  apply (simp add: matrix_vector_column columns_def transpose_def column_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   510
  apply (intro vec.span_sum vec.span_scale)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   511
  apply (force intro: vec.span_base)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   512
  done
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   513
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   514
lemma%unimportant matrix_vector_mult_in_columnspace:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   515
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   516
  shows "(A *v x) \<in> span(columns A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   517
  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
   518
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   519
lemma subspace_orthogonal_to_vector: "subspace {y. orthogonal x y}"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   520
  by (simp add: subspace_def orthogonal_clauses)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   521
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   522
lemma%important orthogonal_nullspace_rowspace:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   523
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   524
  assumes 0: "A *v x = 0" and y: "y \<in> span(rows A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   525
  shows "orthogonal x y"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   526
  using y
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   527
proof%unimportant (induction rule: span_induct)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   528
  case base
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   529
  then show ?case
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   530
    by (simp add: subspace_orthogonal_to_vector)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   531
next
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   532
  case (step v)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   533
  then obtain i where "v = row i A"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   534
    by (auto simp: rows_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   535
  with 0 show ?case
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   536
    unfolding orthogonal_def inner_vec_def matrix_vector_mult_def row_def
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   537
    by (simp add: mult.commute) (metis (no_types) vec_lambda_beta zero_index)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   538
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   539
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   540
lemma%unimportant nullspace_inter_rowspace:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   541
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   542
  shows "A *v x = 0 \<and> x \<in> span(rows A) \<longleftrightarrow> x = 0"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   543
  using orthogonal_nullspace_rowspace orthogonal_self span_zero matrix_vector_mult_0_right
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   544
  by blast
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   545
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   546
lemma%unimportant matrix_vector_mul_injective_on_rowspace:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   547
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   548
  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
   549
  using nullspace_inter_rowspace [of A "x-y"]
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   550
  by (metis diff_eq_diff_eq diff_self matrix_vector_mult_diff_distrib span_diff)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   551
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   552
definition%important rank :: "'a::field^'n^'m=>nat"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   553
  where row_rank_def_gen: "rank A \<equiv> vec.dim(rows A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   554
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   555
lemma%important row_rank_def: "rank A = dim (rows A)" for A::"real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   556
  by%unimportant (auto simp: row_rank_def_gen dim_vec_eq)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   557
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   558
lemma%important dim_rows_le_dim_columns:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   559
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   560
  shows "dim(rows A) \<le> dim(columns A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   561
proof%unimportant -
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   562
  have "dim (span (rows A)) \<le> dim (span (columns A))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   563
  proof -
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   564
    obtain B where "independent B" "span(rows A) \<subseteq> span B"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   565
              and B: "B \<subseteq> span(rows A)""card B = dim (span(rows A))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   566
      using basis_exists [of "span(rows A)"] by metis
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   567
    with span_subspace have eq: "span B = span(rows A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   568
      by auto
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   569
    then have inj: "inj_on ((*v) A) (span B)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   570
      by (simp add: inj_on_def matrix_vector_mul_injective_on_rowspace)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   571
    then have ind: "independent ((*v) A ` B)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   572
      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
   573
    have "dim (span (rows A)) \<le> card ((*v) A ` B)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   574
      unfolding B(2)[symmetric]
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   575
      using inj
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   576
      by (auto simp: card_image inj_on_subset span_superset)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   577
    also have "\<dots> \<le> dim (span (columns A))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   578
      using _ ind
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   579
      by (rule independent_card_le_dim) (auto intro!: matrix_vector_mult_in_columnspace)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   580
    finally show ?thesis .
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   581
  qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   582
  then show ?thesis
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   583
    by (simp add: dim_span)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   584
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   585
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   586
lemma%unimportant column_rank_def:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   587
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   588
  shows "rank A = dim(columns A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   589
  unfolding row_rank_def
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   590
  by (metis columns_transpose dim_rows_le_dim_columns le_antisym rows_transpose)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   591
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   592
lemma%unimportant rank_transpose:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   593
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   594
  shows "rank(transpose A) = rank A"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   595
  by (metis column_rank_def row_rank_def rows_transpose)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   596
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   597
lemma%unimportant matrix_vector_mult_basis:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   598
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   599
  shows "A *v (axis k 1) = column k A"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   600
  by (simp add: cart_eq_inner_axis column_def matrix_mult_dot)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   601
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   602
lemma%unimportant columns_image_basis:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   603
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   604
  shows "columns A = (*v) A ` (range (\<lambda>i. axis i 1))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   605
  by (force simp: columns_def matrix_vector_mult_basis [symmetric])
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   606
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   607
lemma%important rank_dim_range:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   608
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   609
  shows "rank A = dim(range (\<lambda>x. A *v x))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   610
  unfolding column_rank_def
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   611
proof%unimportant (rule span_eq_dim)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   612
  have "span (columns A) \<subseteq> span (range ((*v) A))" (is "?l \<subseteq> ?r")
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   613
    by (simp add: columns_image_basis image_subsetI span_mono)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   614
  then show "?l = ?r"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   615
    by (metis (no_types, lifting) image_subset_iff matrix_vector_mult_in_columnspace
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   616
        span_eq span_span)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   617
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   618
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   619
lemma%unimportant rank_bound:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   620
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   621
  shows "rank A \<le> min CARD('m) (CARD('n))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   622
  by (metis (mono_tags, lifting) dim_subset_UNIV_cart min.bounded_iff
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   623
      column_rank_def row_rank_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   624
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   625
lemma%unimportant full_rank_injective:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   626
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   627
  shows "rank A = CARD('n) \<longleftrightarrow> inj ((*v) A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   628
  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
   629
      dim_eq_full [symmetric] card_cart_basis vec.dimension_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   630
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   631
lemma%unimportant full_rank_surjective:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   632
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   633
  shows "rank A = CARD('m) \<longleftrightarrow> surj ((*v) A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   634
  by (simp add: matrix_right_invertible_surjective [symmetric] left_invertible_transpose [symmetric]
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   635
                matrix_left_invertible_injective full_rank_injective [symmetric] rank_transpose)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   636
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   637
lemma%unimportant rank_I: "rank(mat 1::real^'n^'n) = CARD('n)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   638
  by (simp add: full_rank_injective inj_on_def)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   639
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   640
lemma%unimportant less_rank_noninjective:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   641
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   642
  shows "rank A < CARD('n) \<longleftrightarrow> \<not> inj ((*v) A)"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   643
using less_le rank_bound by (auto simp: full_rank_injective [symmetric])
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   644
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   645
lemma%unimportant matrix_nonfull_linear_equations_eq:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   646
  fixes A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   647
  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
   648
  by (meson matrix_left_invertible_injective full_rank_injective matrix_left_invertible_ker)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   649
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   650
lemma%unimportant rank_eq_0: "rank A = 0 \<longleftrightarrow> A = 0" and rank_0 [simp]: "rank (0::real^'n^'m) = 0"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   651
  for A :: "real^'n^'m"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   652
  by (auto simp: rank_dim_range matrix_eq)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   653
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   654
lemma%important rank_mul_le_right:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   655
  fixes A :: "real^'n^'m" and B :: "real^'p^'n"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   656
  shows "rank(A ** B) \<le> rank B"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   657
proof%unimportant -
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   658
  have "rank(A ** B) \<le> dim ((*v) A ` range ((*v) B))"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   659
    by (auto simp: rank_dim_range image_comp o_def matrix_vector_mul_assoc)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   660
  also have "\<dots> \<le> rank B"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   661
    by (simp add: rank_dim_range dim_image_le)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   662
  finally show ?thesis .
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   663
qed
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   664
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   665
lemma%unimportant rank_mul_le_left:
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   666
  fixes A :: "real^'n^'m" and B :: "real^'p^'n"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   667
  shows "rank(A ** B) \<le> rank A"
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   668
  by (metis matrix_transpose_mul rank_mul_le_right rank_transpose)
d51e5e41fafe Reorg of material
nipkow
parents: 69665
diff changeset
   669
68072
493b818e8e10 added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
diff changeset
   670
end