src/HOL/Matrix/Matrix.thy
author obua
Mon, 10 May 2004 17:10:41 +0200
changeset 14724 021264410f87
parent 14691 e1eedc8cad37
child 14738 83f1a514dcb4
permissions -rw-r--r--
preparation for integration with new Ring_and_Field.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     1
(*  Title:      HOL/Matrix/Matrix.thy
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     2
    ID:         $Id$
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     3
    Author:     Steven Obua
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     4
    License:    2004 Technische Universität München
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     5
*)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     6
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
     7
theory Matrix = MatrixGeneral:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     8
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     9
axclass almost_matrix_element < zero, plus, times
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    10
matrix_add_assoc: "(a+b)+c = a + (b+c)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    11
matrix_add_commute: "a+b = b+a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    12
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    13
matrix_mult_assoc: "(a*b)*c = a*(b*c)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    14
matrix_mult_left_0[simp]: "0 * a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    15
matrix_mult_right_0[simp]: "a * 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    16
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    17
matrix_left_distrib: "(a+b)*c = a*c+b*c"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    18
matrix_right_distrib: "a*(b+c) = a*b+a*c"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    19
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    20
axclass matrix_element < almost_matrix_element
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    21
matrix_add_0[simp]: "0+0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    22
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14662
diff changeset
    23
instance matrix :: (plus) plus ..
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14662
diff changeset
    24
instance matrix :: (times) times ..
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    25
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    26
defs (overloaded)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    27
plus_matrix_def: "A + B == combine_matrix (op +) A B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    28
times_matrix_def: "A * B == mult_matrix (op *) (op +) A B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    29
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    30
instance matrix :: (matrix_element) matrix_element
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    31
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    32
  note combine_matrix_assoc2 = combine_matrix_assoc[simplified associative_def, THEN spec, THEN spec, THEN spec]
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    33
  {
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    34
    fix A::"('a::matrix_element) matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    35
    fix B
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    36
    fix C
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    37
    have "(A + B) + C = A + (B + C)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    38
      apply (simp add: plus_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    39
      apply (rule combine_matrix_assoc2)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    40
      by (simp_all add: matrix_add_assoc)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    41
  }
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    42
  note plus_assoc = this
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    43
  {
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    44
    fix A::"('a::matrix_element) matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    45
    fix B
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    46
    fix C
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    47
    have "(A * B) * C = A * (B * C)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    48
      apply (simp add: times_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    49
      apply (rule mult_matrix_assoc_simple)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    50
      apply (simp_all add: associative_def commutative_def distributive_def l_distributive_def r_distributive_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    51
      apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    52
      apply (simp_all add: matrix_add_assoc)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    53
      apply (simp_all add: matrix_add_commute)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    54
      apply (simp_all add: matrix_mult_assoc)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    55
      by (simp_all add: matrix_left_distrib matrix_right_distrib)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    56
  }
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    57
  note mult_assoc = this
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    58
  note combine_matrix_commute2 = combine_matrix_commute[simplified commutative_def, THEN spec, THEN spec]
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    59
  {
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    60
    fix A::"('a::matrix_element) matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    61
    fix B
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    62
    have "A + B = B + A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    63
      apply (simp add: plus_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    64
      apply (insert combine_matrix_commute2[of "op +"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    65
      apply (rule combine_matrix_commute2)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    66
      by (simp add: matrix_add_commute)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    67
  }
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    68
  note plus_commute = this
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    69
  have plus_zero: "(0::('a::matrix_element) matrix) + 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    70
    apply (simp add: plus_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    71
    apply (rule combine_matrix_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    72
    by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    73
  have mult_left_zero: "!! A. (0::('a::matrix_element) matrix) * A = 0" by(simp add: times_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    74
  have mult_right_zero: "!! A. A * (0::('a::matrix_element) matrix) = 0" by (simp add: times_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    75
  note l_distributive_matrix2 = l_distributive_matrix[simplified l_distributive_def matrix_left_distrib, THEN spec, THEN spec, THEN spec]
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    76
  {
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
    77
    fix A::"('a::matrix_element) matrix"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
    78
    fix B
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    79
    fix C
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    80
    have "(A + B) * C = A * C + B * C"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    81
      apply (simp add: plus_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    82
      apply (simp add: times_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    83
      apply (rule l_distributive_matrix2)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    84
      apply (simp_all add: associative_def commutative_def l_distributive_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    85
      apply (auto)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
    86
      apply (simp_all add: matrix_add_assoc)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    87
      apply (simp_all add: matrix_add_commute)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    88
      by (simp_all add: matrix_left_distrib)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    89
  }
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    90
  note left_distrib = this
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    91
  note r_distributive_matrix2 = r_distributive_matrix[simplified r_distributive_def matrix_right_distrib, THEN spec, THEN spec, THEN spec]
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    92
  {
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
    93
    fix A::"('a::matrix_element) matrix"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
    94
    fix B
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    95
    fix C
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    96
    have "C * (A + B) = C * A + C * B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    97
      apply (simp add: plus_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    98
      apply (simp add: times_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    99
      apply (rule r_distributive_matrix2)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   100
      apply (simp_all add: associative_def commutative_def r_distributive_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   101
      apply (auto)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
   102
      apply (simp_all add: matrix_add_assoc)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   103
      apply (simp_all add: matrix_add_commute)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   104
      by (simp_all add: matrix_right_distrib)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   105
  }
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   106
  note right_distrib = this
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   107
  show "OFCLASS('a matrix, matrix_element_class)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   108
    apply (intro_classes)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   109
    apply (simp_all add: plus_assoc)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   110
    apply (simp_all add: plus_commute)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   111
    apply (simp_all add: plus_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   112
    apply (simp_all add: mult_assoc)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   113
    apply (simp_all add: mult_left_zero mult_right_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   114
    by (simp_all add: left_distrib right_distrib)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   115
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   116
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   117
axclass g_almost_semiring < almost_matrix_element
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   118
g_add_left_0[simp]: "0 + a = a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   119
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   120
lemma g_add_right_0[simp]: "(a::'a::g_almost_semiring) + 0 = a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   121
by (simp add: matrix_add_commute)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   122
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   123
axclass g_semiring < g_almost_semiring
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   124
g_add_leftimp_eq: "a+b = a+c \<Longrightarrow> b = c"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   125
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   126
instance g_almost_semiring < matrix_element
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14662
diff changeset
   127
  by intro_classes simp
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   128
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   129
instance matrix :: (g_almost_semiring) g_almost_semiring
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   130
apply (intro_classes)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   131
by (simp add: plus_matrix_def combine_matrix_def combine_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   132
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   133
lemma RepAbs_matrix_eq_left: " Rep_matrix(Abs_matrix f) = g \<Longrightarrow> \<exists>m. \<forall>j i. m \<le> j \<longrightarrow> f j i = 0 \<Longrightarrow> \<exists>n. \<forall>j i. n \<le> i \<longrightarrow> f j i = 0 \<Longrightarrow> f = g"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   134
by (simp add: RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   135
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   136
lemma RepAbs_matrix_eq_right: "g = Rep_matrix(Abs_matrix f) \<Longrightarrow> \<exists>m. \<forall>j i. m \<le> j \<longrightarrow> f j i = 0 \<Longrightarrow> \<exists>n. \<forall>j i. n \<le> i \<longrightarrow> f j i = 0 \<Longrightarrow> g = f"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   137
by (simp add: RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   138
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   139
instance matrix :: (g_semiring) g_semiring
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   140
apply (intro_classes)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   141
apply (simp add: plus_matrix_def combine_matrix_def combine_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   142
apply (subst Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   143
apply (drule ssubst[OF Rep_matrix_inject, of "% x. x"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   144
apply (drule RepAbs_matrix_eq_left)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   145
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   146
apply (rule_tac x="max (nrows a) (nrows b)" in exI, simp add: nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   147
apply (rule_tac x="max (ncols a) (ncols b)" in exI, simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   148
apply (drule RepAbs_matrix_eq_right)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   149
apply (rule_tac x="max (nrows a) (nrows c)" in exI, simp add: nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   150
apply (rule_tac x="max (ncols a) (ncols c)" in exI, simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   151
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   152
apply (drule_tac x="x" and y="x" in comb, simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   153
apply (drule_tac x="xa" and y="xa" in comb, simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   154
apply (drule g_add_leftimp_eq)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   155
by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   156
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   157
axclass pordered_matrix_element < matrix_element, order, zero
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   158
pordered_add_right_mono: "a <= b \<Longrightarrow> a + c <= b + c"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   159
pordered_mult_left: "0 <= c \<Longrightarrow> a <= b \<Longrightarrow> c*a <= c*b"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   160
pordered_mult_right: "0 <= c \<Longrightarrow> a <= b \<Longrightarrow> a*c <= b*c"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   161
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   162
lemma pordered_add_left_mono: "(a::'a::pordered_matrix_element) <= b \<Longrightarrow> c + a <= c + b"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   163
apply (insert pordered_add_right_mono[of a b c])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   164
by (simp add: matrix_add_commute)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   165
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   166
lemma pordered_add: "(a::'a::pordered_matrix_element) <= b \<Longrightarrow> (c::'a::pordered_matrix_element) <= d \<Longrightarrow> a+c <= b+d"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   167
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   168
  assume p1:"a <= b"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   169
  assume p2:"c <= d"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
   170
  have "a+c <= b+c" by (rule pordered_add_right_mono)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   171
  also have "\<dots> <= b+d" by (rule pordered_add_left_mono)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   172
  ultimately show "a+c <= b+d" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   173
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   174
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
   175
instance matrix :: (pordered_matrix_element) pordered_matrix_element
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   176
apply (intro_classes)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   177
apply (simp_all add: plus_matrix_def times_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   178
apply (rule le_combine_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   179
apply (simp_all)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   180
apply (simp_all add: pordered_add)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   181
apply (rule le_left_mult)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   182
apply (simp_all add: matrix_add_0 g_add_left_0 pordered_add pordered_mult_left matrix_mult_left_0 matrix_mult_right_0)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   183
apply (rule le_right_mult)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   184
by (simp_all add: pordered_add pordered_mult_right)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   185
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   186
axclass pordered_g_semiring < g_semiring, pordered_matrix_element
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   187
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14662
diff changeset
   188
instance matrix :: (pordered_g_semiring) pordered_g_semiring ..
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   189
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   190
lemma nrows_mult: "nrows ((A::('a::matrix_element) matrix) * B) <= nrows A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   191
by (simp add: times_matrix_def mult_nrows)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   192
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   193
lemma ncols_mult: "ncols ((A::('a::matrix_element) matrix) * B) <= ncols B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   194
by (simp add: times_matrix_def mult_ncols)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   195
14724
021264410f87 preparation for integration with new Ring_and_Field.thy
obua
parents: 14691
diff changeset
   196
(*
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   197
constdefs
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   198
  one_matrix :: "nat \<Rightarrow> ('a::semiring) matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   199
  "one_matrix n == Abs_matrix (% j i. if j = i & j < n then 1 else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   200
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   201
lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i & j < n) then 1 else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   202
apply (simp add: one_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   203
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   204
apply (rule exI[of _ n], simp add: split_if)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   205
by (simp add: split_if, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   206
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   207
lemma nrows_one_matrix[simp]: "nrows (one_matrix n) = n" (is "?r = _")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   208
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   209
  have "?r <= n" by (simp add: nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   210
  moreover have "n <= ?r" by (simp add: le_nrows, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   211
  ultimately show "?r = n" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   212
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   213
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   214
lemma ncols_one_matrix[simp]: "ncols (one_matrix n) = n" (is "?r = _")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   215
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   216
  have "?r <= n" by (simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   217
  moreover have "n <= ?r" by (simp add: le_ncols, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   218
  ultimately show "?r = n" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   219
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   220
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   221
lemma one_matrix_mult_right: "ncols A <= n \<Longrightarrow> A * (one_matrix n) = A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   222
apply (subst Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   223
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   224
apply (simp add: times_matrix_def Rep_mult_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   225
apply (rule_tac j1="xa" in ssubst[OF foldseq_almostzero])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   226
apply (simp_all)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   227
by (simp add: max_def ncols)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   228
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   229
lemma one_matrix_mult_left: "nrows A <= n \<Longrightarrow> (one_matrix n) * A = A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   230
apply (subst Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   231
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   232
apply (simp add: times_matrix_def Rep_mult_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   233
apply (rule_tac j1="x" in ssubst[OF foldseq_almostzero])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   234
apply (simp_all)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   235
by (simp add: max_def nrows)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   236
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
   237
constdefs
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   238
  right_inverse_matrix :: "('a::semiring) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
   239
  "right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X)))"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   240
  inverse_matrix :: "('a::semiring) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   241
  "inverse_matrix A X == (right_inverse_matrix A X) \<and> (right_inverse_matrix X A)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   242
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   243
lemma right_inverse_matrix_dim: "right_inverse_matrix A X \<Longrightarrow> nrows A = ncols X"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   244
apply (insert ncols_mult[of A X], insert nrows_mult[of A X])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   245
by (simp add: right_inverse_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   246
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14593
diff changeset
   247
text {* to be continued \dots *}
14724
021264410f87 preparation for integration with new Ring_and_Field.thy
obua
parents: 14691
diff changeset
   248
*)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   249
end