src/HOL/Matrix/SparseMatrix.thy
author wenzelm
Tue, 26 Aug 2008 11:42:46 +0200
changeset 28000 ca56bbb99607
parent 27653 180e28bab764
child 28562 4e74209f113e
permissions -rw-r--r--
replaced /home/isabelle/html-data/isabelle-repos by /home/isabelle-repository/repos;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
16487
2060ebae96f9 proper header;
wenzelm
parents: 16417
diff changeset
     1
(*  Title:      HOL/Matrix/SparseMatrix.thy
2060ebae96f9 proper header;
wenzelm
parents: 16417
diff changeset
     2
    ID:         $Id$
2060ebae96f9 proper header;
wenzelm
parents: 16417
diff changeset
     3
    Author:     Steven Obua
2060ebae96f9 proper header;
wenzelm
parents: 16417
diff changeset
     4
*)
2060ebae96f9 proper header;
wenzelm
parents: 16417
diff changeset
     5
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
     6
theory SparseMatrix
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
     7
imports "./Matrix"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
     8
begin
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
     9
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    10
types 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    11
  'a spvec = "(nat * 'a) list"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    12
  'a spmat = "('a spvec) spvec"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    13
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    14
definition sparse_row_vector :: "('a::ab_group_add) spvec \<Rightarrow> 'a matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
    15
  sparse_row_vector_def: "sparse_row_vector arr = foldl (% m x. m + (singleton_matrix 0 (fst x) (snd x))) 0 arr"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    16
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    17
definition sparse_row_matrix :: "('a::ab_group_add) spmat \<Rightarrow> 'a matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
    18
  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"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    19
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
    20
code_datatype sparse_row_vector sparse_row_matrix
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
    21
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
    22
lemma sparse_row_vector_empty [simp]: "sparse_row_vector [] = 0"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    23
  by (simp add: sparse_row_vector_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    24
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
    25
lemma sparse_row_matrix_empty [simp]: "sparse_row_matrix [] = 0"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    26
  by (simp add: sparse_row_matrix_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    27
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
    28
lemmas [code func] = sparse_row_vector_empty [symmetric]
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
    29
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    30
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
    31
  by (induct l, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    32
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    33
lemma sparse_row_vector_cons[simp]:
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    34
  "sparse_row_vector (a # arr) = (singleton_matrix 0 (fst a) (snd a)) + (sparse_row_vector arr)"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    35
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    36
  apply (auto simp add: sparse_row_vector_def)
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    37
  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"])
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    38
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    39
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    40
lemma sparse_row_vector_append[simp]:
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    41
  "sparse_row_vector (a @ b) = (sparse_row_vector a) + (sparse_row_vector b)"
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    42
  by (induct a) auto
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    43
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    44
lemma nrows_spvec[simp]: "nrows (sparse_row_vector x) <= (Suc 0)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    45
  apply (induct x)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    46
  apply (simp_all add: add_nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    47
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    48
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    49
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
    50
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    51
  apply (auto simp add: sparse_row_matrix_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    52
  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
    53
    "% a m. (move_matrix (sparse_row_vector (snd a)) (int (fst a)) 0) + m"])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    54
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    55
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    56
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
    57
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    58
  apply (auto simp add: sparse_row_matrix_cons)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    59
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    60
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    61
primrec sorted_spvec :: "'a spvec \<Rightarrow> bool" where
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    62
  "sorted_spvec [] = True"
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    63
  | sorted_spvec_step: "sorted_spvec (a#as) = (case as of [] \<Rightarrow> True | b#bs \<Rightarrow> ((fst a < fst b) & (sorted_spvec as)))" 
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    64
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    65
primrec sorted_spmat :: "'a spmat \<Rightarrow> bool" where
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    66
  "sorted_spmat [] = True"
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
    67
  | "sorted_spmat (a#as) = ((sorted_spvec (snd a)) & (sorted_spmat as))"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    68
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    69
declare sorted_spvec.simps [simp del]
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    70
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    71
lemma sorted_spvec_empty[simp]: "sorted_spvec [] = True"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    72
by (simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    73
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    74
lemma sorted_spvec_cons1: "sorted_spvec (a#as) \<Longrightarrow> sorted_spvec as"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    75
apply (induct as)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    76
apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    77
done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    78
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    79
lemma sorted_spvec_cons2: "sorted_spvec (a#b#t) \<Longrightarrow> sorted_spvec (a#t)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    80
apply (induct t)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    81
apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    82
done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    83
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    84
lemma sorted_spvec_cons3: "sorted_spvec(a#b#t) \<Longrightarrow> fst a < fst b"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    85
apply (auto simp add: sorted_spvec.simps)
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_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
    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
done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    94
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    95
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
    96
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    97
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    98
  apply (frule sorted_spvec_cons2, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
    99
  apply (frule sorted_spvec_cons3, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   100
  apply (simp add: sparse_row_matrix_cons neg_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   101
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   102
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   103
primrec minus_spvec :: "('a::ab_group_add) spvec \<Rightarrow> 'a spvec" where
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   104
  "minus_spvec [] = []"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
   105
  | "minus_spvec (a#as) = (fst a, -(snd a))#(minus_spvec as)"
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   106
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   107
primrec abs_spvec ::  "('a::lordered_ab_group_add_abs) spvec \<Rightarrow> 'a spvec" where
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   108
  "abs_spvec [] = []"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
   109
  | "abs_spvec (a#as) = (fst a, abs (snd a))#(abs_spvec as)"
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   110
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   111
lemma sparse_row_vector_minus: 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   112
  "sparse_row_vector (minus_spvec v) = - (sparse_row_vector v)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   113
  apply (induct v)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   114
  apply (simp_all add: sparse_row_vector_cons)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   115
  apply (simp add: Rep_matrix_inject[symmetric])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   116
  apply (rule ext)+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   117
  apply simp
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   118
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   119
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   120
instance matrix :: (lordered_ab_group_add_abs) lordered_ab_group_add_abs
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   121
apply default
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   122
unfolding abs_matrix_def .. (*FIXME move*)
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   123
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   124
lemma sparse_row_vector_abs:
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   125
  "sorted_spvec (v :: 'a::lordered_ring spvec) \<Longrightarrow> sparse_row_vector (abs_spvec v) = abs (sparse_row_vector v)"
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   126
  apply (induct v)
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   127
  apply simp_all
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   128
  apply (frule_tac sorted_spvec_cons1, simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   129
  apply (simp only: Rep_matrix_inject[symmetric])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   130
  apply (rule ext)+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   131
  apply auto
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 15197
diff changeset
   132
  apply (subgoal_tac "Rep_matrix (sparse_row_vector v) 0 a = 0")
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   133
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   134
  apply (rule sorted_sparse_row_vector_zero)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   135
  apply auto
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   136
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   137
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   138
lemma sorted_spvec_minus_spvec:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   139
  "sorted_spvec v \<Longrightarrow> sorted_spvec (minus_spvec v)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   140
  apply (induct v)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   141
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   142
  apply (frule sorted_spvec_cons1, simp)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 15197
diff changeset
   143
  apply (simp add: sorted_spvec.simps split:list.split_asm)
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   144
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   145
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   146
lemma sorted_spvec_abs_spvec:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   147
  "sorted_spvec v \<Longrightarrow> sorted_spvec (abs_spvec v)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   148
  apply (induct v)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   149
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   150
  apply (frule sorted_spvec_cons1, simp)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 15197
diff changeset
   151
  apply (simp add: sorted_spvec.simps split:list.split_asm)
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   152
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   153
  
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   154
definition
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   155
  "smult_spvec y = map (% a. (fst a, y * snd a))"  
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   156
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   157
lemma smult_spvec_empty[simp]: "smult_spvec y [] = []"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   158
  by (simp add: smult_spvec_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   159
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   160
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
   161
  by (simp add: smult_spvec_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   162
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   163
consts addmult_spvec :: "('a::ring) * 'a spvec * 'a spvec \<Rightarrow> 'a spvec"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   164
recdef addmult_spvec "measure (% (y, a, b). length a + (length b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   165
  "addmult_spvec (y, arr, []) = arr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   166
  "addmult_spvec (y, [], brr) = smult_spvec y brr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   167
  "addmult_spvec (y, a#arr, b#brr) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   168
    if (fst a) < (fst b) then (a#(addmult_spvec (y, arr, b#brr))) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   169
    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
   170
    else ((fst a, (snd a)+ y*(snd b))#(addmult_spvec (y, arr,brr)))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   171
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   172
lemma addmult_spvec_empty1[simp]: "addmult_spvec (y, [], a) = smult_spvec y a"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
   173
  by (induct a) auto
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   174
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   175
lemma addmult_spvec_empty2[simp]: "addmult_spvec (y, a, []) = a"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
   176
  by (induct a) auto
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   177
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   178
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
   179
  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
   180
  apply (induct a)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   181
  apply (simp_all add: apply_matrix_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   182
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   183
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   184
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
   185
  apply (induct a)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   186
  apply (simp_all add: smult_spvec_cons scalar_mult_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   187
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   188
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   189
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
   190
  (sparse_row_vector a) + (scalar_mult y (sparse_row_vector b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   191
  apply (rule addmult_spvec.induct[of _ y])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   192
  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
   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)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 15197
diff changeset
   198
  apply (auto simp add: sorted_spvec.simps split:list.split_asm)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   199
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   200
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   201
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
   202
  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
   203
  apply (induct brr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   204
  apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   205
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   206
  apply (auto)
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
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   210
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   211
lemma sorted_spvec_addmult_spvec_helper2: 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   212
 "\<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
   213
       \<Longrightarrow> sorted_spvec ((a, b) # addmult_spvec (y, arr, (aa, ba) # brr))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   214
  apply (induct arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   215
  apply (auto simp add: smult_spvec_def sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   216
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   217
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   218
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   219
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   220
lemma sorted_spvec_addmult_spvec_helper3[rule_format]:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   221
  "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
   222
     \<longrightarrow> sorted_spvec ((aa, b + y * ba) # (addmult_spvec (y, arr, brr)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   223
  apply (rule addmult_spvec.induct[of _ y arr brr])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   224
  apply (simp_all add: sorted_spvec.simps smult_spvec_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   225
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   226
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   227
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
   228
  apply (rule addmult_spvec.induct[of _ y a b])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   229
  apply (simp_all add: sorted_smult_spvec)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   230
  apply (rule conjI, intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   231
  apply (case_tac "~(a < aa)")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   232
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   233
  apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   234
  apply (simp add: sorted_spvec_addmult_spvec_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   235
  apply (intro strip | rule conjI)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   236
  apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   237
  apply (simp add: sorted_spvec_addmult_spvec_helper2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   238
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   239
  apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   240
  apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   241
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   242
  apply (simp_all add: sorted_spvec_addmult_spvec_helper3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   243
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   244
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   245
consts 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   246
  mult_spvec_spmat :: "('a::lordered_ring) spvec * 'a spvec * 'a spmat  \<Rightarrow> 'a spvec"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   247
recdef mult_spvec_spmat "measure (% (c, arr, brr). (length arr) + (length brr))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   248
  "mult_spvec_spmat (c, [], brr) = c"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   249
  "mult_spvec_spmat (c, arr, []) = c"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   250
  "mult_spvec_spmat (c, a#arr, b#brr) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   251
     if ((fst a) < (fst b)) then (mult_spvec_spmat (c, arr, b#brr))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   252
     else (if ((fst b) < (fst a)) then (mult_spvec_spmat (c, a#arr, brr)) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   253
     else (mult_spvec_spmat (addmult_spvec (snd a, c, snd b), arr, brr))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   254
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   255
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
   256
  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
   257
proof -
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   258
  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
   259
  have not_iff: "!! a b. a = b \<Longrightarrow> (~ a) = (~ b)" by simp
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   260
  have max_helper: "!! a b. ~ (a <= max (Suc a) b) \<Longrightarrow> False"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   261
    by arith
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   262
  {
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   263
    fix a 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   264
    fix v
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   265
    assume a:"a < nrows(sparse_row_vector v)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   266
    have b:"nrows(sparse_row_vector v) <= 1" by simp
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   267
    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
   268
    then have "a = 0" by simp
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   269
  }
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   270
  note nrows_helper = this
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   271
  show ?thesis
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   272
    apply (rule mult_spvec_spmat.induct)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   273
    apply simp+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   274
    apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   275
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   276
    apply (frule_tac as=brr in sorted_spvec_cons1)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
   277
    apply (simp add: ring_simps sparse_row_matrix_cons)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15236
diff changeset
   278
    apply (simplesubst Rep_matrix_zero_imp_mult_zero) 
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   279
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   280
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   281
    apply (rule disjI2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   282
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   283
    apply (subst nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   284
    apply (rule  order_trans[of _ 1])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   285
    apply (simp add: comp_1)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   286
    apply (subst Rep_matrix_zero_imp_mult_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   287
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   288
    apply (case_tac "k <= aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   289
    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
   290
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   291
    apply (rule impI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   292
    apply (rule disjI2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   293
    apply (rule nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   294
    apply (rule order_trans[of _ 1])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   295
    apply (simp_all add: comp_1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   296
    
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   297
    apply (intro strip | rule conjI)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   298
    apply (frule_tac as=arr in sorted_spvec_cons1)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
   299
    apply (simp add: ring_simps)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   300
    apply (subst Rep_matrix_zero_imp_mult_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   301
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   302
    apply (rule disjI2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   303
    apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   304
    apply (simp add: sparse_row_matrix_cons neg_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   305
    apply (case_tac "a <= aa")  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   306
    apply (erule sorted_sparse_row_matrix_zero)  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   307
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   308
    apply (intro strip)
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 (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   311
    apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   312
    apply (frule_tac as=brr in sorted_spvec_cons1)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
   313
    apply (simp add: sparse_row_matrix_cons ring_simps sparse_row_vector_addmult_spvec)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   314
    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
   315
    apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   316
    apply (rule sorted_sparse_row_matrix_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   317
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   318
    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
   319
    apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   320
    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
   321
    apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   322
    apply (simp add: neg_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   323
    apply (drule nrows_notzero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   324
    apply (drule nrows_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   325
    apply (arith)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   326
    
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   327
    apply (subst Rep_matrix_inject[symmetric])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   328
    apply (rule ext)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   329
    apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   330
    apply (subst Rep_matrix_mult)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   331
    apply (rule_tac j1=aa in ssubst[OF foldseq_almostzero])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   332
    apply (simp_all)
20432
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20283
diff changeset
   333
    apply (intro strip, rule conjI)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   334
    apply (intro strip)
20432
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20283
diff changeset
   335
    apply (drule_tac max_helper)
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20283
diff changeset
   336
    apply (simp)
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20283
diff changeset
   337
    apply (auto)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   338
    apply (rule zero_imp_mult_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   339
    apply (rule disjI2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   340
    apply (rule nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   341
    apply (rule order_trans[of _ 1])
20432
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20283
diff changeset
   342
    apply (simp)
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20283
diff changeset
   343
    apply (simp)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   344
    done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   345
qed
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   346
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   347
lemma sorted_mult_spvec_spmat[rule_format]: 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   348
  "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
   349
  apply (rule mult_spvec_spmat.induct[of _ c a B])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   350
  apply (simp_all add: sorted_addmult_spvec)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   351
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   352
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   353
consts 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   354
  mult_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   355
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   356
primrec 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   357
  "mult_spmat [] A = []"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   358
  "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
   359
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   360
lemma sparse_row_mult_spmat[rule_format]: 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   361
  "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
   362
  apply (induct A)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
   363
  apply (auto simp add: sparse_row_matrix_cons sparse_row_mult_spvec_spmat ring_simps move_matrix_mult)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   364
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   365
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   366
lemma sorted_spvec_mult_spmat[rule_format]:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   367
  "sorted_spvec (A::('a::lordered_ring) spmat) \<longrightarrow> sorted_spvec (mult_spmat A B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   368
  apply (induct A)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   369
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   370
  apply (drule sorted_spvec_cons1, simp)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 15197
diff changeset
   371
  apply (case_tac A)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   372
  apply (auto simp add: sorted_spvec.simps)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   373
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   374
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   375
lemma sorted_spmat_mult_spmat[rule_format]:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   376
  "sorted_spmat (B::('a::lordered_ring) spmat) \<longrightarrow> sorted_spmat (mult_spmat A B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   377
  apply (induct A)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   378
  apply (auto simp add: sorted_mult_spvec_spmat) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   379
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   380
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   381
consts
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   382
  add_spvec :: "('a::lordered_ab_group_add) spvec * 'a spvec \<Rightarrow> 'a spvec"
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   383
  add_spmat :: "('a::lordered_ab_group_add) spmat * 'a spmat \<Rightarrow> 'a spmat"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   384
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   385
recdef add_spvec "measure (% (a, b). length a + (length b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   386
  "add_spvec (arr, []) = arr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   387
  "add_spvec ([], brr) = brr"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   388
  "add_spvec (a#arr, b#brr) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   389
  if (fst a) < (fst b) then (a#(add_spvec (arr, b#brr))) 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   390
     else (if (fst b < fst a) then (b#(add_spvec (a#arr, brr)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   391
     else ((fst a, (snd a)+(snd b))#(add_spvec (arr,brr)))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   392
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   393
lemma add_spvec_empty1[simp]: "add_spvec ([], a) = a"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   394
  by (induct a, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   395
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   396
lemma add_spvec_empty2[simp]: "add_spvec (a, []) = a"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   397
  by (induct a, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   398
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   399
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
   400
  apply (rule add_spvec.induct[of _ a b])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   401
  apply (simp_all add: singleton_matrix_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   402
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   403
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   404
recdef add_spmat "measure (% (A,B). (length A)+(length B))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   405
  "add_spmat ([], bs) = bs"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   406
  "add_spmat (as, []) = as"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   407
  "add_spmat (a#as, b#bs) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   408
  if fst a < fst b then 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   409
    (a#(add_spmat (as, b#bs)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   410
  else (if fst b < fst a then
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   411
    (b#(add_spmat (a#as, bs)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   412
  else
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   413
    ((fst a, add_spvec (snd a, snd b))#(add_spmat (as, bs)))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   414
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   415
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
   416
  apply (rule add_spmat.induct)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   417
  apply (auto simp add: sparse_row_matrix_cons sparse_row_vector_add move_matrix_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   418
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   419
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
   420
lemmas [code func] = sparse_row_add_spmat [symmetric]
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
   421
lemmas [code func] = sparse_row_vector_add [symmetric]
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 26300
diff changeset
   422
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   423
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
   424
  proof - 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   425
    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
   426
      by (rule add_spvec.induct[of _ _ brr], auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   427
    then show ?thesis
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   428
      by (case_tac brr, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   429
  qed
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   430
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   431
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
   432
  proof - 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   433
    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
   434
      by (rule add_spmat.induct[of _ _ brr], auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   435
    then show ?thesis
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   436
      by (case_tac brr, auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   437
  qed
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   438
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   439
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
   440
  apply (rule add_spvec.induct[of _ arr brr])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   441
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   442
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   443
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   444
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
   445
  apply (rule add_spmat.induct[of _ arr brr])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   446
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   447
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   448
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   449
lemma add_spvec_commute: "add_spvec (a, b) = add_spvec (b, a)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   450
  by (rule add_spvec.induct[of _ a b], auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   451
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   452
lemma add_spmat_commute: "add_spmat (a, b) = add_spmat (b, a)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   453
  apply (rule add_spmat.induct[of _ a b])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   454
  apply (simp_all add: add_spvec_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   455
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   456
  
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   457
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
   458
  apply (drule sorted_add_spvec_helper1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   459
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   460
  apply (case_tac brr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   461
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   462
  apply (drule_tac sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   463
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   464
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   465
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   466
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
   467
  apply (drule sorted_add_spmat_helper1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   468
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   469
  apply (case_tac brr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   470
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   471
  apply (drule_tac sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   472
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   473
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   474
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   475
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
   476
  apply (rule add_spvec.induct[of _ a b])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   477
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   478
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   479
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   480
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   481
  apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   482
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   483
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   484
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   485
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   486
  apply (simp add: sorted_add_spvec_helper2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   487
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   488
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   489
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   490
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   491
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   492
  apply (frule_tac as=arr in sorted_spvec_cons1, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   493
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   494
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   495
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   496
  apply (simp add: sorted_add_spvec_helper2 add_spvec_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   497
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   498
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   499
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   500
  apply (frule_tac as=arr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   501
  apply (frule_tac as=brr in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   502
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   503
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   504
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   505
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   506
  apply (drule_tac sorted_add_spvec_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   507
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   508
  apply (case_tac arr)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   509
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   510
  apply (drule sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   511
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   512
  apply (case_tac brr)
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
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   517
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   518
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
   519
  apply (rule add_spmat.induct[of _ A B])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   520
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   521
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   522
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   523
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   524
  apply (frule_tac as=bs in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   525
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   526
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   527
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   528
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   529
  apply (simp add: sorted_add_spmat_helper2)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   530
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   531
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   532
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   533
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   534
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   535
  apply (frule_tac as=as in sorted_spvec_cons1, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   536
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   537
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   538
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   539
  apply (simp add: sorted_add_spmat_helper2 add_spmat_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   540
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   541
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   542
  apply (clarify)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   543
  apply (frule_tac as=as in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   544
  apply (frule_tac as=bs in sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   545
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   546
  apply (subst sorted_spvec_step)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   547
  apply (simp split: list.split)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   548
  apply (clarify, simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   549
  apply (drule_tac sorted_add_spmat_helper)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   550
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   551
  apply (case_tac as)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   552
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   553
  apply (drule sorted_spvec_cons3)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   554
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   555
  apply (case_tac bs)
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
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   560
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   561
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
   562
  apply (rule add_spmat.induct[of _ A B])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   563
  apply (simp_all add: sorted_spvec_add_spvec)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   564
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   565
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   566
consts
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   567
  le_spvec :: "('a::lordered_ab_group_add) spvec * 'a spvec \<Rightarrow> bool" 
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   568
  le_spmat :: "('a::lordered_ab_group_add) spmat * 'a spmat \<Rightarrow> bool" 
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   569
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   570
recdef le_spvec "measure (% (a,b). (length a) + (length b))" 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   571
  "le_spvec ([], []) = True"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   572
  "le_spvec (a#as, []) = ((snd a <= 0) & (le_spvec (as, [])))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   573
  "le_spvec ([], b#bs) = ((0 <= snd b) & (le_spvec ([], bs)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   574
  "le_spvec (a#as, b#bs) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   575
  if (fst a < fst b) then 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   576
    ((snd a <= 0) & (le_spvec (as, b#bs)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   577
  else (if (fst b < fst a) then
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   578
    ((0 <= snd b) & (le_spvec (a#as, bs)))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   579
  else 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   580
    ((snd a <= snd b) & (le_spvec (as, bs)))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   581
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   582
recdef le_spmat "measure (% (a,b). (length a) + (length b))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   583
  "le_spmat ([], []) = True"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   584
  "le_spmat (a#as, []) = (le_spvec (snd a, []) & (le_spmat (as, [])))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   585
  "le_spmat ([], b#bs) = (le_spvec ([], snd b) & (le_spmat ([], bs)))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   586
  "le_spmat (a#as, b#bs) = (
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   587
  if fst a < fst b then
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   588
    (le_spvec(snd a,[]) & le_spmat(as, b#bs))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   589
  else (if (fst b < fst a) then 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   590
    (le_spvec([], snd b) & le_spmat(a#as, bs))
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   591
  else
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   592
    (le_spvec(snd a, snd b) & le_spmat (as, bs))))"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   593
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   594
constdefs
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   595
  disj_matrices :: "('a::zero) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   596
  "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
   597
24124
4399175e3014 turned simp_depth_limit into configuration option;
wenzelm
parents: 23477
diff changeset
   598
declare [[simp_depth_limit = 6]]
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   599
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   600
lemma disj_matrices_contr1: "disj_matrices A B \<Longrightarrow> Rep_matrix A j i \<noteq> 0 \<Longrightarrow> Rep_matrix B j i = 0"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   601
   by (simp add: disj_matrices_def)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   602
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   603
lemma disj_matrices_contr2: "disj_matrices A B \<Longrightarrow> Rep_matrix B j i \<noteq> 0 \<Longrightarrow> Rep_matrix A j i = 0"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   604
   by (simp add: disj_matrices_def)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   605
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   606
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   607
lemma disj_matrices_add: "disj_matrices A B \<Longrightarrow> disj_matrices C D \<Longrightarrow> disj_matrices A D \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   608
  (A + B <= C + D) = (A <= C & B <= (D::('a::lordered_ab_group_add) matrix))"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   609
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   610
  apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   611
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   612
  apply (erule conjE)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   613
  apply (drule_tac j=j and i=i in spec2)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   614
  apply (case_tac "Rep_matrix B j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   615
  apply (case_tac "Rep_matrix D j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   616
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   617
  apply (simp (no_asm_use) only: le_matrix_def disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   618
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   619
  apply (erule conjE)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   620
  apply (drule_tac j=j and i=i in spec2)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   621
  apply (case_tac "Rep_matrix A j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   622
  apply (case_tac "Rep_matrix C j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   623
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   624
  apply (erule add_mono)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   625
  apply (assumption)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   626
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   627
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   628
lemma disj_matrices_zero1[simp]: "disj_matrices 0 B"
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_zero2[simp]: "disj_matrices A 0"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   632
by (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_commute: "disj_matrices A B = disj_matrices B A"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   635
by (auto simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   636
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   637
lemma disj_matrices_add_le_zero: "disj_matrices A B \<Longrightarrow>
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   638
  (A + B <= 0) = (A <= 0 & (B::('a::lordered_ab_group_add) matrix) <= 0)"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   639
by (rule disj_matrices_add[of A B 0 0, simplified])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   640
 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   641
lemma disj_matrices_add_zero_le: "disj_matrices A B \<Longrightarrow>
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   642
  (0 <= A + B) = (0 <= A & 0 <= (B::('a::lordered_ab_group_add) matrix))"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   643
by (rule disj_matrices_add[of 0 0 A B, simplified])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   644
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   645
lemma disj_matrices_add_x_le: "disj_matrices A B \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   646
  (A <= B + C) = (A <= C & 0 <= (B::('a::lordered_ab_group_add) matrix))"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   647
by (auto simp add: disj_matrices_add[of 0 A B C, simplified])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   648
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   649
lemma disj_matrices_add_le_x: "disj_matrices A B \<Longrightarrow> disj_matrices B C \<Longrightarrow> 
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   650
  (B + A <= C) = (A <= C &  (B::('a::lordered_ab_group_add) matrix) <= 0)"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   651
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
   652
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   653
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
   654
  apply (simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   655
  apply (rule conjI)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   656
  apply (rule neg_imp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   657
  apply (simp)
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
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   662
  apply (rule sorted_sparse_row_vector_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   663
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   664
  done 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   665
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   666
lemma disj_matrices_x_add: "disj_matrices A B \<Longrightarrow> disj_matrices A C \<Longrightarrow> disj_matrices (A::('a::lordered_ab_group_add) matrix) (B+C)"
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   667
  apply (simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   668
  apply (auto)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   669
  apply (drule_tac j=j and i=i in spec2)+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   670
  apply (case_tac "Rep_matrix B j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   671
  apply (case_tac "Rep_matrix C j i = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   672
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   673
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   674
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   675
lemma disj_matrices_add_x: "disj_matrices A B \<Longrightarrow> disj_matrices A C \<Longrightarrow> disj_matrices (B+C) (A::('a::lordered_ab_group_add) matrix)" 
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   676
  by (simp add: disj_matrices_x_add disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   677
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   678
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
   679
  by (auto simp add: disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   680
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   681
lemma disj_move_sparse_vec_mat[simplified disj_matrices_commute]: 
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   682
  "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
   683
  apply (auto simp add: neg_def disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   684
  apply (drule nrows_notzero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   685
  apply (drule less_le_trans[OF _ nrows_spvec])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   686
  apply (subgoal_tac "ja = j")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   687
  apply (simp add: sorted_sparse_row_matrix_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   688
  apply (arith)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   689
  apply (rule nrows)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   690
  apply (rule order_trans[of _ 1 _])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   691
  apply (simp)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   692
  apply (case_tac "nat (int ja - int j) = 0")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   693
  apply (case_tac "ja = j")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   694
  apply (simp add: sorted_sparse_row_matrix_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   695
  apply arith+
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   696
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   697
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   698
lemma disj_move_sparse_row_vector_twice:
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   699
  "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
   700
  apply (auto simp add: neg_def disj_matrices_def)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   701
  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
   702
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   703
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   704
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
   705
  apply (rule le_spvec.induct)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   706
  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
   707
    disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   708
  apply (rule conjI, intro strip)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   709
  apply (simp add: sorted_spvec_cons1)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   710
  apply (subst disj_matrices_add_x_le)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   711
  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
   712
  apply (simp add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   713
  apply (simp, blast)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   714
  apply (intro strip, rule conjI, intro strip)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   715
  apply (simp add: sorted_spvec_cons1)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   716
  apply (subst disj_matrices_add_le_x)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   717
  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
   718
  apply (blast)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   719
  apply (intro strip)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   720
  apply (simp add: sorted_spvec_cons1)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   721
  apply (case_tac "a=aa", simp_all)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   722
  apply (subst disj_matrices_add)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   723
  apply (simp_all add: disj_sparse_row_singleton[OF order_refl] disj_matrices_commute)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   724
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   725
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   726
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
   727
  apply (induct b)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   728
  apply (simp_all add: sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   729
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   730
  apply (subst disj_matrices_add_le_zero)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   731
  apply (simp add: disj_matrices_commute disj_sparse_row_singleton sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   732
  apply (rule_tac y = "snd a" in disj_sparse_row_singleton[OF order_refl])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   733
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   734
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   735
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   736
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
   737
  apply (induct b)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   738
  apply (simp_all add: sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   739
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   740
  apply (subst disj_matrices_add_zero_le)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   741
  apply (simp add: disj_matrices_commute disj_sparse_row_singleton sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   742
  apply (rule_tac y = "snd a" in disj_sparse_row_singleton[OF order_refl])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   743
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   744
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   745
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   746
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
   747
  le_spmat(A, B) = (sparse_row_matrix A <= sparse_row_matrix B)"
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   748
  apply (rule le_spmat.induct)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   749
  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
   750
    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
   751
  apply (rule conjI, intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   752
  apply (simp add: sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   753
  apply (subst disj_matrices_add_x_le)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   754
  apply (rule disj_matrices_add_x)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   755
  apply (simp add: disj_move_sparse_row_vector_twice)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   756
  apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   757
  apply (simp add: disj_move_sparse_vec_mat[OF order_refl] disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   758
  apply (simp, blast)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   759
  apply (intro strip, rule conjI, intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   760
  apply (simp add: sorted_spvec_cons1)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   761
  apply (subst disj_matrices_add_le_x)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   762
  apply (simp add: disj_move_sparse_vec_mat[OF order_refl])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   763
  apply (rule disj_matrices_x_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   764
  apply (simp add: disj_move_sparse_row_vector_twice)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   765
  apply (simp add: disj_move_sparse_vec_mat[OF less_imp_le] disj_matrices_commute)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   766
  apply (simp, blast)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   767
  apply (intro strip)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   768
  apply (case_tac "a=aa")
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   769
  apply (simp_all)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   770
  apply (subst disj_matrices_add)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   771
  apply (simp_all add: disj_matrices_commute disj_move_sparse_vec_mat[OF order_refl])
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   772
  apply (simp add: sorted_spvec_cons1 le_spvec_iff_sparse_row_le)
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   773
  done
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   774
24124
4399175e3014 turned simp_depth_limit into configuration option;
wenzelm
parents: 23477
diff changeset
   775
declare [[simp_depth_limit = 999]]
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   776
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   777
consts 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   778
   abs_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   779
   minus_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   780
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   781
primrec
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   782
  "abs_spmat [] = []"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   783
  "abs_spmat (a#as) = (fst a, abs_spvec (snd a))#(abs_spmat as)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   784
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   785
primrec
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   786
  "minus_spmat [] = []"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   787
  "minus_spmat (a#as) = (fst a, minus_spvec (snd a))#(minus_spmat as)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   788
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   789
lemma sparse_row_matrix_minus:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   790
  "sparse_row_matrix (minus_spmat A) = - (sparse_row_matrix A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   791
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   792
  apply (simp_all add: sparse_row_vector_minus sparse_row_matrix_cons)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   793
  apply (subst Rep_matrix_inject[symmetric])
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   794
  apply (rule ext)+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   795
  apply simp
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   796
  done
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   797
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   798
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
   799
proof -
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   800
  assume x:"x \<noteq> 0"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   801
  have r:"nrows (sparse_row_vector v) <= Suc 0" by (rule nrows_spvec)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   802
  show ?thesis
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   803
    apply (rule nrows)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   804
    apply (subgoal_tac "Suc 0 <= x")
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   805
    apply (insert r)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   806
    apply (simp only:)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   807
    apply (insert x)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   808
    apply arith
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   809
    done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   810
qed
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   811
    
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   812
lemma sparse_row_matrix_abs:
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   813
  "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
   814
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   815
  apply (simp_all add: sparse_row_vector_abs sparse_row_matrix_cons)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   816
  apply (frule_tac sorted_spvec_cons1, simp)
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   817
  apply (simplesubst Rep_matrix_inject[symmetric])
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   818
  apply (rule ext)+
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   819
  apply auto
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   820
  apply (case_tac "x=a")
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   821
  apply (simp)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15236
diff changeset
   822
  apply (simplesubst sorted_sparse_row_matrix_zero)
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   823
  apply auto
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15236
diff changeset
   824
  apply (simplesubst Rep_sparse_row_vector_zero)
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   825
  apply (simp_all add: neg_def)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   826
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   827
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   828
lemma sorted_spvec_minus_spmat: "sorted_spvec A \<Longrightarrow> sorted_spvec (minus_spmat A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   829
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   830
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   831
  apply (frule sorted_spvec_cons1, simp)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 15197
diff changeset
   832
  apply (simp add: sorted_spvec.simps split:list.split_asm)
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   833
  done 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   834
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   835
lemma sorted_spvec_abs_spmat: "sorted_spvec A \<Longrightarrow> sorted_spvec (abs_spmat A)" 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   836
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   837
  apply (simp)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   838
  apply (frule sorted_spvec_cons1, simp)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 15197
diff changeset
   839
  apply (simp add: sorted_spvec.simps split:list.split_asm)
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   840
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   841
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   842
lemma sorted_spmat_minus_spmat: "sorted_spmat A \<Longrightarrow> sorted_spmat (minus_spmat A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   843
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   844
  apply (simp_all add: sorted_spvec_minus_spvec)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   845
  done
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   846
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   847
lemma sorted_spmat_abs_spmat: "sorted_spmat A \<Longrightarrow> sorted_spmat (abs_spmat A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   848
  apply (induct A)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   849
  apply (simp_all add: sorted_spvec_abs_spvec)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   850
  done
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
   851
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   852
constdefs
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   853
  diff_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   854
  "diff_spmat A B == add_spmat (A, minus_spmat B)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   855
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   856
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
   857
  by (simp add: diff_spmat_def sorted_spmat_minus_spmat sorted_spmat_add_spmat)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   858
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   859
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
   860
  by (simp add: diff_spmat_def sorted_spvec_minus_spmat sorted_spvec_add_spmat)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   861
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   862
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
   863
  by (simp add: diff_spmat_def sparse_row_add_spmat sparse_row_matrix_minus)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   864
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   865
constdefs
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   866
  sorted_sparse_matrix :: "'a spmat \<Rightarrow> bool"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   867
  "sorted_sparse_matrix A == (sorted_spvec A) & (sorted_spmat A)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   868
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   869
lemma sorted_sparse_matrix_imp_spvec: "sorted_sparse_matrix A \<Longrightarrow> sorted_spvec A"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   870
  by (simp add: sorted_sparse_matrix_def)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   871
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   872
lemma sorted_sparse_matrix_imp_spmat: "sorted_sparse_matrix A \<Longrightarrow> sorted_spmat A"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   873
  by (simp add: sorted_sparse_matrix_def)
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   874
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   875
lemmas sorted_sp_simps = 
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   876
  sorted_spvec.simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   877
  sorted_spmat.simps
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   878
  sorted_sparse_matrix_def
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   879
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   880
lemma bool1: "(\<not> True) = False"  by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   881
lemma bool2: "(\<not> False) = True"  by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   882
lemma bool3: "((P\<Colon>bool) \<and> True) = P" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   883
lemma bool4: "(True \<and> (P\<Colon>bool)) = P" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   884
lemma bool5: "((P\<Colon>bool) \<and> False) = False" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   885
lemma bool6: "(False \<and> (P\<Colon>bool)) = False" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   886
lemma bool7: "((P\<Colon>bool) \<or> True) = True" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   887
lemma bool8: "(True \<or> (P\<Colon>bool)) = True" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   888
lemma bool9: "((P\<Colon>bool) \<or> False) = P" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   889
lemma bool10: "(False \<or> (P\<Colon>bool)) = P" by blast
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   890
lemmas boolarith = bool1 bool2 bool3 bool4 bool5 bool6 bool7 bool8 bool9 bool10
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   891
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
   892
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
   893
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   894
consts
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   895
  pprt_spvec :: "('a::{lordered_ab_group_add}) spvec \<Rightarrow> 'a spvec"
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   896
  nprt_spvec :: "('a::{lordered_ab_group_add}) spvec \<Rightarrow> 'a spvec"
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   897
  pprt_spmat :: "('a::{lordered_ab_group_add}) spmat \<Rightarrow> 'a spmat"
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 24124
diff changeset
   898
  nprt_spmat :: "('a::{lordered_ab_group_add}) spmat \<Rightarrow> 'a spmat"
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   899
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   900
primrec
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   901
  "pprt_spvec [] = []"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   902
  "pprt_spvec (a#as) = (fst a, pprt (snd a)) # (pprt_spvec as)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   903
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   904
primrec
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   905
  "nprt_spvec [] = []"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   906
  "nprt_spvec (a#as) = (fst a, nprt (snd a)) # (nprt_spvec as)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   907
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   908
primrec 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   909
  "pprt_spmat [] = []"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   910
  "pprt_spmat (a#as) = (fst a, pprt_spvec (snd a))#(pprt_spmat as)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   911
  (*case (pprt_spvec (snd a)) of [] \<Rightarrow> (pprt_spmat as) | y#ys \<Rightarrow> (fst a, y#ys)#(pprt_spmat as))"*)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   912
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   913
primrec 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   914
  "nprt_spmat [] = []"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   915
  "nprt_spmat (a#as) = (fst a, nprt_spvec (snd a))#(nprt_spmat as)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   916
  (*case (nprt_spvec (snd a)) of [] \<Rightarrow> (nprt_spmat as) | y#ys \<Rightarrow> (fst a, y#ys)#(nprt_spmat as))"*)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   917
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   918
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   919
lemma pprt_add: "disj_matrices A (B::(_::lordered_ring) matrix) \<Longrightarrow> pprt (A+B) = pprt A + pprt B"
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 20432
diff changeset
   920
  apply (simp add: pprt_def sup_matrix_def)
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   921
  apply (simp add: Rep_matrix_inject[symmetric])
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   922
  apply (rule ext)+
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   923
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   924
  apply (case_tac "Rep_matrix A x xa \<noteq> 0")
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   925
  apply (simp_all add: disj_matrices_contr1)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   926
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   927
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   928
lemma nprt_add: "disj_matrices A (B::(_::lordered_ring) matrix) \<Longrightarrow> nprt (A+B) = nprt A + nprt B"
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 20432
diff changeset
   929
  apply (simp add: nprt_def inf_matrix_def)
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   930
  apply (simp add: Rep_matrix_inject[symmetric])
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   931
  apply (rule ext)+
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   932
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   933
  apply (case_tac "Rep_matrix A x xa \<noteq> 0")
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   934
  apply (simp_all add: disj_matrices_contr1)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   935
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   936
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   937
lemma pprt_singleton[simp]: "pprt (singleton_matrix j i (x::_::lordered_ring)) = singleton_matrix j i (pprt x)"
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 20432
diff changeset
   938
  apply (simp add: pprt_def sup_matrix_def)
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   939
  apply (simp add: Rep_matrix_inject[symmetric])
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   940
  apply (rule ext)+
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   941
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   942
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   943
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   944
lemma nprt_singleton[simp]: "nprt (singleton_matrix j i (x::_::lordered_ring)) = singleton_matrix j i (nprt x)"
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 20432
diff changeset
   945
  apply (simp add: nprt_def inf_matrix_def)
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   946
  apply (simp add: Rep_matrix_inject[symmetric])
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   947
  apply (rule ext)+
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   948
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   949
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   950
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   951
lemma less_imp_le: "a < b \<Longrightarrow> a <= (b::_::order)" by (simp add: less_def)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   952
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   953
lemma sparse_row_vector_pprt: "sorted_spvec (v :: 'a::lordered_ring spvec) \<Longrightarrow> sparse_row_vector (pprt_spvec v) = pprt (sparse_row_vector v)"
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   954
  apply (induct v)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   955
  apply (simp_all)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   956
  apply (frule sorted_spvec_cons1, auto)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   957
  apply (subst pprt_add)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   958
  apply (subst disj_matrices_commute)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   959
  apply (rule disj_sparse_row_singleton)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   960
  apply auto
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   961
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   962
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   963
lemma sparse_row_vector_nprt: "sorted_spvec (v :: 'a::lordered_ring spvec) \<Longrightarrow> sparse_row_vector (nprt_spvec v) = nprt (sparse_row_vector v)"
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   964
  apply (induct v)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   965
  apply (simp_all)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   966
  apply (frule sorted_spvec_cons1, auto)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   967
  apply (subst nprt_add)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   968
  apply (subst disj_matrices_commute)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   969
  apply (rule disj_sparse_row_singleton)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   970
  apply auto
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   971
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   972
  
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   973
  
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   974
lemma pprt_move_matrix: "pprt (move_matrix (A::('a::lordered_ring) matrix) j i) = move_matrix (pprt A) j i"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   975
  apply (simp add: pprt_def)
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 20432
diff changeset
   976
  apply (simp add: sup_matrix_def)
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   977
  apply (simp add: Rep_matrix_inject[symmetric])
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   978
  apply (rule ext)+
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   979
  apply (simp)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   980
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   981
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   982
lemma nprt_move_matrix: "nprt (move_matrix (A::('a::lordered_ring) matrix) j i) = move_matrix (nprt A) j i"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   983
  apply (simp add: nprt_def)
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 20432
diff changeset
   984
  apply (simp add: inf_matrix_def)
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   985
  apply (simp add: Rep_matrix_inject[symmetric])
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   986
  apply (rule ext)+
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   987
  apply (simp)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   988
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   989
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
   990
lemma sparse_row_matrix_pprt: "sorted_spvec (m :: 'a::lordered_ring spmat) \<Longrightarrow> sorted_spmat m \<Longrightarrow> sparse_row_matrix (pprt_spmat m) = pprt (sparse_row_matrix m)"
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   991
  apply (induct m)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   992
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   993
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   994
  apply (frule sorted_spvec_cons1)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   995
  apply (simp add: sparse_row_matrix_cons sparse_row_vector_pprt)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   996
  apply (subst pprt_add)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   997
  apply (subst disj_matrices_commute)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   998
  apply (rule disj_move_sparse_vec_mat)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
   999
  apply auto
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1000
  apply (simp add: sorted_spvec.simps)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1001
  apply (simp split: list.split)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1002
  apply auto
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1003
  apply (simp add: pprt_move_matrix)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1004
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1005
27653
180e28bab764 more class instantiations
haftmann
parents: 27484
diff changeset
  1006
lemma sparse_row_matrix_nprt: "sorted_spvec (m :: 'a::lordered_ring spmat) \<Longrightarrow> sorted_spmat m \<Longrightarrow> sparse_row_matrix (nprt_spmat m) = nprt (sparse_row_matrix m)"
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1007
  apply (induct m)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1008
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1009
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1010
  apply (frule sorted_spvec_cons1)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1011
  apply (simp add: sparse_row_matrix_cons sparse_row_vector_nprt)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1012
  apply (subst nprt_add)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1013
  apply (subst disj_matrices_commute)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1014
  apply (rule disj_move_sparse_vec_mat)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1015
  apply auto
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1016
  apply (simp add: sorted_spvec.simps)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1017
  apply (simp split: list.split)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1018
  apply auto
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1019
  apply (simp add: nprt_move_matrix)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1020
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1021
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1022
lemma sorted_pprt_spvec: "sorted_spvec v \<Longrightarrow> sorted_spvec (pprt_spvec v)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1023
  apply (induct v)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1024
  apply (simp)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1025
  apply (frule sorted_spvec_cons1)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1026
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1027
  apply (simp add: sorted_spvec.simps split:list.split_asm)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1028
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1029
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1030
lemma sorted_nprt_spvec: "sorted_spvec v \<Longrightarrow> sorted_spvec (nprt_spvec v)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1031
  apply (induct v)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1032
  apply (simp)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1033
  apply (frule sorted_spvec_cons1)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1034
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1035
  apply (simp add: sorted_spvec.simps split:list.split_asm)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1036
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1037
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1038
lemma sorted_spvec_pprt_spmat: "sorted_spvec m \<Longrightarrow> sorted_spvec (pprt_spmat m)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1039
  apply (induct m)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1040
  apply (simp)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1041
  apply (frule sorted_spvec_cons1)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1042
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1043
  apply (simp add: sorted_spvec.simps split:list.split_asm)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1044
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1045
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1046
lemma sorted_spvec_nprt_spmat: "sorted_spvec m \<Longrightarrow> sorted_spvec (nprt_spmat m)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1047
  apply (induct m)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1048
  apply (simp)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1049
  apply (frule sorted_spvec_cons1)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1050
  apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1051
  apply (simp add: sorted_spvec.simps split:list.split_asm)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1052
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1053
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1054
lemma sorted_spmat_pprt_spmat: "sorted_spmat m \<Longrightarrow> sorted_spmat (pprt_spmat m)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1055
  apply (induct m)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1056
  apply (simp_all add: sorted_pprt_spvec)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1057
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1058
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1059
lemma sorted_spmat_nprt_spmat: "sorted_spmat m \<Longrightarrow> sorted_spmat (nprt_spmat m)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1060
  apply (induct m)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1061
  apply (simp_all add: sorted_nprt_spvec)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1062
  done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1063
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1064
constdefs
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1065
  mult_est_spmat :: "('a::lordered_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1066
  "mult_est_spmat r1 r2 s1 s2 == 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1067
  add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2), add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2), 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1068
  add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1), mult_spmat (nprt_spmat s1) (nprt_spmat r1))))"  
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1069
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1070
lemmas sparse_row_matrix_op_simps =
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1071
  sorted_sparse_matrix_imp_spmat sorted_sparse_matrix_imp_spvec
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1072
  sparse_row_add_spmat sorted_spvec_add_spmat sorted_spmat_add_spmat
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1073
  sparse_row_diff_spmat sorted_spvec_diff_spmat sorted_spmat_diff_spmat
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1074
  sparse_row_matrix_minus sorted_spvec_minus_spmat sorted_spmat_minus_spmat
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1075
  sparse_row_mult_spmat sorted_spvec_mult_spmat sorted_spmat_mult_spmat
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1076
  sparse_row_matrix_abs sorted_spvec_abs_spmat sorted_spmat_abs_spmat
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1077
  le_spmat_iff_sparse_row_le
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1078
  sparse_row_matrix_pprt sorted_spvec_pprt_spmat sorted_spmat_pprt_spmat
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1079
  sparse_row_matrix_nprt sorted_spvec_nprt_spmat sorted_spmat_nprt_spmat
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1080
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1081
lemma zero_eq_Numeral0: "(0::_::number_ring) = Numeral0" by simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1082
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1083
lemmas sparse_row_matrix_arith_simps[simplified zero_eq_Numeral0] = 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1084
  mult_spmat.simps mult_spvec_spmat.simps 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1085
  addmult_spvec.simps 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1086
  smult_spvec_empty smult_spvec_cons
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1087
  add_spmat.simps add_spvec.simps
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1088
  minus_spmat.simps minus_spvec.simps
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1089
  abs_spmat.simps abs_spvec.simps
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1090
  diff_spmat_def
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1091
  le_spmat.simps le_spvec.simps
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1092
  pprt_spmat.simps pprt_spvec.simps
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1093
  nprt_spmat.simps nprt_spvec.simps
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1094
  mult_est_spmat_def
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1095
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1096
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1097
(*lemma spm_linprog_dual_estimate_1:
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1098
  assumes  
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1099
  "sorted_sparse_matrix A1"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1100
  "sorted_sparse_matrix A2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1101
  "sorted_sparse_matrix c1"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1102
  "sorted_sparse_matrix c2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1103
  "sorted_sparse_matrix y"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1104
  "sorted_spvec b"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1105
  "sorted_spvec r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1106
  "le_spmat ([], y)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1107
  "A * x \<le> sparse_row_matrix (b::('a::lordered_ring) spmat)"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1108
  "sparse_row_matrix A1 <= A"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1109
  "A <= sparse_row_matrix A2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1110
  "sparse_row_matrix c1 <= c"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1111
  "c <= sparse_row_matrix c2"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1112
  "abs x \<le> sparse_row_matrix r"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1113
  shows
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1114
  "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
  1115
  abs_spmat (diff_spmat (mult_spmat y A1) c1)), diff_spmat c2 c1)) r))"
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15009
diff changeset
  1116
  by (insert prems, simp add: sparse_row_matrix_op_simps linprog_dual_estimate_1[where A=A])
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  1117
*)
15009
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
  1118
8c89f588c7aa support for sparse matrices
obua
parents:
diff changeset
  1119
end