src/HOL/Matrix/SparseMatrix.thy
author obua
Fri, 03 Sep 2004 17:10:36 +0200
changeset 15178 5f621aa35c25
parent 15009 8c89f588c7aa
child 15197 19e735596e51
permissions -rw-r--r--
Matrix theory, linear programming
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     1
theory SparseMatrix = Matrix:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     2
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     3
types 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     4
  'a spvec = "(nat * 'a) list"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     5
  'a spmat = "('a spvec) spvec"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     6
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     7
consts
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     8
  sparse_row_vector :: "('a::lordered_ring) spvec \<Rightarrow> 'a matrix"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     9
  sparse_row_matrix :: "('a::lordered_ring) spmat \<Rightarrow> 'a matrix"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    10
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    11
defs
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    12
  sparse_row_vector_def : "sparse_row_vector arr == foldl (% m x. m + (singleton_matrix 0 (fst x) (snd x))) 0 arr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    13
  sparse_row_matrix_def : "sparse_row_matrix arr == foldl (% m r. m + (move_matrix (sparse_row_vector (snd r)) (int (fst r)) 0)) 0 arr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    14
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    15
lemma sparse_row_vector_empty[simp]: "sparse_row_vector [] = 0"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    16
  by (simp add: sparse_row_vector_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    17
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    18
lemma sparse_row_matrix_empty[simp]: "sparse_row_matrix [] = 0"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    19
  by (simp add: sparse_row_matrix_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    20
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    21
lemma foldl_distrstart[rule_format]: "! a x y. (f (g x y) a = g x (f y a)) \<Longrightarrow> ! x y. (foldl f (g x y) l = g x (foldl f y l))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    22
  by (induct l, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    23
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    24
lemma sparse_row_vector_cons[simp]: "sparse_row_vector (a#arr) = (singleton_matrix 0 (fst a) (snd a)) + (sparse_row_vector arr)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    25
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    26
  apply (auto simp add: sparse_row_vector_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    27
  apply (simp add: foldl_distrstart[of "\<lambda>m x. m + singleton_matrix 0 (fst x) (snd x)" "\<lambda>x m. singleton_matrix 0 (fst x) (snd x) + m"])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    28
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    29
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    30
lemma sparse_row_vector_append[simp]: "sparse_row_vector (a @ b) = (sparse_row_vector a) + (sparse_row_vector b)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    31
  by (induct a, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    32
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    33
lemma nrows_spvec[simp]: "nrows (sparse_row_vector x) <= (Suc 0)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    34
  apply (induct x)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    35
  apply (simp_all add: add_nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    36
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    37
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    38
lemma sparse_row_matrix_cons: "sparse_row_matrix (a#arr) = ((move_matrix (sparse_row_vector (snd a)) (int (fst a)) 0)) + sparse_row_matrix arr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    39
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    40
  apply (auto simp add: sparse_row_matrix_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    41
  apply (simp add: foldl_distrstart[of "\<lambda>m x. m + (move_matrix (sparse_row_vector (snd x)) (int (fst x)) 0)" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    42
    "% a m. (move_matrix (sparse_row_vector (snd a)) (int (fst a)) 0) + m"])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    43
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    44
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    45
lemma sparse_row_matrix_append: "sparse_row_matrix (arr@brr) = (sparse_row_matrix arr) + (sparse_row_matrix brr)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    46
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    47
  apply (auto simp add: sparse_row_matrix_cons)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    48
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    49
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    50
consts
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    51
  sorted_spvec :: "'a spvec \<Rightarrow> bool"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    52
  sorted_spmat :: "'a spmat \<Rightarrow> bool"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    53
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    54
primrec
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    55
  "sorted_spmat [] = True"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    56
  "sorted_spmat (a#as) = ((sorted_spvec (snd a)) & (sorted_spmat as))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    57
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    58
primrec
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    59
  "sorted_spvec [] = True"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    60
sorted_spvec_step:  "sorted_spvec (a#as) = (case as of [] \<Rightarrow> True | b#bs \<Rightarrow> ((fst a < fst b) & (sorted_spvec as)))" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    61
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    62
declare sorted_spvec.simps [simp del]
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    63
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    64
lemma sorted_spvec_empty[simp]: "sorted_spvec [] = True"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    65
by (simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    66
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    67
lemma sorted_spvec_cons1: "sorted_spvec (a#as) \<Longrightarrow> sorted_spvec as"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    68
apply (induct as)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    69
apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    70
done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    71
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    72
lemma sorted_spvec_cons2: "sorted_spvec (a#b#t) \<Longrightarrow> sorted_spvec (a#t)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    73
apply (induct t)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    74
apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    75
done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    76
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    77
lemma sorted_spvec_cons3: "sorted_spvec(a#b#t) \<Longrightarrow> fst a < fst b"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    78
apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    79
done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    80
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    81
lemma sorted_sparse_row_vector_zero[rule_format]: "m <= n \<longrightarrow> sorted_spvec ((n,a)#arr) \<longrightarrow> Rep_matrix (sparse_row_vector arr) j m = 0"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    82
apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    83
apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    84
apply (frule sorted_spvec_cons2,simp)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    85
apply (frule sorted_spvec_cons3, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    86
done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    87
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    88
lemma sorted_sparse_row_matrix_zero[rule_format]: "m <= n \<longrightarrow> sorted_spvec ((n,a)#arr) \<longrightarrow> Rep_matrix (sparse_row_matrix arr) m j = 0"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    89
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    90
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    91
  apply (frule sorted_spvec_cons2, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    92
  apply (frule sorted_spvec_cons3, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    93
  apply (simp add: sparse_row_matrix_cons neg_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    94
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    95
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    96
consts
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
    97
  abs_spvec :: "('a::lordered_ring) spvec \<Rightarrow> 'a spvec"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
    98
  minus_spvec ::  "('a::lordered_ring) spvec \<Rightarrow> 'a spvec"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    99
  smult_spvec :: "('a::lordered_ring) \<Rightarrow> 'a spvec \<Rightarrow> 'a spvec" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   100
  addmult_spvec :: "('a::lordered_ring) * 'a spvec * 'a spvec \<Rightarrow> 'a spvec"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   101
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   102
primrec 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   103
  "minus_spvec [] = []"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   104
  "minus_spvec (a#as) = (fst a, -(snd a))#(minus_spvec as)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   105
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   106
primrec 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   107
  "abs_spvec [] = []"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   108
  "abs_spvec (a#as) = (fst a, abs (snd a))#(abs_spvec as)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   109
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   110
lemma sparse_row_vector_minus: 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   111
  "sparse_row_vector (minus_spvec v) = - (sparse_row_vector v)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   112
  apply (induct v)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   113
  apply (simp_all add: sparse_row_vector_cons)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   114
  apply (simp add: Rep_matrix_inject[symmetric])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   115
  apply (rule ext)+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   116
  apply simp
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   117
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   118
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   119
lemma sparse_row_vector_abs:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   120
  "sorted_spvec v \<Longrightarrow> sparse_row_vector (abs_spvec v) = abs (sparse_row_vector v)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   121
  apply (induct v)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   122
  apply (simp_all add: sparse_row_vector_cons)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   123
  apply (frule_tac sorted_spvec_cons1, simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   124
  apply (simp only: Rep_matrix_inject[symmetric])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   125
  apply (rule ext)+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   126
  apply auto
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   127
  apply (subgoal_tac "Rep_matrix (sparse_row_vector list) 0 a = 0")
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   128
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   129
  apply (rule sorted_sparse_row_vector_zero)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   130
  apply auto
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   131
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   132
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   133
lemma sorted_spvec_minus_spvec:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   134
  "sorted_spvec v \<Longrightarrow> sorted_spvec (minus_spvec v)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   135
  apply (induct v)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   136
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   137
  apply (frule sorted_spvec_cons1, simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   138
  apply (simp add: sorted_spvec.simps)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   139
  apply (case_tac list)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   140
  apply (simp_all)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   141
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   142
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   143
lemma sorted_spvec_abs_spvec:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   144
  "sorted_spvec v \<Longrightarrow> sorted_spvec (abs_spvec v)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   145
  apply (induct v)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   146
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   147
  apply (frule sorted_spvec_cons1, simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   148
  apply (simp add: sorted_spvec.simps)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   149
  apply (case_tac list)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   150
  apply (simp_all)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   151
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   152
  
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   153
defs
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   154
  smult_spvec_def: "smult_spvec y arr == map (% a. (fst a, y * snd a)) arr"  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   155
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   156
lemma smult_spvec_empty[simp]: "smult_spvec y [] = []"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   157
  by (simp add: smult_spvec_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   158
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   159
lemma smult_spvec_cons: "smult_spvec y (a#arr) = (fst a, y * (snd a)) # (smult_spvec y arr)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   160
  by (simp add: smult_spvec_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   161
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   162
recdef addmult_spvec "measure (% (y, a, b). length a + (length b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   163
  "addmult_spvec (y, arr, []) = arr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   164
  "addmult_spvec (y, [], brr) = smult_spvec y brr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   165
  "addmult_spvec (y, a#arr, b#brr) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   166
    if (fst a) < (fst b) then (a#(addmult_spvec (y, arr, b#brr))) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   167
    else (if (fst b < fst a) then ((fst b, y * (snd b))#(addmult_spvec (y, a#arr, brr)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   168
    else ((fst a, (snd a)+ y*(snd b))#(addmult_spvec (y, arr,brr)))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   169
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   170
lemma addmult_spvec_empty1[simp]: "addmult_spvec (y, [], a) = smult_spvec y a"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   171
  by (induct a, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   172
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   173
lemma addmult_spvec_empty2[simp]: "addmult_spvec (y, a, []) = a"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   174
  by (induct a, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   175
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   176
lemma sparse_row_vector_map: "(! x y. f (x+y) = (f x) + (f y)) \<Longrightarrow> (f::'a\<Rightarrow>('a::lordered_ring)) 0 = 0 \<Longrightarrow> 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   177
  sparse_row_vector (map (% x. (fst x, f (snd x))) a) = apply_matrix f (sparse_row_vector a)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   178
  apply (induct a)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   179
  apply (simp_all add: apply_matrix_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   180
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   181
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   182
lemma sparse_row_vector_smult: "sparse_row_vector (smult_spvec y a) = scalar_mult y (sparse_row_vector a)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   183
  apply (induct a)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   184
  apply (simp_all add: smult_spvec_cons scalar_mult_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   185
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   186
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   187
lemma sparse_row_vector_addmult_spvec: "sparse_row_vector (addmult_spvec (y::'a::lordered_ring, a, b)) = 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   188
  (sparse_row_vector a) + (scalar_mult y (sparse_row_vector b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   189
  apply (rule addmult_spvec.induct[of _ y])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   190
  apply (simp add: scalar_mult_add smult_spvec_cons sparse_row_vector_smult singleton_matrix_add)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   191
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   192
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   193
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   194
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   195
lemma sorted_smult_spvec[rule_format]: "sorted_spvec a \<Longrightarrow> sorted_spvec (smult_spvec y a)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   196
  apply (auto simp add: smult_spvec_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   197
  apply (induct a)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   198
  apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   199
  apply (case_tac list)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   200
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   201
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   202
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   203
lemma sorted_spvec_addmult_spvec_helper: "\<lbrakk>sorted_spvec (addmult_spvec (y, (a, b) # arr, brr)); aa < a; sorted_spvec ((a, b) # arr); 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   204
  sorted_spvec ((aa, ba) # brr)\<rbrakk> \<Longrightarrow> sorted_spvec ((aa, y * ba) # addmult_spvec (y, (a, b) # arr, brr))"  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   205
  apply (induct brr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   206
  apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   207
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   208
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   209
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   210
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   211
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   212
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   213
lemma sorted_spvec_addmult_spvec_helper2: 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   214
 "\<lbrakk>sorted_spvec (addmult_spvec (y, arr, (aa, ba) # brr)); a < aa; sorted_spvec ((a, b) # arr); sorted_spvec ((aa, ba) # brr)\<rbrakk>
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   215
       \<Longrightarrow> sorted_spvec ((a, b) # addmult_spvec (y, arr, (aa, ba) # brr))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   216
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   217
  apply (auto simp add: smult_spvec_def sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   218
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   219
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   220
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   221
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   222
lemma sorted_spvec_addmult_spvec_helper3[rule_format]:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   223
  "sorted_spvec (addmult_spvec (y, arr, brr)) \<longrightarrow> sorted_spvec ((aa, b) # arr) \<longrightarrow> sorted_spvec ((aa, ba) # brr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   224
     \<longrightarrow> sorted_spvec ((aa, b + y * ba) # (addmult_spvec (y, arr, brr)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   225
  apply (rule addmult_spvec.induct[of _ y arr brr])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   226
  apply (simp_all add: sorted_spvec.simps smult_spvec_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   227
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   228
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   229
lemma sorted_addmult_spvec[rule_format]: "sorted_spvec a \<longrightarrow> sorted_spvec b \<longrightarrow> sorted_spvec (addmult_spvec (y, a, b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   230
  apply (rule addmult_spvec.induct[of _ y a b])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   231
  apply (simp_all add: sorted_smult_spvec)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   232
  apply (rule conjI, intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   233
  apply (case_tac "~(a < aa)")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   234
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   235
  apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   236
  apply (simp add: sorted_spvec_addmult_spvec_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   237
  apply (intro strip | rule conjI)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   238
  apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   239
  apply (simp add: sorted_spvec_addmult_spvec_helper2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   240
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   241
  apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   242
  apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   243
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   244
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   245
  apply (simp_all add: sorted_spvec_addmult_spvec_helper3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   246
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   247
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   248
consts 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   249
  mult_spvec_spmat :: "('a::lordered_ring) spvec * 'a spvec * 'a spmat  \<Rightarrow> 'a spvec"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   250
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   251
recdef mult_spvec_spmat "measure (% (c, arr, brr). (length arr) + (length brr))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   252
  "mult_spvec_spmat (c, [], brr) = c"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   253
  "mult_spvec_spmat (c, arr, []) = c"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   254
  "mult_spvec_spmat (c, a#arr, b#brr) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   255
     if ((fst a) < (fst b)) then (mult_spvec_spmat (c, arr, b#brr))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   256
     else (if ((fst b) < (fst a)) then (mult_spvec_spmat (c, a#arr, brr)) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   257
     else (mult_spvec_spmat (addmult_spvec (snd a, c, snd b), arr, brr))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   258
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   259
lemma sparse_row_mult_spvec_spmat[rule_format]: "sorted_spvec (a::('a::lordered_ring) spvec) \<longrightarrow> sorted_spvec B \<longrightarrow> 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   260
  sparse_row_vector (mult_spvec_spmat (c, a, B)) = (sparse_row_vector c) + (sparse_row_vector a) * (sparse_row_matrix B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   261
proof -
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   262
  have comp_1: "!! a b. a < b \<Longrightarrow> Suc 0 <= nat ((int b)-(int a))" by arith
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   263
  have not_iff: "!! a b. a = b \<Longrightarrow> (~ a) = (~ b)" by simp
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   264
  have max_helper: "!! a b. ~ (a <= max (Suc a) b) \<Longrightarrow> False"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   265
    by arith
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   266
  {
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   267
    fix a 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   268
    fix v
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   269
    assume a:"a < nrows(sparse_row_vector v)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   270
    have b:"nrows(sparse_row_vector v) <= 1" by simp
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   271
    note dummy = less_le_trans[of a "nrows (sparse_row_vector v)" 1, OF a b]   
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   272
    then have "a = 0" by simp
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   273
  }
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   274
  note nrows_helper = this
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   275
  show ?thesis
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   276
    apply (rule mult_spvec_spmat.induct)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   277
    apply simp+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   278
    apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   279
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   280
    apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   281
    apply (simp add: ring_eq_simps sparse_row_matrix_cons)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   282
    apply (subst Rep_matrix_zero_imp_mult_zero) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   283
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   284
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   285
    apply (rule disjI2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   286
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   287
    apply (subst nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   288
    apply (rule  order_trans[of _ 1])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   289
    apply (simp add: comp_1)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   290
    apply (subst Rep_matrix_zero_imp_mult_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   291
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   292
    apply (case_tac "k <= aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   293
    apply (rule_tac m1 = k and n1 = a and a1 = b in ssubst[OF sorted_sparse_row_vector_zero])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   294
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   295
    apply (rule impI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   296
    apply (rule disjI2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   297
    apply (rule nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   298
    apply (rule order_trans[of _ 1])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   299
    apply (simp_all add: comp_1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   300
    
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   301
    apply (intro strip | rule conjI)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   302
    apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   303
    apply (simp add: ring_eq_simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   304
    apply (subst Rep_matrix_zero_imp_mult_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   305
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   306
    apply (rule disjI2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   307
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   308
    apply (simp add: sparse_row_matrix_cons neg_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   309
    apply (case_tac "a <= aa")  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   310
    apply (erule sorted_sparse_row_matrix_zero)  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   311
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   312
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   313
    apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   314
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   315
    apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   316
    apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   317
    apply (simp add: sparse_row_matrix_cons ring_eq_simps sparse_row_vector_addmult_spvec)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   318
    apply (rule_tac B1 = "sparse_row_matrix brr" in ssubst[OF Rep_matrix_zero_imp_mult_zero])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   319
    apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   320
    apply (rule sorted_sparse_row_matrix_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   321
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   322
    apply (rule_tac A1 = "sparse_row_vector arr" in ssubst[OF Rep_matrix_zero_imp_mult_zero])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   323
    apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   324
    apply (rule_tac m=k and n = aa and a = b and arr=arr in sorted_sparse_row_vector_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   325
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   326
    apply (simp add: neg_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   327
    apply (drule nrows_notzero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   328
    apply (drule nrows_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   329
    apply (arith)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   330
    
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   331
    apply (subst Rep_matrix_inject[symmetric])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   332
    apply (rule ext)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   333
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   334
    apply (subst Rep_matrix_mult)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   335
    apply (rule_tac j1=aa in ssubst[OF foldseq_almostzero])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   336
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   337
    apply (intro strip, rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   338
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   339
    apply (drule_tac max_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   340
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   341
    apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   342
    apply (rule zero_imp_mult_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   343
    apply (rule disjI2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   344
    apply (rule nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   345
    apply (rule order_trans[of _ 1])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   346
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   347
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   348
    done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   349
qed
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   350
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   351
lemma sorted_mult_spvec_spmat[rule_format]: 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   352
  "sorted_spvec (c::('a::lordered_ring) spvec) \<longrightarrow> sorted_spmat B \<longrightarrow> sorted_spvec (mult_spvec_spmat (c, a, B))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   353
  apply (rule mult_spvec_spmat.induct[of _ c a B])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   354
  apply (simp_all add: sorted_addmult_spvec)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   355
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   356
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   357
consts 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   358
  mult_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   359
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   360
primrec 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   361
  "mult_spmat [] A = []"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   362
  "mult_spmat (a#as) A = (fst a, mult_spvec_spmat ([], snd a, A))#(mult_spmat as A)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   363
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   364
lemma sparse_row_mult_spmat[rule_format]: 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   365
  "sorted_spmat A \<longrightarrow> sorted_spvec B \<longrightarrow> sparse_row_matrix (mult_spmat A B) = (sparse_row_matrix A) * (sparse_row_matrix B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   366
  apply (induct A)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   367
  apply (auto simp add: sparse_row_matrix_cons sparse_row_mult_spvec_spmat ring_eq_simps move_matrix_mult)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   368
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   369
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   370
lemma sorted_spvec_mult_spmat[rule_format]:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   371
  "sorted_spvec (A::('a::lordered_ring) spmat) \<longrightarrow> sorted_spvec (mult_spmat A B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   372
  apply (induct A)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   373
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   374
  apply (drule sorted_spvec_cons1, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   375
  apply (case_tac list)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   376
  apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   377
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   378
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   379
lemma sorted_spmat_mult_spmat[rule_format]:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   380
  "sorted_spmat (B::('a::lordered_ring) spmat) \<longrightarrow> sorted_spmat (mult_spmat A B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   381
  apply (induct A)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   382
  apply (auto simp add: sorted_mult_spvec_spmat) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   383
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   384
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   385
consts
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   386
  add_spvec :: "('a::lordered_ab_group) spvec * 'a spvec \<Rightarrow> 'a spvec"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   387
  add_spmat :: "('a::lordered_ab_group) spmat * 'a spmat \<Rightarrow> 'a spmat"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   388
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   389
recdef add_spvec "measure (% (a, b). length a + (length b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   390
  "add_spvec (arr, []) = arr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   391
  "add_spvec ([], brr) = brr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   392
  "add_spvec (a#arr, b#brr) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   393
  if (fst a) < (fst b) then (a#(add_spvec (arr, b#brr))) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   394
     else (if (fst b < fst a) then (b#(add_spvec (a#arr, brr)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   395
     else ((fst a, (snd a)+(snd b))#(add_spvec (arr,brr)))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   396
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   397
lemma add_spvec_empty1[simp]: "add_spvec ([], a) = a"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   398
  by (induct a, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   399
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   400
lemma add_spvec_empty2[simp]: "add_spvec (a, []) = a"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   401
  by (induct a, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   402
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   403
lemma sparse_row_vector_add: "sparse_row_vector (add_spvec (a,b)) = (sparse_row_vector a) + (sparse_row_vector b)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   404
  apply (rule add_spvec.induct[of _ a b])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   405
  apply (simp_all add: singleton_matrix_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   406
  apply (case_tac "a = aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   407
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   408
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   409
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   410
recdef add_spmat "measure (% (A,B). (length A)+(length B))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   411
  "add_spmat ([], bs) = bs"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   412
  "add_spmat (as, []) = as"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   413
  "add_spmat (a#as, b#bs) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   414
  if fst a < fst b then 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   415
    (a#(add_spmat (as, b#bs)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   416
  else (if fst b < fst a then
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   417
    (b#(add_spmat (a#as, bs)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   418
  else
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   419
    ((fst a, add_spvec (snd a, snd b))#(add_spmat (as, bs)))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   420
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   421
lemma sparse_row_add_spmat: "sparse_row_matrix (add_spmat (A, B)) = (sparse_row_matrix A) + (sparse_row_matrix B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   422
  apply (rule add_spmat.induct)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   423
  apply (auto simp add: sparse_row_matrix_cons sparse_row_vector_add move_matrix_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   424
  apply (case_tac "a=aa", simp, simp)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   425
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   426
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   427
lemma sorted_add_spvec_helper1[rule_format]: "add_spvec ((a,b)#arr, brr) = (ab, bb) # list \<longrightarrow> (ab = a | (brr \<noteq> [] & ab = fst (hd brr)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   428
  proof - 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   429
    have "(! x ab a. x = (a,b)#arr \<longrightarrow> add_spvec (x, brr) = (ab, bb) # list \<longrightarrow> (ab = a | (ab = fst (hd brr))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   430
      by (rule add_spvec.induct[of _ _ brr], auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   431
    then show ?thesis
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   432
      by (case_tac brr, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   433
  qed
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   434
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   435
lemma sorted_add_spmat_helper1[rule_format]: "add_spmat ((a,b)#arr, brr) = (ab, bb) # list \<longrightarrow> (ab = a | (brr \<noteq> [] & ab = fst (hd brr)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   436
  proof - 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   437
    have "(! x ab a. x = (a,b)#arr \<longrightarrow> add_spmat (x, brr) = (ab, bb) # list \<longrightarrow> (ab = a | (ab = fst (hd brr))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   438
      by (rule add_spmat.induct[of _ _ brr], auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   439
    then show ?thesis
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   440
      by (case_tac brr, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   441
  qed
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   442
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   443
lemma sorted_add_spvec_helper[rule_format]: "add_spvec (arr, brr) = (ab, bb) # list \<longrightarrow> ((arr \<noteq> [] & ab = fst (hd arr)) | (brr \<noteq> [] & ab = fst (hd brr)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   444
  apply (rule add_spvec.induct[of _ arr brr])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   445
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   446
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   447
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   448
lemma sorted_add_spmat_helper[rule_format]: "add_spmat (arr, brr) = (ab, bb) # list \<longrightarrow> ((arr \<noteq> [] & ab = fst (hd arr)) | (brr \<noteq> [] & ab = fst (hd brr)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   449
  apply (rule add_spmat.induct[of _ arr brr])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   450
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   451
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   452
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   453
lemma add_spvec_commute: "add_spvec (a, b) = add_spvec (b, a)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   454
  by (rule add_spvec.induct[of _ a b], auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   455
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   456
lemma add_spmat_commute: "add_spmat (a, b) = add_spmat (b, a)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   457
  apply (rule add_spmat.induct[of _ a b])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   458
  apply (simp_all add: add_spvec_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   459
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   460
  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   461
lemma sorted_add_spvec_helper2: "add_spvec ((a,b)#arr, brr) = (ab, bb) # list \<Longrightarrow> aa < a \<Longrightarrow> sorted_spvec ((aa, ba) # brr) \<Longrightarrow> aa < ab"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   462
  apply (drule sorted_add_spvec_helper1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   463
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   464
  apply (case_tac brr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   465
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   466
  apply (drule_tac sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   467
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   468
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   469
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   470
lemma sorted_add_spmat_helper2: "add_spmat ((a,b)#arr, brr) = (ab, bb) # list \<Longrightarrow> aa < a \<Longrightarrow> sorted_spvec ((aa, ba) # brr) \<Longrightarrow> aa < ab"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   471
  apply (drule sorted_add_spmat_helper1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   472
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   473
  apply (case_tac brr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   474
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   475
  apply (drule_tac sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   476
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   477
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   478
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   479
lemma sorted_spvec_add_spvec[rule_format]: "sorted_spvec a \<longrightarrow> sorted_spvec b \<longrightarrow> sorted_spvec (add_spvec (a, b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   480
  apply (rule add_spvec.induct[of _ a b])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   481
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   482
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   483
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   484
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   485
  apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   486
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   487
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   488
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   489
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   490
  apply (simp add: sorted_add_spvec_helper2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   491
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   492
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   493
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   494
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   495
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   496
  apply (frule_tac as=arr in sorted_spvec_cons1, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   497
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   498
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   499
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   500
  apply (simp add: sorted_add_spvec_helper2 add_spvec_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   501
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   502
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   503
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   504
  apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   505
  apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   506
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   507
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   508
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   509
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   510
  apply (drule_tac sorted_add_spvec_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   511
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   512
  apply (case_tac arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   513
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   514
  apply (drule sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   515
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   516
  apply (case_tac brr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   517
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   518
  apply (drule sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   519
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   520
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   521
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   522
lemma sorted_spvec_add_spmat[rule_format]: "sorted_spvec A \<longrightarrow> sorted_spvec B \<longrightarrow> sorted_spvec (add_spmat (A, B))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   523
  apply (rule add_spmat.induct[of _ A B])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   524
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   525
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   526
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   527
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   528
  apply (frule_tac as=bs in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   529
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   530
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   531
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   532
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   533
  apply (simp add: sorted_add_spmat_helper2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   534
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   535
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   536
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   537
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   538
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   539
  apply (frule_tac as=as in sorted_spvec_cons1, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   540
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   541
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   542
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   543
  apply (simp add: sorted_add_spmat_helper2 add_spmat_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   544
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   545
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   546
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   547
  apply (frule_tac as=as in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   548
  apply (frule_tac as=bs in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   549
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   550
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   551
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   552
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   553
  apply (drule_tac sorted_add_spmat_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   554
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   555
  apply (case_tac as)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   556
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   557
  apply (drule sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   558
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   559
  apply (case_tac bs)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   560
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   561
  apply (drule sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   562
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   563
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   564
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   565
lemma sorted_spmat_add_spmat[rule_format]: "sorted_spmat A \<longrightarrow> sorted_spmat B \<longrightarrow> sorted_spmat (add_spmat (A, B))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   566
  apply (rule add_spmat.induct[of _ A B])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   567
  apply (simp_all add: sorted_spvec_add_spvec)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   568
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   569
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   570
consts
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   571
  le_spvec :: "('a::lordered_ab_group) spvec * 'a spvec \<Rightarrow> bool" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   572
  le_spmat :: "('a::lordered_ab_group) spmat * 'a spmat \<Rightarrow> bool" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   573
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   574
recdef le_spvec "measure (% (a,b). (length a) + (length b))" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   575
  "le_spvec ([], []) = True"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   576
  "le_spvec (a#as, []) = ((snd a <= 0) & (le_spvec (as, [])))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   577
  "le_spvec ([], b#bs) = ((0 <= snd b) & (le_spvec ([], bs)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   578
  "le_spvec (a#as, b#bs) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   579
  if (fst a < fst b) then 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   580
    ((snd a <= 0) & (le_spvec (as, b#bs)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   581
  else (if (fst b < fst a) then
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   582
    ((0 <= snd b) & (le_spvec (a#as, bs)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   583
  else 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   584
    ((snd a <= snd b) & (le_spvec (as, bs)))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   585
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   586
recdef le_spmat "measure (% (a,b). (length a) + (length b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   587
  "le_spmat ([], []) = True"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   588
  "le_spmat (a#as, []) = (le_spvec (snd a, []) & (le_spmat (as, [])))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   589
  "le_spmat ([], b#bs) = (le_spvec ([], snd b) & (le_spmat ([], bs)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   590
  "le_spmat (a#as, b#bs) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   591
  if fst a < fst b then
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   592
    (le_spvec(snd a,[]) & le_spmat(as, b#bs))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   593
  else (if (fst b < fst a) then 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   594
    (le_spvec([], snd b) & le_spmat(a#as, bs))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   595
  else
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   596
    (le_spvec(snd a, snd b) & le_spmat (as, bs))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   597
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   598
constdefs
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   599
  disj_matrices :: "('a::zero) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   600
  "disj_matrices A B == (! j i. (Rep_matrix A j i \<noteq> 0) \<longrightarrow> (Rep_matrix B j i = 0)) & (! j i. (Rep_matrix B j i \<noteq> 0) \<longrightarrow> (Rep_matrix A j i = 0))"  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   601
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   602
ML {* simp_depth_limit := 2 *}
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   603
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   604
lemma disj_matrices_add: "disj_matrices A B \<Longrightarrow> disj_matrices C D \<Longrightarrow> disj_matrices A D \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   605
  (A + B <= C + D) = (A <= C & B <= (D::('a::lordered_ab_group) matrix))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   606
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   607
  apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   608
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   609
  apply (erule conjE)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   610
  apply (drule_tac j=j and i=i in spec2)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   611
  apply (case_tac "Rep_matrix B j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   612
  apply (case_tac "Rep_matrix D j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   613
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   614
  apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   615
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   616
  apply (erule conjE)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   617
  apply (drule_tac j=j and i=i in spec2)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   618
  apply (case_tac "Rep_matrix A j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   619
  apply (case_tac "Rep_matrix C j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   620
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   621
  apply (erule add_mono)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   622
  apply (assumption)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   623
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   624
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   625
lemma disj_matrices_zero1[simp]: "disj_matrices 0 B"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   626
by (simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   627
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   628
lemma disj_matrices_zero2[simp]: "disj_matrices A 0"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   629
by (simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   630
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   631
lemma disj_matrices_commute: "disj_matrices A B = disj_matrices B A"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   632
by (auto simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   633
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   634
lemma disj_matrices_add_le_zero: "disj_matrices A B \<Longrightarrow>
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   635
  (A + B <= 0) = (A <= 0 & (B::('a::lordered_ab_group) matrix) <= 0)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   636
by (rule disj_matrices_add[of A B 0 0, simplified])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   637
 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   638
lemma disj_matrices_add_zero_le: "disj_matrices A B \<Longrightarrow>
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   639
  (0 <= A + B) = (0 <= A & 0 <= (B::('a::lordered_ab_group) matrix))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   640
by (rule disj_matrices_add[of 0 0 A B, simplified])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   641
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   642
lemma disj_matrices_add_x_le: "disj_matrices A B \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   643
  (A <= B + C) = (A <= C & 0 <= (B::('a::lordered_ab_group) matrix))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   644
by (auto simp add: disj_matrices_add[of 0 A B C, simplified])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   645
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   646
lemma disj_matrices_add_le_x: "disj_matrices A B \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   647
  (B + A <= C) = (A <= C &  (B::('a::lordered_ab_group) matrix) <= 0)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   648
by (auto simp add: disj_matrices_add[of B A 0 C,simplified] disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   649
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   650
lemma disj_sparse_row_singleton: "i <= j \<Longrightarrow> sorted_spvec((j,y)#v) \<Longrightarrow> disj_matrices (sparse_row_vector v) (singleton_matrix 0 i x)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   651
  apply (simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   652
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   653
  apply (rule neg_imp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   654
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   655
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   656
  apply (rule sorted_sparse_row_vector_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   657
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   658
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   659
  apply (rule sorted_sparse_row_vector_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   660
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   661
  done 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   662
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   663
lemma disj_matrices_x_add: "disj_matrices A B \<Longrightarrow> disj_matrices A C \<Longrightarrow> disj_matrices (A::('a::lordered_ab_group) matrix) (B+C)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   664
  apply (simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   665
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   666
  apply (drule_tac j=j and i=i in spec2)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   667
  apply (case_tac "Rep_matrix B j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   668
  apply (case_tac "Rep_matrix C j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   669
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   670
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   671
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   672
lemma disj_matrices_add_x: "disj_matrices A B \<Longrightarrow> disj_matrices A C \<Longrightarrow> disj_matrices (B+C) (A::('a::lordered_ab_group) matrix)" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   673
  by (simp add: disj_matrices_x_add disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   674
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   675
lemma disj_singleton_matrices[simp]: "disj_matrices (singleton_matrix j i x) (singleton_matrix u v y) = (j \<noteq> u | i \<noteq> v | x = 0 | y = 0)" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   676
  by (auto simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   677
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   678
lemma disj_move_sparse_vec_mat[simplified disj_matrices_commute]: 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   679
  "j <= a \<Longrightarrow> sorted_spvec((a,c)#as) \<Longrightarrow> disj_matrices (move_matrix (sparse_row_vector b) (int j) i) (sparse_row_matrix as)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   680
  apply (auto simp add: neg_def disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   681
  apply (drule nrows_notzero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   682
  apply (drule less_le_trans[OF _ nrows_spvec])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   683
  apply (subgoal_tac "ja = j")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   684
  apply (simp add: sorted_sparse_row_matrix_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   685
  apply (arith)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   686
  apply (rule nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   687
  apply (rule order_trans[of _ 1 _])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   688
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   689
  apply (case_tac "nat (int ja - int j) = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   690
  apply (case_tac "ja = j")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   691
  apply (simp add: sorted_sparse_row_matrix_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   692
  apply arith+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   693
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   694
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   695
lemma disj_move_sparse_row_vector_twice:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   696
  "j \<noteq> u \<Longrightarrow> disj_matrices (move_matrix (sparse_row_vector a) j i) (move_matrix (sparse_row_vector b) u v)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   697
  apply (auto simp add: neg_def disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   698
  apply (rule nrows, rule order_trans[of _ 1], simp, drule nrows_notzero, drule less_le_trans[OF _ nrows_spvec], arith)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   699
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   700
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   701
lemma le_spvec_iff_sparse_row_le[rule_format]: "(sorted_spvec a) \<longrightarrow> (sorted_spvec b) \<longrightarrow> (le_spvec (a,b)) = (sparse_row_vector a <= sparse_row_vector b)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   702
  apply (rule le_spvec.induct)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   703
  apply (simp_all add: sorted_spvec_cons1 disj_matrices_add_le_zero disj_matrices_add_zero_le 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   704
    disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   705
  apply (rule conjI, intro strip)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   706
  apply (simp add: sorted_spvec_cons1)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   707
  apply (subst disj_matrices_add_x_le)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   708
  apply (simp add: disj_sparse_row_singleton[OF less_imp_le] disj_matrices_x_add disj_matrices_commute)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   709
  apply (simp add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   710
  apply (simp, blast)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   711
  apply (intro strip, rule conjI, intro strip)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   712
  apply (simp add: sorted_spvec_cons1)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   713
  apply (subst disj_matrices_add_le_x)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   714
  apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_sparse_row_singleton[OF less_imp_le] disj_matrices_commute disj_matrices_x_add)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   715
  apply (blast)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   716
  apply (intro strip)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   717
  apply (simp add: sorted_spvec_cons1)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   718
  apply (case_tac "a=aa", simp_all)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   719
  apply (subst disj_matrices_add)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   720
  apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   721
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   722
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   723
lemma le_spvec_empty2_sparse_row[rule_format]: "(sorted_spvec b) \<longrightarrow> (le_spvec (b,[]) = (sparse_row_vector b <= 0))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   724
  apply (induct b)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   725
  apply (simp_all add: sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   726
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   727
  apply (subst disj_matrices_add_le_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   728
  apply (simp add: disj_matrices_commute disj_sparse_row_singleton sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   729
  apply (rule_tac y = "snd a" in disj_sparse_row_singleton[OF order_refl])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   730
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   731
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   732
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   733
lemma le_spvec_empty1_sparse_row[rule_format]: "(sorted_spvec b) \<longrightarrow> (le_spvec ([],b) = (0 <= sparse_row_vector b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   734
  apply (induct b)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   735
  apply (simp_all add: sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   736
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   737
  apply (subst disj_matrices_add_zero_le)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   738
  apply (simp add: disj_matrices_commute disj_sparse_row_singleton sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   739
  apply (rule_tac y = "snd a" in disj_sparse_row_singleton[OF order_refl])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   740
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   741
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   742
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   743
lemma le_spmat_iff_sparse_row_le[rule_format]: "(sorted_spvec A) \<longrightarrow> (sorted_spmat A) \<longrightarrow> (sorted_spvec B) \<longrightarrow> (sorted_spmat B) \<longrightarrow> 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   744
  le_spmat(A, B) = (sparse_row_matrix A <= sparse_row_matrix B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   745
  apply (rule le_spmat.induct)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   746
  apply (simp add: sparse_row_matrix_cons disj_matrices_add_le_zero disj_matrices_add_zero_le disj_move_sparse_vec_mat[OF order_refl] 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   747
    disj_matrices_commute sorted_spvec_cons1 le_spvec_empty2_sparse_row le_spvec_empty1_sparse_row)+ 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   748
  apply (rule conjI, intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   749
  apply (simp add: sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   750
  apply (subst disj_matrices_add_x_le)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   751
  apply (rule disj_matrices_add_x)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   752
  apply (simp add: disj_move_sparse_row_vector_twice)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   753
  apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   754
  apply (simp add: disj_move_sparse_vec_mat[OF order_refl] disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   755
  apply (simp, blast)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   756
  apply (intro strip, rule conjI, intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   757
  apply (simp add: sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   758
  apply (subst disj_matrices_add_le_x)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   759
  apply (simp add: disj_move_sparse_vec_mat[OF order_refl])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   760
  apply (rule disj_matrices_x_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   761
  apply (simp add: disj_move_sparse_row_vector_twice)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   762
  apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   763
  apply (simp, blast)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   764
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   765
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   766
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   767
  apply (subst disj_matrices_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   768
  apply (simp_all add: disj_matrices_commute disj_move_sparse_vec_mat[OF order_refl])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   769
  apply (simp add: sorted_spvec_cons1 le_spvec_iff_sparse_row_le)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   770
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   771
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   772
ML {* simp_depth_limit := 999 *}
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   773
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   774
consts 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   775
   abs_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   776
   minus_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   777
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   778
primrec
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   779
  "abs_spmat [] = []"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   780
  "abs_spmat (a#as) = (fst a, abs_spvec (snd a))#(abs_spmat as)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   781
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   782
primrec
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   783
  "minus_spmat [] = []"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   784
  "minus_spmat (a#as) = (fst a, minus_spvec (snd a))#(minus_spmat as)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   785
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   786
lemma sparse_row_matrix_minus:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   787
  "sparse_row_matrix (minus_spmat A) = - (sparse_row_matrix A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   788
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   789
  apply (simp_all add: sparse_row_vector_minus sparse_row_matrix_cons)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   790
  apply (subst Rep_matrix_inject[symmetric])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   791
  apply (rule ext)+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   792
  apply simp
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   793
  done
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   794
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   795
lemma Rep_sparse_row_vector_zero: "x \<noteq> 0 \<Longrightarrow> Rep_matrix (sparse_row_vector v) x y = 0"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   796
proof -
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   797
  assume x:"x \<noteq> 0"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   798
  have r:"nrows (sparse_row_vector v) <= Suc 0" by (rule nrows_spvec)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   799
  show ?thesis
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   800
    apply (rule nrows)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   801
    apply (subgoal_tac "Suc 0 <= x")
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   802
    apply (insert r)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   803
    apply (simp only:)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   804
    apply (insert x)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   805
    apply arith
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   806
    done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   807
qed
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   808
    
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   809
lemma sparse_row_matrix_abs:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   810
  "sorted_spvec A \<Longrightarrow> sorted_spmat A \<Longrightarrow> sparse_row_matrix (abs_spmat A) = abs (sparse_row_matrix A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   811
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   812
  apply (simp_all add: sparse_row_vector_abs sparse_row_matrix_cons)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   813
  apply (frule_tac sorted_spvec_cons1, simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   814
  apply (subst Rep_matrix_inject[symmetric])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   815
  apply (rule ext)+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   816
  apply auto
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   817
  apply (case_tac "x=a")
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   818
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   819
  apply (subst sorted_sparse_row_matrix_zero)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   820
  apply auto
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   821
  apply (subst Rep_sparse_row_vector_zero)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   822
  apply (simp_all add: neg_def)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   823
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   824
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   825
lemma sorted_spvec_minus_spmat: "sorted_spvec A \<Longrightarrow> sorted_spvec (minus_spmat A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   826
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   827
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   828
  apply (frule sorted_spvec_cons1, simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   829
  apply (simp add: sorted_spvec.simps)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   830
  apply (case_tac list)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   831
  apply (simp_all)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   832
  done 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   833
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   834
lemma sorted_spvec_abs_spmat: "sorted_spvec A \<Longrightarrow> sorted_spvec (abs_spmat A)" 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   835
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   836
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   837
  apply (frule sorted_spvec_cons1, simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   838
  apply (simp add: sorted_spvec.simps)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   839
  apply (case_tac list)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   840
  apply (simp_all)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   841
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   842
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   843
lemma sorted_spmat_minus_spmat: "sorted_spmat A \<Longrightarrow> sorted_spmat (minus_spmat A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   844
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   845
  apply (simp_all add: sorted_spvec_minus_spvec)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   846
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   847
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   848
lemma sorted_spmat_abs_spmat: "sorted_spmat A \<Longrightarrow> sorted_spmat (abs_spmat A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   849
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   850
  apply (simp_all add: sorted_spvec_abs_spvec)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   851
  done
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   852
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   853
constdefs
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   854
  diff_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   855
  "diff_spmat A B == add_spmat (A, minus_spmat B)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   856
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   857
lemma sorted_spmat_diff_spmat: "sorted_spmat A \<Longrightarrow> sorted_spmat B \<Longrightarrow> sorted_spmat (diff_spmat A B)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   858
  by (simp add: diff_spmat_def sorted_spmat_minus_spmat sorted_spmat_add_spmat)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   859
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   860
lemma sorted_spvec_diff_spmat: "sorted_spvec A \<Longrightarrow> sorted_spvec B \<Longrightarrow> sorted_spvec (diff_spmat A B)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   861
  by (simp add: diff_spmat_def sorted_spvec_minus_spmat sorted_spvec_add_spmat)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   862
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   863
lemma sparse_row_diff_spmat: "sparse_row_matrix (diff_spmat A B ) = (sparse_row_matrix A) - (sparse_row_matrix B)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   864
  by (simp add: diff_spmat_def sparse_row_add_spmat sparse_row_matrix_minus)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   865
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   866
constdefs
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   867
  sorted_sparse_matrix :: "'a spmat \<Rightarrow> bool"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   868
  "sorted_sparse_matrix A == (sorted_spvec A) & (sorted_spmat A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   869
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   870
lemma sorted_sparse_matrix_imp_spvec: "sorted_sparse_matrix A \<Longrightarrow> sorted_spvec A"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   871
  by (simp add: sorted_sparse_matrix_def)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   872
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   873
lemma sorted_sparse_matrix_imp_spmat: "sorted_sparse_matrix A \<Longrightarrow> sorted_spmat A"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   874
  by (simp add: sorted_sparse_matrix_def)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   875
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   876
lemmas sparse_row_matrix_op_simps =
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   877
  sorted_sparse_matrix_imp_spmat sorted_sparse_matrix_imp_spvec
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   878
  sparse_row_add_spmat sorted_spvec_add_spmat sorted_spmat_add_spmat
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   879
  sparse_row_diff_spmat sorted_spvec_diff_spmat sorted_spmat_diff_spmat
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   880
  sparse_row_matrix_minus sorted_spvec_minus_spmat sorted_spmat_minus_spmat
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   881
  sparse_row_mult_spmat sorted_spvec_mult_spmat sorted_spmat_mult_spmat
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   882
  sparse_row_matrix_abs sorted_spvec_abs_spmat sorted_spmat_abs_spmat
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   883
  le_spmat_iff_sparse_row_le
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   884
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   885
lemma zero_eq_Numeral0: "(0::_::number_ring) = Numeral0" by simp
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   886
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   887
lemmas sparse_row_matrix_arith_simps[simplified zero_eq_Numeral0] = 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   888
  mult_spmat.simps mult_spvec_spmat.simps 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   889
  addmult_spvec.simps 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   890
  smult_spvec_empty smult_spvec_cons
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   891
  add_spmat.simps add_spvec.simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   892
  minus_spmat.simps minus_spvec.simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   893
  abs_spmat.simps abs_spvec.simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   894
  diff_spmat_def
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   895
  le_spmat.simps le_spvec.simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   896
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   897
lemmas sorted_sp_simps = 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   898
  sorted_spvec.simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   899
  sorted_spmat.simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   900
  sorted_sparse_matrix_def
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   901
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   902
lemma bool1: "(\<not> True) = False"  by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   903
lemma bool2: "(\<not> False) = True"  by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   904
lemma bool3: "((P\<Colon>bool) \<and> True) = P" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   905
lemma bool4: "(True \<and> (P\<Colon>bool)) = P" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   906
lemma bool5: "((P\<Colon>bool) \<and> False) = False" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   907
lemma bool6: "(False \<and> (P\<Colon>bool)) = False" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   908
lemma bool7: "((P\<Colon>bool) \<or> True) = True" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   909
lemma bool8: "(True \<or> (P\<Colon>bool)) = True" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   910
lemma bool9: "((P\<Colon>bool) \<or> False) = P" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   911
lemma bool10: "(False \<or> (P\<Colon>bool)) = P" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   912
lemmas boolarith = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   913
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   914
lemma if_case_eq: "(if b then x else y) = (case b of True => x | False => y)" by simp
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   915
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   916
lemma spm_linprog_dual_estimate_1:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   917
  assumes  
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   918
  "sorted_sparse_matrix A1"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   919
  "sorted_sparse_matrix A2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   920
  "sorted_sparse_matrix c1"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   921
  "sorted_sparse_matrix c2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   922
  "sorted_sparse_matrix y"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   923
  "sorted_spvec b"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   924
  "sorted_spvec r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   925
  "le_spmat ([], y)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   926
  "A * x \<le> sparse_row_matrix (b::('a::lordered_ring) spmat)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   927
  "sparse_row_matrix A1 <= A"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   928
  "A <= sparse_row_matrix A2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   929
  "sparse_row_matrix c1 <= c"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   930
  "c <= sparse_row_matrix c2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   931
  "abs x \<le> sparse_row_matrix r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   932
  shows
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   933
  "c * x \<le> sparse_row_matrix (add_spmat (mult_spmat y b, mult_spmat (add_spmat (add_spmat (mult_spmat y (diff_spmat A2 A1), 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   934
  abs_spmat (diff_spmat (mult_spmat y A1) c1)), diff_spmat c2 c1)) r))"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   935
  by (insert prems, simp add: sparse_row_matrix_op_simps linprog_dual_estimate_1[where A=A])
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   936
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   937
end