src/HOL/Matrix_LP/Matrix.thy
author nipkow
Mon, 24 Sep 2018 14:30:09 +0200
changeset 69064 5840724b1d71
parent 67613 ce654b0e6d69
child 73463 552a9dd5b4a2
permissions -rw-r--r--
Prefix form of infix with * on either side no longer needs special treatment because (* and *) are no longer comment brackets in terms.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47455
26315a545e26 updated headers;
wenzelm
parents: 46988
diff changeset
     1
(*  Title:      HOL/Matrix_LP/Matrix.thy
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     2
    Author:     Steven Obua
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     3
*)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     4
17915
e38947f9ba5e isatool fixheaders;
wenzelm
parents: 16733
diff changeset
     5
theory Matrix
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63566
diff changeset
     6
imports Main "HOL-Library.Lattice_Algebras"
17915
e38947f9ba5e isatool fixheaders;
wenzelm
parents: 16733
diff changeset
     7
begin
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
     8
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 41413
diff changeset
     9
type_synonym 'a infmatrix = "nat \<Rightarrow> nat \<Rightarrow> 'a"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    10
44174
d1d79f0e1ea6 make more HOL theories work with separate set type
huffman
parents: 42463
diff changeset
    11
definition nonzero_positions :: "(nat \<Rightarrow> nat \<Rightarrow> 'a::zero) \<Rightarrow> (nat \<times> nat) set" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    12
  "nonzero_positions A = {pos. A (fst pos) (snd pos) ~= 0}"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    13
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44174
diff changeset
    14
definition "matrix = {(f::(nat \<Rightarrow> nat \<Rightarrow> 'a::zero)). finite (nonzero_positions f)}"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44174
diff changeset
    15
61260
e6f03fae14d5 explicit indication of overloaded typedefs;
wenzelm
parents: 61169
diff changeset
    16
typedef (overloaded) 'a matrix = "matrix :: (nat \<Rightarrow> nat \<Rightarrow> 'a::zero) set"
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44174
diff changeset
    17
  unfolding matrix_def
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44174
diff changeset
    18
proof
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44174
diff changeset
    19
  show "(\<lambda>j i. 0) \<in> {(f::(nat \<Rightarrow> nat \<Rightarrow> 'a::zero)). finite (nonzero_positions f)}"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    20
    by (simp add: nonzero_positions_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    21
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    22
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    23
declare Rep_matrix_inverse[simp]
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    24
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    25
lemma finite_nonzero_positions : "finite (nonzero_positions (Rep_matrix A))"
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44174
diff changeset
    26
  by (induct A) (simp add: Abs_matrix_inverse matrix_def)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    27
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
    28
definition nrows :: "('a::zero) matrix \<Rightarrow> nat" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    29
  "nrows A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image fst) (nonzero_positions (Rep_matrix A))))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
    30
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
    31
definition ncols :: "('a::zero) matrix \<Rightarrow> nat" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    32
  "ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    33
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    34
lemma nrows:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    35
  assumes hyp: "nrows A \<le> m"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
    36
  shows "(Rep_matrix A m n) = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    37
proof cases
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    38
  assume "nonzero_positions(Rep_matrix A) = {}"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    39
  then show "(Rep_matrix A m n) = 0" by (simp add: nonzero_positions_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    40
next
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    41
  assume a: "nonzero_positions(Rep_matrix A) \<noteq> {}"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    42
  let ?S = "fst`(nonzero_positions(Rep_matrix A))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    43
  have c: "finite (?S)" by (simp add: finite_nonzero_positions)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    44
  from hyp have d: "Max (?S) < m" by (simp add: a nrows_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    45
  have "m \<notin> ?S"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    46
    proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    47
      have "m \<in> ?S \<Longrightarrow> m <= Max(?S)" by (simp add: Max_ge [OF c])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    48
      moreover from d have "~(m <= Max ?S)" by (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    49
      ultimately show "m \<notin> ?S" by (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    50
    qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    51
  thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    52
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    53
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
    54
definition transpose_infmatrix :: "'a infmatrix \<Rightarrow> 'a infmatrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    55
  "transpose_infmatrix A j i == A i j"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
    56
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
    57
definition transpose_matrix :: "('a::zero) matrix \<Rightarrow> 'a matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    58
  "transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    59
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    60
declare transpose_infmatrix_def[simp]
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    61
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    62
lemma transpose_infmatrix_twice[simp]: "transpose_infmatrix (transpose_infmatrix A) = A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    63
by ((rule ext)+, simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    64
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    65
lemma transpose_infmatrix: "transpose_infmatrix (% j i. P j i) = (% j i. P i j)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    66
  apply (rule ext)+
46985
bd955d9f464b tuned proofs;
wenzelm
parents: 46702
diff changeset
    67
  by simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    68
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    69
lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    70
apply (rule Abs_matrix_inverse)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    71
apply (simp add: matrix_def nonzero_positions_def image_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    72
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    73
  let ?A = "{pos. Rep_matrix x (snd pos) (fst pos) \<noteq> 0}"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    74
  let ?swap = "% pos. (snd pos, fst pos)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    75
  let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \<noteq> 0}"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    76
  have swap_image: "?swap`?A = ?B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    77
    apply (simp add: image_def)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 38526
diff changeset
    78
    apply (rule set_eqI)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    79
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    80
    proof
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    81
      fix y
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    82
      assume hyp: "\<exists>a b. Rep_matrix x b a \<noteq> 0 \<and> y = (b, a)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    83
      thus "Rep_matrix x (fst y) (snd y) \<noteq> 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    84
        proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    85
          from hyp obtain a b where "(Rep_matrix x b a \<noteq> 0 & y = (b,a))" by blast
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    86
          then show "Rep_matrix x (fst y) (snd y) \<noteq> 0" by (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    87
        qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    88
    next
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    89
      fix y
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    90
      assume hyp: "Rep_matrix x (fst y) (snd y) \<noteq> 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    91
      show "\<exists> a b. (Rep_matrix x b a \<noteq> 0 & y = (b,a))"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32491
diff changeset
    92
        by (rule exI[of _ "snd y"], rule exI[of _ "fst y"]) (simp add: hyp)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    93
    qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    94
  then have "finite (?swap`?A)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    95
    proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    96
      have "finite (nonzero_positions (Rep_matrix x))" by (simp add: finite_nonzero_positions)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    97
      then have "finite ?B" by (simp add: nonzero_positions_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    98
      with swap_image show "finite (?swap`?A)" by (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
    99
    qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   100
  moreover
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   101
  have "inj_on ?swap ?A" by (simp add: inj_on_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   102
  ultimately show "finite ?A"by (rule finite_imageD[of ?swap ?A])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   103
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   104
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   105
lemma infmatrixforward: "(x::'a infmatrix) = y \<Longrightarrow> \<forall> a b. x a b = y a b" by auto
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   106
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   107
lemma transpose_infmatrix_inject: "(transpose_infmatrix A = transpose_infmatrix B) = (A = B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   108
apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   109
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   110
apply (simp add: transpose_infmatrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   111
apply (drule infmatrixforward)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   112
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   113
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   114
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   115
lemma transpose_matrix_inject: "(transpose_matrix A = transpose_matrix B) = (A = B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   116
apply (simp add: transpose_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   117
apply (subst Rep_matrix_inject[THEN sym])+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   118
apply (simp only: transpose_infmatrix_closed transpose_infmatrix_inject)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   119
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   120
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   121
lemma transpose_matrix[simp]: "Rep_matrix(transpose_matrix A) j i = Rep_matrix A i j"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   122
by (simp add: transpose_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   123
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   124
lemma transpose_transpose_id[simp]: "transpose_matrix (transpose_matrix A) = A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   125
by (simp add: transpose_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   126
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   127
lemma nrows_transpose[simp]: "nrows (transpose_matrix A) = ncols A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   128
by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   129
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   130
lemma ncols_transpose[simp]: "ncols (transpose_matrix A) = nrows A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   131
by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   132
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   133
lemma ncols: "ncols A <= n \<Longrightarrow> Rep_matrix A m n = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   134
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   135
  assume "ncols A <= n"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   136
  then have "nrows (transpose_matrix A) <= n" by (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   137
  then have "Rep_matrix (transpose_matrix A) n m = 0" by (rule nrows)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   138
  thus "Rep_matrix A m n = 0" by (simp add: transpose_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   139
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   140
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   141
lemma ncols_le: "(ncols A <= n) = (\<forall>j i. n <= i \<longrightarrow> (Rep_matrix A j i) = 0)" (is "_ = ?st")
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   142
apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   143
apply (simp add: ncols)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   144
proof (simp add: ncols_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   145
  let ?P = "nonzero_positions (Rep_matrix A)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   146
  let ?p = "snd`?P"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   147
  have a:"finite ?p" by (simp add: finite_nonzero_positions)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   148
  let ?m = "Max ?p"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   149
  assume "~(Suc (?m) <= n)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   150
  then have b:"n <= ?m" by (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   151
  fix a b
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   152
  assume "(a,b) \<in> ?P"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   153
  then have "?p \<noteq> {}" by (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   154
  with a have "?m \<in>  ?p" by (simp)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   155
  moreover have "\<forall>x. (x \<in> ?p \<longrightarrow> (\<exists>y. (Rep_matrix A y x) \<noteq> 0))" by (simp add: nonzero_positions_def image_def)
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   156
  ultimately have "\<exists>y. (Rep_matrix A y ?m) \<noteq> 0" by (simp)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   157
  moreover assume ?st
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   158
  ultimately show "False" using b by (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   159
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   160
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   161
lemma less_ncols: "(n < ncols A) = (\<exists>j i. n <= i & (Rep_matrix A j i) \<noteq> 0)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   162
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   163
  have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   164
  show ?thesis by (simp add: a ncols_le)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   165
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   166
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   167
lemma le_ncols: "(n <= ncols A) = (\<forall> m. (\<forall> j i. m <= i \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n <= m)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   168
apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   169
apply (subgoal_tac "ncols A <= m")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   170
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   171
apply (simp add: ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   172
apply (drule_tac x="ncols A" in spec)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   173
by (simp add: ncols)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   174
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   175
lemma nrows_le: "(nrows A <= n) = (\<forall>j i. n <= j \<longrightarrow> (Rep_matrix A j i) = 0)" (is ?s)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   176
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   177
  have "(nrows A <= n) = (ncols (transpose_matrix A) <= n)" by (simp)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   178
  also have "\<dots> = (\<forall>j i. n <= i \<longrightarrow> (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le)
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   179
  also have "\<dots> = (\<forall>j i. n <= i \<longrightarrow> (Rep_matrix A i j) = 0)" by (simp)
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   180
  finally show "(nrows A <= n) = (\<forall>j i. n <= j \<longrightarrow> (Rep_matrix A j i) = 0)" by (auto)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   181
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   182
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   183
lemma less_nrows: "(m < nrows A) = (\<exists>j i. m <= j & (Rep_matrix A j i) \<noteq> 0)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   184
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   185
  have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   186
  show ?thesis by (simp add: a nrows_le)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   187
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   188
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   189
lemma le_nrows: "(n <= nrows A) = (\<forall> m. (\<forall> j i. m <= j \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n <= m)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   190
apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   191
apply (subgoal_tac "nrows A <= m")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   192
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   193
apply (simp add: nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   194
apply (drule_tac x="nrows A" in spec)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   195
by (simp add: nrows)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   196
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   197
lemma nrows_notzero: "Rep_matrix A m n \<noteq> 0 \<Longrightarrow> m < nrows A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   198
apply (case_tac "nrows A <= m")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   199
apply (simp_all add: nrows)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   200
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   201
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   202
lemma ncols_notzero: "Rep_matrix A m n \<noteq> 0 \<Longrightarrow> n < ncols A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   203
apply (case_tac "ncols A <= n")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   204
apply (simp_all add: ncols)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   205
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   206
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   207
lemma finite_natarray1: "finite {x. x < (n::nat)}"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   208
apply (induct n)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   209
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   210
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   211
  fix n
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 38526
diff changeset
   212
  have "{x. x < Suc n} = insert n {x. x < n}"  by (rule set_eqI, simp, arith)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   213
  moreover assume "finite {x. x < n}"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   214
  ultimately show "finite {x. x < Suc n}" by (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   215
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   216
50027
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49834
diff changeset
   217
lemma finite_natarray2: "finite {(x, y). x < (m::nat) & y < (n::nat)}"
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49834
diff changeset
   218
by simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   219
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   220
lemma RepAbs_matrix:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   221
  assumes aem: "\<exists>m. \<forall>j i. m <= j \<longrightarrow> x j i = 0" (is ?em) and aen:"\<exists>n. \<forall>j i. (n <= i \<longrightarrow> x j i = 0)" (is ?en)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   222
  shows "(Rep_matrix (Abs_matrix x)) = x"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   223
apply (rule Abs_matrix_inverse)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   224
apply (simp add: matrix_def nonzero_positions_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   225
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   226
  from aem obtain m where a: "\<forall>j i. m <= j \<longrightarrow> x j i = 0" by (blast)
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   227
  from aen obtain n where b: "\<forall>j i. n <= i \<longrightarrow> x j i = 0" by (blast)
50027
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49834
diff changeset
   228
  let ?u = "{(i, j). x i j \<noteq> 0}"
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49834
diff changeset
   229
  let ?v = "{(i, j). i < m & j < n}"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   230
  have c: "!! (m::nat) a. ~(m <= a) \<Longrightarrow> a < m" by (arith)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   231
  from a b have "(?u \<inter> (-?v)) = {}"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   232
    apply (simp)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 38526
diff changeset
   233
    apply (rule set_eqI)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   234
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   235
    apply auto
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   236
    by (rule c, auto)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   237
  then have d: "?u \<subseteq> ?v" by blast
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   238
  moreover have "finite ?v" by (simp add: finite_natarray2)
50027
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49834
diff changeset
   239
  moreover have "{pos. x (fst pos) (snd pos) \<noteq> 0} = ?u" by auto
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49834
diff changeset
   240
  ultimately show "finite {pos. x (fst pos) (snd pos) \<noteq> 0}"
7747a9f4c358 adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
bulwahn
parents: 49834
diff changeset
   241
    by (metis (lifting) finite_subset)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   242
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   243
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   244
definition apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   245
  "apply_infmatrix f == % A. (% j i. f (A j i))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   246
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   247
definition apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   248
  "apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   249
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   250
definition combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   251
  "combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   252
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   253
definition combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('c::zero) matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   254
  "combine_matrix f == % A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   255
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   256
lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   257
by (simp add: apply_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   258
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   259
lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   260
by (simp add: combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   261
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   262
definition commutative :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   263
"commutative f == \<forall>x y. f x y = f y x"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   264
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   265
definition associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   266
"associative f == \<forall>x y z. f (f x y) z = f x (f y z)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   267
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   268
text\<open>
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   269
To reason about associativity and commutativity of operations on matrices,
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   270
let's take a step back and look at the general situtation: Assume that we have
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   271
sets $A$ and $B$ with $B \subset A$ and an abstraction $u: A \rightarrow B$. This abstraction has to fulfill $u(b) = b$ for all $b \in B$, but is arbitrary otherwise.
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   272
Each function $f: A \times A \rightarrow A$ now induces a function $f': B \times B \rightarrow B$ by $f' = u \circ f$.
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   273
It is obvious that commutativity of $f$ implies commutativity of $f'$: $f' x y = u (f x y) = u (f y x) = f' y x.$
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   274
\<close>
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   275
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   276
lemma combine_infmatrix_commute:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   277
  "commutative f \<Longrightarrow> commutative (combine_infmatrix f)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   278
by (simp add: commutative_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   279
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   280
lemma combine_matrix_commute:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   281
"commutative f \<Longrightarrow> commutative (combine_matrix f)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   282
by (simp add: combine_matrix_def commutative_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   283
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   284
text\<open>
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   285
On the contrary, given an associative function $f$ we cannot expect $f'$ to be associative. A counterexample is given by $A=\ganz$, $B=\{-1, 0, 1\}$,
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   286
as $f$ we take addition on $\ganz$, which is clearly associative. The abstraction is given by  $u(a) = 0$ for $a \notin B$. Then we have
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   287
\[ f' (f' 1 1) -1 = u(f (u (f 1 1)) -1) = u(f (u 2) -1) = u (f 0 -1) = -1, \]
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   288
but on the other hand we have
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   289
\[ f' 1 (f' 1 -1) = u (f 1 (u (f 1 -1))) = u (f 1 0) = 1.\]
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   290
A way out of this problem is to assume that $f(A\times A)\subset A$ holds, and this is what we are going to do:
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   291
\<close>
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   292
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   293
lemma nonzero_positions_combine_infmatrix[simp]: "f 0 0 = 0 \<Longrightarrow> nonzero_positions (combine_infmatrix f A B) \<subseteq> (nonzero_positions A) \<union> (nonzero_positions B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   294
by (rule subsetI, simp add: nonzero_positions_def combine_infmatrix_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   295
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   296
lemma finite_nonzero_positions_Rep[simp]: "finite (nonzero_positions (Rep_matrix A))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   297
by (insert Rep_matrix [of A], simp add: matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   298
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   299
lemma combine_infmatrix_closed [simp]:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   300
  "f 0 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))) = combine_infmatrix f (Rep_matrix A) (Rep_matrix B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   301
apply (rule Abs_matrix_inverse)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   302
apply (simp add: matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   303
apply (rule finite_subset[of _ "(nonzero_positions (Rep_matrix A)) \<union> (nonzero_positions (Rep_matrix B))"])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   304
by (simp_all)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   305
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   306
text \<open>We need the next two lemmas only later, but it is analog to the above one, so we prove them now:\<close>
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   307
lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \<Longrightarrow> nonzero_positions (apply_infmatrix f A) \<subseteq> nonzero_positions A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   308
by (rule subsetI, simp add: nonzero_positions_def apply_infmatrix_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   309
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   310
lemma apply_infmatrix_closed [simp]:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   311
  "f 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (apply_infmatrix f (Rep_matrix A))) = apply_infmatrix f (Rep_matrix A)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   312
apply (rule Abs_matrix_inverse)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   313
apply (simp add: matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   314
apply (rule finite_subset[of _ "nonzero_positions (Rep_matrix A)"])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   315
by (simp_all)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   316
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   317
lemma combine_infmatrix_assoc[simp]: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_infmatrix f)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   318
by (simp add: associative_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   319
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   320
lemma comb: "f = g \<Longrightarrow> x = y \<Longrightarrow> f x = g y"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   321
by (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   322
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   323
lemma combine_matrix_assoc: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_matrix f)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   324
apply (simp(no_asm) add: associative_def combine_matrix_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   325
apply (rule comb [of Abs_matrix Abs_matrix])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   326
by (auto, insert combine_infmatrix_assoc[of f], simp add: associative_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   327
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   328
lemma Rep_apply_matrix[simp]: "f 0 = 0 \<Longrightarrow> Rep_matrix (apply_matrix f A) j i = f (Rep_matrix A j i)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   329
by (simp add: apply_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   330
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   331
lemma Rep_combine_matrix[simp]: "f 0 0 = 0 \<Longrightarrow> Rep_matrix (combine_matrix f A B) j i = f (Rep_matrix A j i) (Rep_matrix B j i)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   332
  by(simp add: combine_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   333
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   334
lemma combine_nrows_max: "f 0 0 = 0  \<Longrightarrow> nrows (combine_matrix f A B) <= max (nrows A) (nrows B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   335
by (simp add: nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   336
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   337
lemma combine_ncols_max: "f 0 0 = 0 \<Longrightarrow> ncols (combine_matrix f A B) <= max (ncols A) (ncols B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   338
by (simp add: ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   339
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   340
lemma combine_nrows: "f 0 0 = 0 \<Longrightarrow> nrows A <= q \<Longrightarrow> nrows B <= q \<Longrightarrow> nrows(combine_matrix f A B) <= q"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   341
  by (simp add: nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   342
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   343
lemma combine_ncols: "f 0 0 = 0 \<Longrightarrow> ncols A <= q \<Longrightarrow> ncols B <= q \<Longrightarrow> ncols(combine_matrix f A B) <= q"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   344
  by (simp add: ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   345
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   346
definition zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   347
  "zero_r_neutral f == \<forall>a. f a 0 = a"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   348
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   349
definition zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   350
  "zero_l_neutral f == \<forall>a. f 0 a = a"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   351
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   352
definition zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   353
  "zero_closed f == (\<forall>x. f x 0 = 0) & (\<forall>y. f 0 y = 0)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   354
38273
d31a34569542 modernized some specifications;
wenzelm
parents: 37765
diff changeset
   355
primrec foldseq :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
d31a34569542 modernized some specifications;
wenzelm
parents: 37765
diff changeset
   356
where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   357
  "foldseq f s 0 = s 0"
38273
d31a34569542 modernized some specifications;
wenzelm
parents: 37765
diff changeset
   358
| "foldseq f s (Suc n) = f (s 0) (foldseq f (% k. s(Suc k)) n)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   359
38273
d31a34569542 modernized some specifications;
wenzelm
parents: 37765
diff changeset
   360
primrec foldseq_transposed ::  "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
d31a34569542 modernized some specifications;
wenzelm
parents: 37765
diff changeset
   361
where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   362
  "foldseq_transposed f s 0 = s 0"
38273
d31a34569542 modernized some specifications;
wenzelm
parents: 37765
diff changeset
   363
| "foldseq_transposed f s (Suc n) = f (foldseq_transposed f s n) (s (Suc n))"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   364
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   365
lemma foldseq_assoc : "associative f \<Longrightarrow> foldseq f = foldseq_transposed f"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   366
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   367
  assume a:"associative f"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   368
  then have sublemma: "\<And>n. \<forall>N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   369
  proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   370
    fix n
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   371
    show "\<forall>N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   372
    proof (induct n)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   373
      show "\<forall>N s. N <= 0 \<longrightarrow> foldseq f s N = foldseq_transposed f s N" by simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   374
    next
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   375
      fix n
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   376
      assume b: "\<forall>N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   377
      have c:"\<And>N s. N <= n \<Longrightarrow> foldseq f s N = foldseq_transposed f s N" by (simp add: b)
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   378
      show "\<forall>N t. N <= Suc n \<longrightarrow> foldseq f t N = foldseq_transposed f t N"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   379
      proof (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   380
        fix N t
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   381
        assume Nsuc: "N <= Suc n"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   382
        show "foldseq f t N = foldseq_transposed f t N"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   383
        proof cases
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   384
          assume "N <= n"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   385
          then show "foldseq f t N = foldseq_transposed f t N" by (simp add: b)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   386
        next
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   387
          assume "~(N <= n)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   388
          with Nsuc have Nsuceq: "N = Suc n" by simp
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   389
          have neqz: "n \<noteq> 0 \<Longrightarrow> \<exists>m. n = Suc m & Suc m <= n" by arith
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   390
          have assocf: "!! x y z. f x (f y z) = f (f x y) z" by (insert a, simp add: associative_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   391
          show "foldseq f t N = foldseq_transposed f t N"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   392
            apply (simp add: Nsuceq)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   393
            apply (subst c)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   394
            apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   395
            apply (case_tac "n = 0")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   396
            apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   397
            apply (drule neqz)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   398
            apply (erule exE)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   399
            apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   400
            apply (subst assocf)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   401
            proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   402
              fix m
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   403
              assume "n = Suc m & Suc m <= n"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   404
              then have mless: "Suc m <= n" by arith
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   405
              then have step1: "foldseq_transposed f (% k. t (Suc k)) m = foldseq f (% k. t (Suc k)) m" (is "?T1 = ?T2")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   406
                apply (subst c)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   407
                by simp+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   408
              have step2: "f (t 0) ?T2 = foldseq f t (Suc m)" (is "_ = ?T3") by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   409
              have step3: "?T3 = foldseq_transposed f t (Suc m)" (is "_ = ?T4")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   410
                apply (subst c)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   411
                by (simp add: mless)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   412
              have step4: "?T4 = f (foldseq_transposed f t m) (t (Suc m))" (is "_=?T5") by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   413
              from step1 step2 step3 step4 show sowhat: "f (f (t 0) ?T1) (t (Suc (Suc m))) = f ?T5 (t (Suc (Suc m)))" by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   414
            qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   415
          qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   416
        qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   417
      qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   418
    qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   419
    show "foldseq f = foldseq_transposed f" by ((rule ext)+, insert sublemma, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   420
  qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   421
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   422
lemma foldseq_distr: "\<lbrakk>associative f; commutative f\<rbrakk> \<Longrightarrow> foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   423
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   424
  assume assoc: "associative f"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   425
  assume comm: "commutative f"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   426
  from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   427
  from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   428
  from assoc comm have c: "!! x y z. f x (f y z) = f y (f x z)" by (simp add: commutative_def associative_def)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   429
  have "\<And>n. (\<forall>u v. foldseq f (%k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   430
    apply (induct_tac n)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   431
    apply (simp+, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   432
    by (simp add: a b c)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   433
  then show "foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   434
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   435
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   436
theorem "\<lbrakk>associative f; associative g; \<forall>a b c d. g (f a b) (f c d) = f (g a c) (g b d); \<exists>x y. (f x) \<noteq> (f y); \<exists>x y. (g x) \<noteq> (g y); f x x = x; g x x = x\<rbrakk> \<Longrightarrow> f=g | (\<forall>y. f y x = y) | (\<forall>y. g y x = y)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   437
oops
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   438
(* Model found
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   439
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   440
Trying to find a model that refutes: \<lbrakk>associative f; associative g;
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   441
 \<forall>a b c d. g (f a b) (f c d) = f (g a c) (g b d); \<exists>x y. f x \<noteq> f y;
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   442
 \<exists>x y. g x \<noteq> g y; f x x = x; g x x = x\<rbrakk>
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   443
\<Longrightarrow> f = g \<or> (\<forall>y. f y x = y) \<or> (\<forall>y. g y x = y)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   444
Searching for a model of size 1, translating term... invoking SAT solver... no model found.
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   445
Searching for a model of size 2, translating term... invoking SAT solver... no model found.
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   446
Searching for a model of size 3, translating term... invoking SAT solver...
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   447
Model found:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   448
Size of types: 'a: 3
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   449
x: a1
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   450
g: (a0\<mapsto>(a0\<mapsto>a1, a1\<mapsto>a0, a2\<mapsto>a1), a1\<mapsto>(a0\<mapsto>a0, a1\<mapsto>a1, a2\<mapsto>a0), a2\<mapsto>(a0\<mapsto>a1, a1\<mapsto>a0, a2\<mapsto>a1))
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   451
f: (a0\<mapsto>(a0\<mapsto>a0, a1\<mapsto>a0, a2\<mapsto>a0), a1\<mapsto>(a0\<mapsto>a1, a1\<mapsto>a1, a2\<mapsto>a1), a2\<mapsto>(a0\<mapsto>a0, a1\<mapsto>a0, a2\<mapsto>a0))
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   452
*)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   453
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   454
lemma foldseq_zero:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   455
assumes fz: "f 0 0 = 0" and sz: "\<forall>i. i <= n \<longrightarrow> s i = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   456
shows "foldseq f s n = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   457
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   458
  have "\<And>n. \<forall>s. (\<forall>i. i <= n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s n = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   459
    apply (induct_tac n)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   460
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   461
    by (simp add: fz)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   462
  then show "foldseq f s n = 0" by (simp add: sz)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   463
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   464
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   465
lemma foldseq_significant_positions:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   466
  assumes p: "\<forall>i. i <= N \<longrightarrow> S i = T i"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   467
  shows "foldseq f S N = foldseq f T N"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   468
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   469
  have "\<And>m. \<forall>s t. (\<forall>i. i<=m \<longrightarrow> s i = t i) \<longrightarrow> foldseq f s m = foldseq f t m"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   470
    apply (induct_tac m)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   471
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   472
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   473
    apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   474
    proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   475
      fix n
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   476
      fix s::"nat\<Rightarrow>'a"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   477
      fix t::"nat\<Rightarrow>'a"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   478
      assume a: "\<forall>s t. (\<forall>i\<le>n. s i = t i) \<longrightarrow> foldseq f s n = foldseq f t n"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   479
      assume b: "\<forall>i\<le>Suc n. s i = t i"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   480
      have c:"!! a b. a = b \<Longrightarrow> f (t 0) a = f (t 0) b" by blast
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   481
      have d:"!! s t. (\<forall>i\<le>n. s i = t i) \<Longrightarrow> foldseq f s n = foldseq f t n" by (simp add: a)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   482
      show "f (t 0) (foldseq f (\<lambda>k. s (Suc k)) n) = f (t 0) (foldseq f (\<lambda>k. t (Suc k)) n)" by (rule c, simp add: d b)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   483
    qed
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   484
  with p show ?thesis by simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   485
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   486
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   487
lemma foldseq_tail:
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   488
  assumes "M <= N"
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   489
  shows "foldseq f S N = foldseq f (% k. (if k < M then (S k) else (foldseq f (% k. S(k+M)) (N-M)))) M"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   490
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   491
  have suc: "\<And>a b. \<lbrakk>a <= Suc b; a \<noteq> Suc b\<rbrakk> \<Longrightarrow> a <= b" by arith
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   492
  have a: "\<And>a b c . a = b \<Longrightarrow> f c a = f c b" by blast
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   493
  have "\<And>n. \<forall>m s. m <= n \<longrightarrow> foldseq f s n = foldseq f (% k. (if k < m then (s k) else (foldseq f (% k. s(k+m)) (n-m)))) m"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   494
    apply (induct_tac n)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   495
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   496
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   497
    apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   498
    apply (case_tac "m = Suc na")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   499
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   500
    apply (rule a)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   501
    apply (rule foldseq_significant_positions)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   502
    apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   503
    apply (drule suc, simp+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   504
    proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   505
      fix na m s
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   506
      assume suba:"\<forall>m\<le>na. \<forall>s. foldseq f s na = foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (na - m))m"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   507
      assume subb:"m <= na"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   508
      from suba have subc:"!! m s. m <= na \<Longrightarrow>foldseq f s na = foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (na - m))m" by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   509
      have subd: "foldseq f (\<lambda>k. if k < m then s (Suc k) else foldseq f (\<lambda>k. s (Suc (k + m))) (na - m)) m =
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   510
        foldseq f (% k. s(Suc k)) na"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   511
        by (rule subc[of m "% k. s(Suc k)", THEN sym], simp add: subb)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   512
      from subb have sube: "m \<noteq> 0 \<Longrightarrow> \<exists>mm. m = Suc mm & mm <= na" by arith
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   513
      show "f (s 0) (foldseq f (\<lambda>k. if k < m then s (Suc k) else foldseq f (\<lambda>k. s (Suc (k + m))) (na - m)) m) =
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   514
        foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (Suc na - m)) m"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   515
        apply (simp add: subd)
38526
a9ce311eb6b9 tuned proof
haftmann
parents: 38273
diff changeset
   516
        apply (cases "m = 0")
63566
e5abbdee461a more accurate cong del;
wenzelm
parents: 63167
diff changeset
   517
         apply simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   518
        apply (drule sube)
63566
e5abbdee461a more accurate cong del;
wenzelm
parents: 63167
diff changeset
   519
        apply auto
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   520
        apply (rule a)
63566
e5abbdee461a more accurate cong del;
wenzelm
parents: 63167
diff changeset
   521
        apply (simp add: subc cong del: if_weak_cong)
e5abbdee461a more accurate cong del;
wenzelm
parents: 63167
diff changeset
   522
        done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   523
    qed
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   524
  then show ?thesis using assms by simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   525
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   526
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   527
lemma foldseq_zerotail:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   528
  assumes
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   529
  fz: "f 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   530
  and sz: "\<forall>i.  n <= i \<longrightarrow> s i = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   531
  and nm: "n <= m"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   532
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   533
  "foldseq f s n = foldseq f s m"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   534
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   535
  show "foldseq f s n = foldseq f s m"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   536
    apply (simp add: foldseq_tail[OF nm, of f s])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   537
    apply (rule foldseq_significant_positions)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   538
    apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   539
    apply (subst foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   540
    by (simp add: fz sz)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   541
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   542
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   543
lemma foldseq_zerotail2:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   544
  assumes "\<forall>x. f x 0 = x"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   545
  and "\<forall>i. n < i \<longrightarrow> s i = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   546
  and nm: "n <= m"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   547
  shows "foldseq f s n = foldseq f s m"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   548
proof -
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   549
  have "f 0 0 = 0" by (simp add: assms)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   550
  have b: "\<And>m n. n <= m \<Longrightarrow> m \<noteq> n \<Longrightarrow> \<exists>k. m-n = Suc k" by arith
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   551
  have c: "0 <= m" by simp
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   552
  have d: "\<And>k. k \<noteq> 0 \<Longrightarrow> \<exists>l. k = Suc l" by arith
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   553
  show ?thesis
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   554
    apply (subst foldseq_tail[OF nm])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   555
    apply (rule foldseq_significant_positions)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   556
    apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   557
    apply (case_tac "m=n")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   558
    apply (simp+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   559
    apply (drule b[OF nm])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   560
    apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   561
    apply (case_tac "k=0")
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   562
    apply (simp add: assms)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   563
    apply (drule d)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   564
    apply (auto)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   565
    apply (simp add: assms foldseq_zero)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   566
    done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   567
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   568
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   569
lemma foldseq_zerostart:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   570
  "\<forall>x. f 0 (f 0 x) = f 0 x \<Longrightarrow> \<forall>i. i <= n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   571
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   572
  assume f00x: "\<forall>x. f 0 (f 0 x) = f 0 x"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   573
  have "\<forall>s. (\<forall>i. i<=n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   574
    apply (induct n)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   575
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   576
    apply (rule allI, rule impI)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   577
    proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   578
      fix n
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   579
      fix s
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   580
      have a:"foldseq f s (Suc (Suc n)) = f (s 0) (foldseq f (% k. s(Suc k)) (Suc n))" by simp
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   581
      assume b: "\<forall>s. ((\<forall>i\<le>n. s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n)))"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   582
      from b have c:"!! s. (\<forall>i\<le>n. s i = 0) \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" by simp
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   583
      assume d: "\<forall>i. i <= Suc n \<longrightarrow> s i = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   584
      show "foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   585
        apply (subst a)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   586
        apply (subst c)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   587
        by (simp add: d f00x)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   588
    qed
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   589
  then show "\<forall>i. i <= n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" by simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   590
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   591
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   592
lemma foldseq_zerostart2:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   593
  "\<forall>x. f 0 x = x \<Longrightarrow> \<forall>i. i < n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s n = s n"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   594
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   595
  assume a: "\<forall>i. i<n \<longrightarrow> s i = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   596
  assume x: "\<forall>x. f 0 x = x"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   597
  from x have f00x: "\<forall>x. f 0 (f 0 x) = f 0 x" by blast
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   598
  have b: "\<And>i l. i < Suc l = (i <= l)" by arith
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   599
  have d: "\<And>k. k \<noteq> 0 \<Longrightarrow> \<exists>l. k = Suc l" by arith
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   600
  show "foldseq f s n = s n"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   601
  apply (case_tac "n=0")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   602
  apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   603
  apply (insert a)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   604
  apply (drule d)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   605
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   606
  apply (simp add: b)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   607
  apply (insert f00x)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   608
  apply (drule foldseq_zerostart)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   609
  by (simp add: x)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   610
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   611
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   612
lemma foldseq_almostzero:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   613
  assumes f0x: "\<forall>x. f 0 x = x" and fx0: "\<forall>x. f x 0 = x" and s0: "\<forall>i. i \<noteq> j \<longrightarrow> s i = 0"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   614
  shows "foldseq f s n = (if (j <= n) then (s j) else 0)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   615
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   616
  from s0 have a: "\<forall>i. i < j \<longrightarrow> s i = 0" by simp
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   617
  from s0 have b: "\<forall>i. j < i \<longrightarrow> s i = 0" by simp
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   618
  show ?thesis
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   619
    apply auto
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   620
    apply (subst foldseq_zerotail2[of f, OF fx0, of j, OF b, of n, THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   621
    apply simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   622
    apply (subst foldseq_zerostart2)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   623
    apply (simp add: f0x a)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   624
    apply (subst foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   625
    by (simp add: s0 f0x)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   626
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   627
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   628
lemma foldseq_distr_unary:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   629
  assumes "!! a b. g (f a b) = f (g a) (g b)"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   630
  shows "g(foldseq f s n) = foldseq f (% x. g(s x)) n"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   631
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   632
  have "\<forall>s. g(foldseq f s n) = foldseq f (% x. g(s x)) n"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   633
    apply (induct_tac n)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   634
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   635
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   636
    apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   637
    apply (drule_tac x="% k. s (Suc k)" in spec)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   638
    by (simp add: assms)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   639
  then show ?thesis by simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   640
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   641
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   642
definition mult_matrix_n :: "nat \<Rightarrow> (('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   643
  "mult_matrix_n n fmul fadd A B == Abs_matrix(% j i. foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) n)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   644
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   645
definition mult_matrix :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   646
  "mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   647
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   648
lemma mult_matrix_n:
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   649
  assumes "ncols A \<le>  n" (is ?An) "nrows B \<le> n" (is ?Bn) "fadd 0 0 = 0" "fmul 0 0 = 0"
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   650
  shows c:"mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   651
proof -
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   652
  show ?thesis using assms
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   653
    apply (simp add: mult_matrix_def mult_matrix_n_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   654
    apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   655
    apply (rule foldseq_zerotail, simp_all add: nrows_le ncols_le assms)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   656
    done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   657
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   658
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   659
lemma mult_matrix_nm:
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   660
  assumes "ncols A <= n" "nrows B <= n" "ncols A <= m" "nrows B <= m" "fadd 0 0 = 0" "fmul 0 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   661
  shows "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   662
proof -
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   663
  from assms have "mult_matrix_n n fmul fadd A B = mult_matrix fmul fadd A B"
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   664
    by (simp add: mult_matrix_n)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   665
  also from assms have "\<dots> = mult_matrix_n m fmul fadd A B"
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   666
    by (simp add: mult_matrix_n[THEN sym])
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   667
  finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   668
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   669
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   670
definition r_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   671
  "r_distributive fmul fadd == \<forall>a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   672
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   673
definition l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   674
  "l_distributive fmul fadd == \<forall>a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   675
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   676
definition distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   677
  "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   678
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   679
lemma max1: "!! a x y. (a::nat) <= x \<Longrightarrow> a <= max x y" by (arith)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   680
lemma max2: "!! b x y. (b::nat) <= y \<Longrightarrow> b <= max x y" by (arith)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   681
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   682
lemma r_distributive_matrix:
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   683
 assumes
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   684
  "r_distributive fmul fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   685
  "associative fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   686
  "commutative fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   687
  "fadd 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   688
  "\<forall>a. fmul a 0 = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   689
  "\<forall>a. fmul 0 a = 0"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   690
 shows "r_distributive (mult_matrix fmul fadd) (combine_matrix fadd)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   691
proof -
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   692
  from assms show ?thesis
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   693
    apply (simp add: r_distributive_def mult_matrix_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   694
    proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   695
      fix a::"'a matrix"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   696
      fix u::"'b matrix"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   697
      fix v::"'b matrix"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   698
      let ?mx = "max (ncols a) (max (nrows u) (nrows v))"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   699
      from assms show "mult_matrix_n (max (ncols a) (nrows (combine_matrix fadd u v))) fmul fadd a (combine_matrix fadd u v) =
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   700
        combine_matrix fadd (mult_matrix_n (max (ncols a) (nrows u)) fmul fadd a u) (mult_matrix_n (max (ncols a) (nrows v)) fmul fadd a v)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   701
        apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   702
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   703
        apply (subst mult_matrix_nm[of _ _ v ?mx fadd fmul])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   704
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   705
        apply (subst mult_matrix_nm[of _ _ u ?mx fadd fmul])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   706
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   707
        apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   708
        apply (simp add: combine_matrix_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   709
        apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   710
        apply (simplesubst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   711
        apply (simp, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   712
        apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   713
        apply (rule exI[of _ "ncols v"], simp add: ncols_le foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   714
        apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   715
        apply (simp, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   716
        apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   717
        apply (rule exI[of _ "ncols u"], simp add: ncols_le foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   718
        done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   719
    qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   720
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   721
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   722
lemma l_distributive_matrix:
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   723
 assumes
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   724
  "l_distributive fmul fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   725
  "associative fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   726
  "commutative fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   727
  "fadd 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   728
  "\<forall>a. fmul a 0 = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   729
  "\<forall>a. fmul 0 a = 0"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   730
 shows "l_distributive (mult_matrix fmul fadd) (combine_matrix fadd)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   731
proof -
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   732
  from assms show ?thesis
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   733
    apply (simp add: l_distributive_def mult_matrix_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   734
    proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   735
      fix a::"'b matrix"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   736
      fix u::"'a matrix"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   737
      fix v::"'a matrix"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   738
      let ?mx = "max (nrows a) (max (ncols u) (ncols v))"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   739
      from assms show "mult_matrix_n (max (ncols (combine_matrix fadd u v)) (nrows a)) fmul fadd (combine_matrix fadd u v) a =
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   740
               combine_matrix fadd (mult_matrix_n (max (ncols u) (nrows a)) fmul fadd u a) (mult_matrix_n (max (ncols v) (nrows a)) fmul fadd v a)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   741
        apply (subst mult_matrix_nm[of v _ _ ?mx fadd fmul])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   742
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   743
        apply (subst mult_matrix_nm[of u _ _ ?mx fadd fmul])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   744
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   745
        apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   746
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   747
        apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   748
        apply (simp add: combine_matrix_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   749
        apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   750
        apply (simplesubst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   751
        apply (simp, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   752
        apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   753
        apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   754
        apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   755
        apply (simp, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   756
        apply (rule exI[of _ "nrows u"], simp add: nrows_le foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   757
        apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   758
        done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   759
    qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   760
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   761
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   762
instantiation matrix :: (zero) zero
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   763
begin
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   764
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 35818
diff changeset
   765
definition zero_matrix_def: "0 = Abs_matrix (\<lambda>j i. 0)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   766
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   767
instance ..
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   768
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   769
end
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   770
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   771
lemma Rep_zero_matrix_def[simp]: "Rep_matrix 0 j i = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   772
  apply (simp add: zero_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   773
  apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   774
  by (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   775
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   776
lemma zero_matrix_def_nrows[simp]: "nrows 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   777
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   778
  have a:"!! (x::nat). x <= 0 \<Longrightarrow> x = 0" by (arith)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   779
  show "nrows 0 = 0" by (rule a, subst nrows_le, simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   780
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   781
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   782
lemma zero_matrix_def_ncols[simp]: "ncols 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   783
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   784
  have a:"!! (x::nat). x <= 0 \<Longrightarrow> x = 0" by (arith)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   785
  show "ncols 0 = 0" by (rule a, subst ncols_le, simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   786
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   787
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   788
lemma combine_matrix_zero_l_neutral: "zero_l_neutral f \<Longrightarrow> zero_l_neutral (combine_matrix f)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   789
  by (simp add: zero_l_neutral_def combine_matrix_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   790
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   791
lemma combine_matrix_zero_r_neutral: "zero_r_neutral f \<Longrightarrow> zero_r_neutral (combine_matrix f)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   792
  by (simp add: zero_r_neutral_def combine_matrix_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   793
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   794
lemma mult_matrix_zero_closed: "\<lbrakk>fadd 0 0 = 0; zero_closed fmul\<rbrakk> \<Longrightarrow> zero_closed (mult_matrix fmul fadd)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   795
  apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   796
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   797
  by (subst foldseq_zero, (simp add: zero_matrix_def)+)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   798
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   799
lemma mult_matrix_n_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd A 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   800
  apply (simp add: mult_matrix_n_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   801
  apply (subst foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   802
  by (simp_all add: zero_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   803
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   804
lemma mult_matrix_n_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd 0 A = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   805
  apply (simp add: mult_matrix_n_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   806
  apply (subst foldseq_zero)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   807
  by (simp_all add: zero_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   808
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   809
lemma mult_matrix_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd 0 A = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   810
by (simp add: mult_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   811
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   812
lemma mult_matrix_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; \<forall>a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd A 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   813
by (simp add: mult_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   814
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   815
lemma apply_matrix_zero[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   816
  apply (simp add: apply_matrix_def apply_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   817
  by (simp add: zero_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   818
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   819
lemma combine_matrix_zero: "f 0 0 = 0 \<Longrightarrow> combine_matrix f 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   820
  apply (simp add: combine_matrix_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   821
  by (simp add: zero_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   822
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   823
lemma transpose_matrix_zero[simp]: "transpose_matrix 0 = 0"
46985
bd955d9f464b tuned proofs;
wenzelm
parents: 46702
diff changeset
   824
apply (simp add: transpose_matrix_def zero_matrix_def RepAbs_matrix)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   825
apply (subst Rep_matrix_inject[symmetric], (rule ext)+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   826
apply (simp add: RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   827
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   828
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   829
lemma apply_zero_matrix_def[simp]: "apply_matrix (% x. 0) A = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   830
  apply (simp add: apply_matrix_def apply_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   831
  by (simp add: zero_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   832
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   833
definition singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   834
  "singleton_matrix j i a == Abs_matrix(% m n. if j = m & i = n then a else 0)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   835
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   836
definition move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix" where
46702
202a09ba37d8 avoid using constant Int.neg
huffman
parents: 45694
diff changeset
   837
  "move_matrix A y x == Abs_matrix(% j i. if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat ((int j)-y)) (nat ((int i)-x)))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   838
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   839
definition take_rows :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   840
  "take_rows A r == Abs_matrix(% j i. if (j < r) then (Rep_matrix A j i) else 0)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   841
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   842
definition take_columns :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   843
  "take_columns A c == Abs_matrix(% j i. if (i < c) then (Rep_matrix A j i) else 0)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   844
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   845
definition column_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   846
  "column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   847
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
   848
definition row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   849
  "row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   850
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   851
lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   852
apply (simp add: singleton_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   853
apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   854
apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   855
apply (rule exI[of _ "Suc m"], simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   856
apply (rule exI[of _ "Suc n"], simp+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   857
by (subst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   858
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   859
lemma apply_singleton_matrix[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f (singleton_matrix j i x) = (singleton_matrix j i (f x))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   860
apply (subst Rep_matrix_inject[symmetric])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   861
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   862
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   863
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   864
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   865
lemma singleton_matrix_zero[simp]: "singleton_matrix j i 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   866
  by (simp add: singleton_matrix_def zero_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   867
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   868
lemma nrows_singleton[simp]: "nrows(singleton_matrix j i e) = (if e = 0 then 0 else Suc j)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   869
proof-
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   870
have th: "\<not> (\<forall>m. m \<le> j)" "\<exists>n. \<not> n \<le> i" by arith+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   871
from th show ?thesis 
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   872
apply (auto)
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   873
apply (rule le_antisym)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   874
apply (subst nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   875
apply (simp add: singleton_matrix_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   876
apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   877
apply auto
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   878
apply (simp add: Suc_le_eq)
61824
dcbe9f756ae0 not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents: 61260
diff changeset
   879
apply (rule not_le_imp_less)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   880
apply (subst nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   881
by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   882
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   883
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   884
lemma ncols_singleton[simp]: "ncols(singleton_matrix j i e) = (if e = 0 then 0 else Suc i)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   885
proof-
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   886
have th: "\<not> (\<forall>m. m \<le> j)" "\<exists>n. \<not> n \<le> i" by arith+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   887
from th show ?thesis 
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   888
apply (auto)
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   889
apply (rule le_antisym)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   890
apply (subst ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   891
apply (simp add: singleton_matrix_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   892
apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   893
apply auto
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   894
apply (simp add: Suc_le_eq)
61824
dcbe9f756ae0 not_leE -> not_le_imp_less and other tidying
paulson <lp15@cam.ac.uk>
parents: 61260
diff changeset
   895
apply (rule not_le_imp_less)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   896
apply (subst ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   897
by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   898
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   899
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   900
lemma combine_singleton: "f 0 0 = 0 \<Longrightarrow> combine_matrix f (singleton_matrix j i a) (singleton_matrix j i b) = singleton_matrix j i (f a b)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   901
apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   902
apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   903
apply (rule exI[of _ "Suc j"], simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   904
apply (rule exI[of _ "Suc i"], simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   905
apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   906
apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   907
apply (rule exI[of _ "Suc j"], simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   908
apply (rule exI[of _ "Suc i"], simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   909
by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   910
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   911
lemma transpose_singleton[simp]: "transpose_matrix (singleton_matrix j i a) = singleton_matrix i j a"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   912
apply (subst Rep_matrix_inject[symmetric], (rule ext)+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   913
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   914
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   915
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   916
lemma Rep_move_matrix[simp]:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   917
  "Rep_matrix (move_matrix A y x) j i =
46702
202a09ba37d8 avoid using constant Int.neg
huffman
parents: 45694
diff changeset
   918
  (if (((int j)-y) < 0) | (((int i)-x) < 0) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   919
apply (simp add: move_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   920
apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   921
by (subst RepAbs_matrix,
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61824
diff changeset
   922
  rule exI[of _ "(nrows A)+(nat \<bar>y\<bar>)"], auto, rule nrows, arith,
1135b8de26c3 more symbols;
wenzelm
parents: 61824
diff changeset
   923
  rule exI[of _ "(ncols A)+(nat \<bar>x\<bar>)"], auto, rule ncols, arith)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   924
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   925
lemma move_matrix_0_0[simp]: "move_matrix A 0 0 = A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   926
by (simp add: move_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   927
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   928
lemma move_matrix_ortho: "move_matrix A j i = move_matrix (move_matrix A j 0) 0 i"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   929
apply (subst Rep_matrix_inject[symmetric])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   930
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   931
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   932
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   933
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   934
lemma transpose_move_matrix[simp]:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   935
  "transpose_matrix (move_matrix A x y) = move_matrix (transpose_matrix A) y x"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   936
apply (subst Rep_matrix_inject[symmetric], (rule ext)+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   937
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   938
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   939
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   940
lemma move_matrix_singleton[simp]: "move_matrix (singleton_matrix u v x) j i = 
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   941
  (if (j + int u < 0) | (i + int v < 0) then 0 else (singleton_matrix (nat (j + int u)) (nat (i + int v)) x))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   942
  apply (subst Rep_matrix_inject[symmetric])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   943
  apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   944
  apply (case_tac "j + int u < 0")
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   945
  apply (simp, arith)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   946
  apply (case_tac "i + int v < 0")
46702
202a09ba37d8 avoid using constant Int.neg
huffman
parents: 45694
diff changeset
   947
  apply (simp, arith)
202a09ba37d8 avoid using constant Int.neg
huffman
parents: 45694
diff changeset
   948
  apply simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   949
  apply arith
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   950
  done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   951
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   952
lemma Rep_take_columns[simp]:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   953
  "Rep_matrix (take_columns A c) j i =
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   954
  (if i < c then (Rep_matrix A j i) else 0)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   955
apply (simp add: take_columns_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   956
apply (simplesubst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   957
apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   958
apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   959
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   960
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   961
lemma Rep_take_rows[simp]:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   962
  "Rep_matrix (take_rows A r) j i =
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   963
  (if j < r then (Rep_matrix A j i) else 0)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   964
apply (simp add: take_rows_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   965
apply (simplesubst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   966
apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   967
apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   968
done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   969
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   970
lemma Rep_column_of_matrix[simp]:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   971
  "Rep_matrix (column_of_matrix A c) j i = (if i = 0 then (Rep_matrix A j c) else 0)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   972
  by (simp add: column_of_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   973
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   974
lemma Rep_row_of_matrix[simp]:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   975
  "Rep_matrix (row_of_matrix A r) j i = (if j = 0 then (Rep_matrix A r i) else 0)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   976
  by (simp add: row_of_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   977
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   978
lemma column_of_matrix: "ncols A <= n \<Longrightarrow> column_of_matrix A n = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   979
apply (subst Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   980
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   981
by (simp add: ncols)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   982
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   983
lemma row_of_matrix: "nrows A <= n \<Longrightarrow> row_of_matrix A n = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   984
apply (subst Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   985
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   986
by (simp add: nrows)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   987
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   988
lemma mult_matrix_singleton_right[simp]:
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   989
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   990
  "\<forall>x. fmul x 0 = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   991
  "\<forall>x. fmul 0 x = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   992
  "\<forall>x. fadd 0 x = x"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
   993
  "\<forall>x. fadd x 0 = x"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   994
  shows "(mult_matrix fmul fadd A (singleton_matrix j i e)) = apply_matrix (% x. fmul x e) (move_matrix (column_of_matrix A j) 0 (int i))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   995
  apply (simp add: mult_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   996
  apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   997
  apply (auto)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
   998
  apply (simp add: assms)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
   999
  apply (simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1000
  apply (rule comb[of "Abs_matrix" "Abs_matrix"], auto, (rule ext)+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1001
  apply (subst foldseq_almostzero[of _ j])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1002
  apply (simp add: assms)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1003
  apply (auto)
29700
22faf21db3df added some simp rules
nipkow
parents: 29667
diff changeset
  1004
  done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1005
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1006
lemma mult_matrix_ext:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1007
  assumes
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1008
  eprem:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1009
  "\<exists>e. (\<forall>a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1010
  and fprems:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1011
  "\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1012
  "\<forall>a. fmul a 0 = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1013
  "\<forall>a. fadd a 0 = a"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1014
  "\<forall>a. fadd 0 a = a"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1015
  and contraprems:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1016
  "mult_matrix fmul fadd A = mult_matrix fmul fadd B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1017
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1018
  "A = B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1019
proof(rule contrapos_np[of "False"], simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1020
  assume a: "A \<noteq> B"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1021
  have b: "\<And>f g. (\<forall>x y. f x y = g x y) \<Longrightarrow> f = g" by ((rule ext)+, auto)
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1022
  have "\<exists>j i. (Rep_matrix A j i) \<noteq> (Rep_matrix B j i)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1023
    apply (rule contrapos_np[of "False"], simp+)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1024
    apply (insert b[of "Rep_matrix A" "Rep_matrix B"], simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1025
    by (simp add: Rep_matrix_inject a)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1026
  then obtain J I where c:"(Rep_matrix A J I) \<noteq> (Rep_matrix B J I)" by blast
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1027
  from eprem obtain e where eprops:"(\<forall>a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)" by blast
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1028
  let ?S = "singleton_matrix I 0 e"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1029
  let ?comp = "mult_matrix fmul fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1030
  have d: "!!x f g. f = g \<Longrightarrow> f x = g x" by blast
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1031
  have e: "(% x. fmul x e) 0 = 0" by (simp add: assms)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1032
  have "~(?comp A ?S = ?comp B ?S)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1033
    apply (rule notI)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1034
    apply (simp add: fprems eprops)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1035
    apply (simp add: Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1036
    apply (drule d[of _ _ "J"], drule d[of _ _ "0"])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1037
    by (simp add: e c eprops)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1038
  with contraprems show "False" by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1039
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1040
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1041
definition foldmatrix :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1042
  "foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1043
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1044
definition foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1045
  "foldmatrix_transposed f g A m n == foldseq g (% j. foldseq_transposed f (A j) n) m"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1046
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1047
lemma foldmatrix_transpose:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1048
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1049
  "\<forall>a b c d. g(f a b) (f c d) = f (g a c) (g b d)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1050
  shows
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1051
  "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1052
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1053
  have forall:"\<And>P x. (\<forall>x. P x) \<Longrightarrow> P x" by auto
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1054
  have tworows:"\<forall>A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1055
    apply (induct n)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1056
    apply (simp add: foldmatrix_def foldmatrix_transposed_def assms)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1057
    apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1058
    by (drule_tac x="(% j i. A j (Suc i))" in forall, simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1059
  show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1060
    apply (simp add: foldmatrix_def foldmatrix_transposed_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1061
    apply (induct m, simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1062
    apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1063
    apply (insert tworows)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1064
    apply (drule_tac x="% j i. (if j = 0 then (foldseq_transposed g (\<lambda>u. A u i) m) else (A (Suc m) i))" in spec)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1065
    by (simp add: foldmatrix_def foldmatrix_transposed_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1066
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1067
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1068
lemma foldseq_foldseq:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1069
assumes
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1070
  "associative f"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1071
  "associative g"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1072
  "\<forall>a b c d. g(f a b) (f c d) = f (g a c) (g b d)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1073
shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1074
  "foldseq g (% j. foldseq f (A j) n) m = foldseq f (% j. foldseq g ((transpose_infmatrix A) j) m) n"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1075
  apply (insert foldmatrix_transpose[of g f A m n])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1076
  by (simp add: foldmatrix_def foldmatrix_transposed_def foldseq_assoc[THEN sym] assms)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1077
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1078
lemma mult_n_nrows:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1079
assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1080
"\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1081
"\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1082
"fadd 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1083
shows "nrows (mult_matrix_n n fmul fadd A B) \<le> nrows A"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1084
apply (subst nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1085
apply (simp add: mult_matrix_n_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1086
apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1087
apply (rule_tac x="nrows A" in exI)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1088
apply (simp add: nrows assms foldseq_zero)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1089
apply (rule_tac x="ncols B" in exI)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1090
apply (simp add: ncols assms foldseq_zero)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1091
apply (simp add: nrows assms foldseq_zero)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1092
done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1093
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1094
lemma mult_n_ncols:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1095
assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1096
"\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1097
"\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1098
"fadd 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1099
shows "ncols (mult_matrix_n n fmul fadd A B) \<le> ncols B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1100
apply (subst ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1101
apply (simp add: mult_matrix_n_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1102
apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1103
apply (rule_tac x="nrows A" in exI)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1104
apply (simp add: nrows assms foldseq_zero)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1105
apply (rule_tac x="ncols B" in exI)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1106
apply (simp add: ncols assms foldseq_zero)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1107
apply (simp add: ncols assms foldseq_zero)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1108
done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1109
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1110
lemma mult_nrows:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1111
assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1112
"\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1113
"\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1114
"fadd 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1115
shows "nrows (mult_matrix fmul fadd A B) \<le> nrows A"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1116
by (simp add: mult_matrix_def mult_n_nrows assms)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1117
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1118
lemma mult_ncols:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1119
assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1120
"\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1121
"\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1122
"fadd 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1123
shows "ncols (mult_matrix fmul fadd A B) \<le> ncols B"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1124
by (simp add: mult_matrix_def mult_n_ncols assms)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1125
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1126
lemma nrows_move_matrix_le: "nrows (move_matrix A j i) <= nat((int (nrows A)) + j)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1127
  apply (auto simp add: nrows_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1128
  apply (rule nrows)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1129
  apply (arith)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1130
  done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1131
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1132
lemma ncols_move_matrix_le: "ncols (move_matrix A j i) <= nat((int (ncols A)) + i)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1133
  apply (auto simp add: ncols_le)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1134
  apply (rule ncols)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1135
  apply (arith)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1136
  done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1137
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1138
lemma mult_matrix_assoc:
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1139
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1140
  "\<forall>a. fmul1 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1141
  "\<forall>a. fmul1 a 0 = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1142
  "\<forall>a. fmul2 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1143
  "\<forall>a. fmul2 a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1144
  "fadd1 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1145
  "fadd2 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1146
  "\<forall>a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1147
  "associative fadd1"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1148
  "associative fadd2"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1149
  "\<forall>a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1150
  "\<forall>a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1151
  "\<forall>a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1152
  shows "mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B) C = mult_matrix fmul1 fadd1 A (mult_matrix fmul2 fadd2 B C)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1153
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1154
  have comb_left:  "!! A B x y. A = B \<Longrightarrow> (Rep_matrix (Abs_matrix A)) x y = (Rep_matrix(Abs_matrix B)) x y" by blast
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1155
  have fmul2fadd1fold: "!! x s n. fmul2 (foldseq fadd1 s n)  x = foldseq fadd1 (% k. fmul2 (s k) x) n"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1156
    by (rule_tac g1 = "% y. fmul2 y x" in ssubst [OF foldseq_distr_unary], insert assms, simp_all)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1157
  have fmul1fadd2fold: "!! x s n. fmul1 x (foldseq fadd2 s n) = foldseq fadd2 (% k. fmul1 x (s k)) n"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1158
    using assms by (rule_tac g1 = "% y. fmul1 x y" in ssubst [OF foldseq_distr_unary], simp_all)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1159
  let ?N = "max (ncols A) (max (ncols B) (max (nrows B) (nrows C)))"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1160
  show ?thesis
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1161
    apply (simp add: Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1162
    apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1163
    apply (simp add: mult_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1164
    apply (simplesubst mult_matrix_nm[of _ "max (ncols (mult_matrix_n (max (ncols A) (nrows B)) fmul1 fadd1 A B)) (nrows C)" _ "max (ncols B) (nrows C)"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1165
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1166
    apply (simplesubst mult_matrix_nm[of _ "max (ncols A) (nrows (mult_matrix_n (max (ncols B) (nrows C)) fmul2 fadd2 B C))" _ "max (ncols A) (nrows B)"])
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1167
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1168
    apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1169
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1170
    apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1171
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1172
    apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1173
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1174
    apply (simplesubst mult_matrix_nm[of _ _ _ "?N"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1175
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows assms)+
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1176
    apply (simp add: mult_matrix_n_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1177
    apply (rule comb_left)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1178
    apply ((rule ext)+, simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1179
    apply (simplesubst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1180
    apply (rule exI[of _ "nrows B"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1181
    apply (simp add: nrows assms foldseq_zero)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1182
    apply (rule exI[of _ "ncols C"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1183
    apply (simp add: assms ncols foldseq_zero)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1184
    apply (subst RepAbs_matrix)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1185
    apply (rule exI[of _ "nrows A"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1186
    apply (simp add: nrows assms foldseq_zero)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1187
    apply (rule exI[of _ "ncols B"])
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1188
    apply (simp add: assms ncols foldseq_zero)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1189
    apply (simp add: fmul2fadd1fold fmul1fadd2fold assms)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1190
    apply (subst foldseq_foldseq)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1191
    apply (simp add: assms)+
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1192
    apply (simp add: transpose_infmatrix)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1193
    done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1194
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1195
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1196
lemma
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1197
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1198
  "\<forall>a. fmul1 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1199
  "\<forall>a. fmul1 a 0 = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1200
  "\<forall>a. fmul2 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1201
  "\<forall>a. fmul2 a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1202
  "fadd1 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1203
  "fadd2 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1204
  "\<forall>a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1205
  "associative fadd1"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1206
  "associative fadd2"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1207
  "\<forall>a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1208
  "\<forall>a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1209
  "\<forall>a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1210
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1211
  "(mult_matrix fmul1 fadd1 A) o (mult_matrix fmul2 fadd2 B) = mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1212
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1213
apply (simp add: comp_def )
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1214
apply (simp add: mult_matrix_assoc assms)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1215
done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1216
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1217
lemma mult_matrix_assoc_simple:
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1218
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1219
  "\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1220
  "\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1221
  "fadd 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1222
  "associative fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1223
  "commutative fadd"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1224
  "associative fmul"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1225
  "distributive fmul fadd"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1226
  shows "mult_matrix fmul fadd (mult_matrix fmul fadd A B) C = mult_matrix fmul fadd A (mult_matrix fmul fadd B C)"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1227
proof -
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1228
  have "!! a b c d. fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1229
    using assms by (simp add: associative_def commutative_def)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1230
  then show ?thesis
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1231
    apply (subst mult_matrix_assoc)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1232
    using assms
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1233
    apply simp_all
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1234
    apply (simp_all add: associative_def distributive_def l_distributive_def r_distributive_def)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1235
    done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1236
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1237
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1238
lemma transpose_apply_matrix: "f 0 = 0 \<Longrightarrow> transpose_matrix (apply_matrix f A) = apply_matrix f (transpose_matrix A)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1239
apply (simp add: Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1240
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1241
by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1242
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1243
lemma transpose_combine_matrix: "f 0 0 = 0 \<Longrightarrow> transpose_matrix (combine_matrix f A B) = combine_matrix f (transpose_matrix A) (transpose_matrix B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1244
apply (simp add: Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1245
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1246
by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1247
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1248
lemma Rep_mult_matrix:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1249
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1250
  "\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1251
  "\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1252
  "fadd 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1253
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1254
  "Rep_matrix(mult_matrix fmul fadd A B) j i =
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1255
  foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) (max (ncols A) (nrows B))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1256
apply (simp add: mult_matrix_def mult_matrix_n_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1257
apply (subst RepAbs_matrix)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1258
apply (rule exI[of _ "nrows A"], insert assms, simp add: nrows foldseq_zero)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1259
apply (rule exI[of _ "ncols B"], insert assms, simp add: ncols foldseq_zero)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1260
apply simp
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1261
done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1262
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1263
lemma transpose_mult_matrix:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1264
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1265
  "\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1266
  "\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1267
  "fadd 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1268
  "\<forall>x y. fmul y x = fmul x y"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1269
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1270
  "transpose_matrix (mult_matrix fmul fadd A B) = mult_matrix fmul fadd (transpose_matrix B) (transpose_matrix A)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1271
  apply (simp add: Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1272
  apply (rule ext)+
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1273
  using assms
54864
a064732223ad abolished slightly odd global lattice interpretation for min/max
haftmann
parents: 54230
diff changeset
  1274
  apply (simp add: Rep_mult_matrix ac_simps)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1275
  done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1276
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1277
lemma column_transpose_matrix: "column_of_matrix (transpose_matrix A) n = transpose_matrix (row_of_matrix A n)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1278
apply (simp add:  Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1279
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1280
by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1281
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1282
lemma take_columns_transpose_matrix: "take_columns (transpose_matrix A) n = transpose_matrix (take_rows A n)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1283
apply (simp add: Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1284
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1285
by simp
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1286
27580
haftmann
parents: 27484
diff changeset
  1287
instantiation matrix :: ("{zero, ord}") ord
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1288
begin
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1289
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1290
definition
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1291
  le_matrix_def: "A \<le> B \<longleftrightarrow> (\<forall>j i. Rep_matrix A j i \<le> Rep_matrix B j i)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1292
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1293
definition
61076
bdc1e2f0a86a eliminated \<Colon>;
wenzelm
parents: 57512
diff changeset
  1294
  less_def: "A < (B::'a matrix) \<longleftrightarrow> A \<le> B \<and> \<not> B \<le> A"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1295
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1296
instance ..
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1297
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1298
end
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1299
27580
haftmann
parents: 27484
diff changeset
  1300
instance matrix :: ("{zero, order}") order
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1301
apply intro_classes
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1302
apply (simp_all add: le_matrix_def less_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1303
apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1304
apply (drule_tac x=j in spec, drule_tac x=j in spec)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1305
apply (drule_tac x=i in spec, drule_tac x=i in spec)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1306
apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1307
apply (simp add: Rep_matrix_inject[THEN sym])
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1308
apply (rule ext)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1309
apply (drule_tac x=xa in spec, drule_tac x=xa in spec)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1310
apply (drule_tac x=xb in spec, drule_tac x=xb in spec)
28637
7aabaf1ba263 reactivated HOL-Matrix;
wenzelm
parents: 28562
diff changeset
  1311
apply simp
7aabaf1ba263 reactivated HOL-Matrix;
wenzelm
parents: 28562
diff changeset
  1312
done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1313
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1314
lemma le_apply_matrix:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1315
  assumes
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1316
  "f 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1317
  "\<forall>x y. x <= y \<longrightarrow> f x <= f y"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1318
  "(a::('a::{ord, zero}) matrix) <= b"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1319
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1320
  "apply_matrix f a <= apply_matrix f b"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1321
  using assms by (simp add: le_matrix_def)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1322
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1323
lemma le_combine_matrix:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1324
  assumes
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1325
  "f 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1326
  "\<forall>a b c d. a <= b & c <= d \<longrightarrow> f a c <= f b d"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1327
  "A <= B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1328
  "C <= D"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1329
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1330
  "combine_matrix f A C <= combine_matrix f B D"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1331
  using assms by (simp add: le_matrix_def)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1332
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1333
lemma le_left_combine_matrix:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1334
  assumes
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1335
  "f 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1336
  "\<forall>a b c. a <= b \<longrightarrow> f c a <= f c b"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1337
  "A <= B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1338
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1339
  "combine_matrix f C A <= combine_matrix f C B"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1340
  using assms by (simp add: le_matrix_def)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1341
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1342
lemma le_right_combine_matrix:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1343
  assumes
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1344
  "f 0 0 = 0"
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1345
  "\<forall>a b c. a <= b \<longrightarrow> f a c <= f b c"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1346
  "A <= B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1347
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1348
  "combine_matrix f A C <= combine_matrix f B C"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1349
  using assms by (simp add: le_matrix_def)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1350
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1351
lemma le_transpose_matrix: "(A <= B) = (transpose_matrix A <= transpose_matrix B)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1352
  by (simp add: le_matrix_def, auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1353
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1354
lemma le_foldseq:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1355
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1356
  "\<forall>a b c d . a <= b & c <= d \<longrightarrow> f a c <= f b d"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1357
  "\<forall>i. i <= n \<longrightarrow> s i <= t i"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1358
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1359
  "foldseq f s n <= foldseq f t n"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1360
proof -
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1361
  have "\<forall>s t. (\<forall>i. i<=n \<longrightarrow> s i <= t i) \<longrightarrow> foldseq f s n <= foldseq f t n"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1362
    by (induct n) (simp_all add: assms)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1363
  then show "foldseq f s n <= foldseq f t n" using assms by simp
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1364
qed
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1365
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1366
lemma le_left_mult:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1367
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1368
  "\<forall>a b c d. a <= b & c <= d \<longrightarrow> fadd a c <= fadd b d"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1369
  "\<forall>c a b.   0 <= c & a <= b \<longrightarrow> fmul c a <= fmul c b"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1370
  "\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1371
  "\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1372
  "fadd 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1373
  "0 <= C"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1374
  "A <= B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1375
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1376
  "mult_matrix fmul fadd C A <= mult_matrix fmul fadd C B"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1377
  using assms
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1378
  apply (simp add: le_matrix_def Rep_mult_matrix)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1379
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1380
  apply (simplesubst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1381
  apply (rule le_foldseq)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1382
  apply (auto)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1383
  done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1384
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1385
lemma le_right_mult:
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1386
  assumes
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1387
  "\<forall>a b c d. a <= b & c <= d \<longrightarrow> fadd a c <= fadd b d"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1388
  "\<forall>c a b. 0 <= c & a <= b \<longrightarrow> fmul a c <= fmul b c"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1389
  "\<forall>a. fmul 0 a = 0"
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1390
  "\<forall>a. fmul a 0 = 0"
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1391
  "fadd 0 0 = 0"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1392
  "0 <= C"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1393
  "A <= B"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1394
  shows
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1395
  "mult_matrix fmul fadd A C <= mult_matrix fmul fadd B C"
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1396
  using assms
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1397
  apply (simp add: le_matrix_def Rep_mult_matrix)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1398
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1399
  apply (simplesubst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1400
  apply (rule le_foldseq)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1401
  apply (auto)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1402
  done
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1403
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1404
lemma spec2: "\<forall>j i. P j i \<Longrightarrow> P j i" by blast
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1405
lemma neg_imp: "(\<not> Q \<longrightarrow> \<not> P) \<Longrightarrow> P \<longrightarrow> Q" by blast
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1406
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1407
lemma singleton_matrix_le[simp]: "(singleton_matrix j i a <= singleton_matrix j i b) = (a <= (b::_::order))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1408
  by (auto simp add: le_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1409
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1410
lemma singleton_le_zero[simp]: "(singleton_matrix j i x <= 0) = (x <= (0::'a::{order,zero}))"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1411
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1412
  apply (simp add: le_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1413
  apply (drule_tac j=j and i=i in spec2)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1414
  apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1415
  apply (simp add: le_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1416
  done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1417
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1418
lemma singleton_ge_zero[simp]: "(0 <= singleton_matrix j i x) = ((0::'a::{order,zero}) <= x)"
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1419
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1420
  apply (simp add: le_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1421
  apply (drule_tac j=j and i=i in spec2)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1422
  apply (simp)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1423
  apply (simp add: le_matrix_def)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1424
  done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1425
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1426
lemma move_matrix_le_zero[simp]: "0 <= j \<Longrightarrow> 0 <= i \<Longrightarrow> (move_matrix A j i <= 0) = (A <= (0::('a::{order,zero}) matrix))"
46702
202a09ba37d8 avoid using constant Int.neg
huffman
parents: 45694
diff changeset
  1427
  apply (auto simp add: le_matrix_def)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1428
  apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1429
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1430
  done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1431
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1432
lemma move_matrix_zero_le[simp]: "0 <= j \<Longrightarrow> 0 <= i \<Longrightarrow> (0 <= move_matrix A j i) = ((0::('a::{order,zero}) matrix) <= A)"
46702
202a09ba37d8 avoid using constant Int.neg
huffman
parents: 45694
diff changeset
  1433
  apply (auto simp add: le_matrix_def)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1434
  apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1435
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1436
  done
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1437
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1438
lemma move_matrix_le_move_matrix_iff[simp]: "0 <= j \<Longrightarrow> 0 <= i \<Longrightarrow> (move_matrix A j i <= move_matrix B j i) = (A <= (B::('a::{order,zero}) matrix))"
46702
202a09ba37d8 avoid using constant Int.neg
huffman
parents: 45694
diff changeset
  1439
  apply (auto simp add: le_matrix_def)
27484
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1440
  apply (drule_tac j="ja+(nat j)" and i="ia+(nat i)" in spec2)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1441
  apply (auto)
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1442
  done  
dbb9981c3d18 added marginal setup for code generation
haftmann
parents: 25764
diff changeset
  1443
27580
haftmann
parents: 27484
diff changeset
  1444
instantiation matrix :: ("{lattice, zero}") lattice
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1445
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1446
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 35818
diff changeset
  1447
definition "inf = combine_matrix inf"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1448
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 35818
diff changeset
  1449
definition "sup = combine_matrix sup"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1450
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1451
instance
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61076
diff changeset
  1452
  by standard (auto simp add: le_infI le_matrix_def inf_matrix_def sup_matrix_def)
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 22422
diff changeset
  1453
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1454
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1455
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1456
instantiation matrix :: ("{plus, zero}") plus
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1457
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1458
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1459
definition
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
  1460
  plus_matrix_def: "A + B = combine_matrix (+) A B"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1461
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1462
instance ..
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1463
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1464
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1465
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1466
instantiation matrix :: ("{uminus, zero}") uminus
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1467
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1468
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1469
definition
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 35818
diff changeset
  1470
  minus_matrix_def: "- A = apply_matrix uminus A"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1471
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1472
instance ..
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1473
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1474
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1475
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1476
instantiation matrix :: ("{minus, zero}") minus
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1477
begin
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1478
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1479
definition
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
  1480
  diff_matrix_def: "A - B = combine_matrix (-) A B"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1481
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1482
instance ..
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1483
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1484
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1485
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1486
instantiation matrix :: ("{plus, times, zero}") times
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1487
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1488
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1489
definition
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67613
diff changeset
  1490
  times_matrix_def: "A * B = mult_matrix ((*)) (+) A B"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1491
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1492
instance ..
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1493
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1494
end
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1495
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1496
instantiation matrix :: ("{lattice, uminus, zero}") abs
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1497
begin
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1498
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1499
definition
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61824
diff changeset
  1500
  abs_matrix_def: "\<bar>A :: 'a matrix\<bar> = sup A (- A)"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1501
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1502
instance ..
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1503
878c37886eed removed some legacy instantiations
haftmann
parents: 25502
diff changeset
  1504
end
23879
4776af8be741 split class abs from class minus
haftmann
parents: 23477
diff changeset
  1505
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1506
instance matrix :: (monoid_add) monoid_add
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1507
proof
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1508
  fix A B C :: "'a matrix"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1509
  show "A + B + C = A + (B + C)"    
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1510
    apply (simp add: plus_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1511
    apply (rule combine_matrix_assoc[simplified associative_def, THEN spec, THEN spec, THEN spec])
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 54864
diff changeset
  1512
    apply (simp_all add: add.assoc)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1513
    done
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1514
  show "0 + A = A"
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1515
    apply (simp add: plus_matrix_def)
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1516
    apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec])
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1517
    apply (simp)
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1518
    done
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1519
  show "A + 0 = A"
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1520
    apply (simp add: plus_matrix_def)
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1521
    apply (rule combine_matrix_zero_r_neutral [simplified zero_r_neutral_def, THEN spec])
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1522
    apply (simp)
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1523
    done
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1524
qed
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1525
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1526
instance matrix :: (comm_monoid_add) comm_monoid_add
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1527
proof
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1528
  fix A B :: "'a matrix"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1529
  show "A + B = B + A"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1530
    apply (simp add: plus_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1531
    apply (rule combine_matrix_commute[simplified commutative_def, THEN spec, THEN spec])
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 54864
diff changeset
  1532
    apply (simp_all add: add.commute)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1533
    done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1534
  show "0 + A = A"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1535
    apply (simp add: plus_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1536
    apply (rule combine_matrix_zero_l_neutral[simplified zero_l_neutral_def, THEN spec])
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1537
    apply (simp)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1538
    done
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1539
qed
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1540
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1541
instance matrix :: (group_add) group_add
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1542
proof
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1543
  fix A B :: "'a matrix"
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1544
  show "- A + A = 0" 
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1545
    by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 50027
diff changeset
  1546
  show "A + - B = A - B"
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 50027
diff changeset
  1547
    by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject [symmetric] ext)
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1548
qed
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1549
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1550
instance matrix :: (ab_group_add) ab_group_add
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1551
proof
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1552
  fix A B :: "'a matrix"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1553
  show "- A + A = 0" 
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1554
    by (simp add: plus_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1555
  show "A - B = A + - B" 
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1556
    by (simp add: plus_matrix_def diff_matrix_def minus_matrix_def Rep_matrix_inject[symmetric] ext)
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1557
qed
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1558
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34872
diff changeset
  1559
instance matrix :: (ordered_ab_group_add) ordered_ab_group_add
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1560
proof
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1561
  fix A B C :: "'a matrix"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1562
  assume "A <= B"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1563
  then show "C + A <= C + B"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1564
    apply (simp add: plus_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1565
    apply (rule le_left_combine_matrix)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1566
    apply (simp_all)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1567
    done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1568
qed
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1569
  
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34872
diff changeset
  1570
instance matrix :: (lattice_ab_group_add) semilattice_inf_ab_group_add ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34872
diff changeset
  1571
instance matrix :: (lattice_ab_group_add) semilattice_sup_ab_group_add ..
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1572
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1573
instance matrix :: (semiring_0) semiring_0
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1574
proof
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1575
  fix A B C :: "'a matrix"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1576
  show "A * B * C = A * (B * C)"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1577
    apply (simp add: times_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1578
    apply (rule mult_matrix_assoc)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28637
diff changeset
  1579
    apply (simp_all add: associative_def algebra_simps)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1580
    done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1581
  show "(A + B) * C = A * C + B * C"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1582
    apply (simp add: times_matrix_def plus_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1583
    apply (rule l_distributive_matrix[simplified l_distributive_def, THEN spec, THEN spec, THEN spec])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28637
diff changeset
  1584
    apply (simp_all add: associative_def commutative_def algebra_simps)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1585
    done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1586
  show "A * (B + C) = A * B + A * C"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1587
    apply (simp add: times_matrix_def plus_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1588
    apply (rule r_distributive_matrix[simplified r_distributive_def, THEN spec, THEN spec, THEN spec])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28637
diff changeset
  1589
    apply (simp_all add: associative_def commutative_def algebra_simps)
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1590
    done
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1591
  show "0 * A = 0" by (simp add: times_matrix_def)
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1592
  show "A * 0 = 0" by (simp add: times_matrix_def)
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1593
qed
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1594
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1595
instance matrix :: (ring) ring ..
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1596
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34872
diff changeset
  1597
instance matrix :: (ordered_ring) ordered_ring
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1598
proof
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1599
  fix A B C :: "'a matrix"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1600
  assume a: "A \<le> B"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1601
  assume b: "0 \<le> C"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1602
  from a b show "C * A \<le> C * B"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1603
    apply (simp add: times_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1604
    apply (rule le_left_mult)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1605
    apply (simp_all add: add_mono mult_left_mono)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1606
    done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1607
  from a b show "A * C \<le> B * C"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1608
    apply (simp add: times_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1609
    apply (rule le_right_mult)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1610
    apply (simp_all add: add_mono mult_right_mono)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1611
    done
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1612
qed
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1613
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34872
diff changeset
  1614
instance matrix :: (lattice_ring) lattice_ring
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1615
proof
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34872
diff changeset
  1616
  fix A B C :: "('a :: lattice_ring) matrix"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61824
diff changeset
  1617
  show "\<bar>A\<bar> = sup A (-A)" 
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1618
    by (simp add: abs_matrix_def)
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1619
qed
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1620
25303
0699e20feabd renamed lordered_*_* to lordered_*_add_*; further localization
haftmann
parents: 23879
diff changeset
  1621
lemma Rep_matrix_add[simp]:
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1622
  "Rep_matrix ((a::('a::monoid_add)matrix)+b) j i  = (Rep_matrix a j i) + (Rep_matrix b j i)"
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1623
  by (simp add: plus_matrix_def)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1624
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1625
lemma Rep_matrix_mult: "Rep_matrix ((a::('a::semiring_0) matrix) * b) j i = 
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
  1626
  foldseq (+) (% k.  (Rep_matrix a j k) * (Rep_matrix b k i)) (max (ncols a) (nrows b))"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1627
apply (simp add: times_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1628
apply (simp add: Rep_mult_matrix)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1629
done
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1630
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1631
lemma apply_matrix_add: "\<forall>x y. f (x+y) = (f x) + (f y) \<Longrightarrow> f 0 = (0::'a)
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1632
  \<Longrightarrow> apply_matrix f ((a::('a::monoid_add) matrix) + b) = (apply_matrix f a) + (apply_matrix f b)"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1633
apply (subst Rep_matrix_inject[symmetric])
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1634
apply (rule ext)+
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1635
apply (simp)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1636
done
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1637
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1638
lemma singleton_matrix_add: "singleton_matrix j i ((a::_::monoid_add)+b) = (singleton_matrix j i a) + (singleton_matrix j i b)"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1639
apply (subst Rep_matrix_inject[symmetric])
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1640
apply (rule ext)+
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1641
apply (simp)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1642
done
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1643
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1644
lemma nrows_mult: "nrows ((A::('a::semiring_0) matrix) * B) <= nrows A"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1645
by (simp add: times_matrix_def mult_nrows)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1646
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1647
lemma ncols_mult: "ncols ((A::('a::semiring_0) matrix) * B) <= ncols B"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1648
by (simp add: times_matrix_def mult_ncols)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1649
22422
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 21312
diff changeset
  1650
definition
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 21312
diff changeset
  1651
  one_matrix :: "nat \<Rightarrow> ('a::{zero,one}) matrix" where
ee19cdb07528 stepping towards uniform lattice theory development in HOL
haftmann
parents: 21312
diff changeset
  1652
  "one_matrix n = Abs_matrix (% j i. if j = i & j < n then 1 else 0)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1653
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1654
lemma Rep_one_matrix[simp]: "Rep_matrix (one_matrix n) j i = (if (j = i & j < n) then 1 else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1655
apply (simp add: one_matrix_def)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15178
diff changeset
  1656
apply (simplesubst RepAbs_matrix)
62390
842917225d56 more canonical names
nipkow
parents: 61945
diff changeset
  1657
apply (rule exI[of _ n], simp add: if_split)+
842917225d56 more canonical names
nipkow
parents: 61945
diff changeset
  1658
by (simp add: if_split)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1659
20633
e98f59806244 renamed axclass_xxxx axclasses;
wenzelm
parents: 17915
diff changeset
  1660
lemma nrows_one_matrix[simp]: "nrows ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _")
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1661
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1662
  have "?r <= n" by (simp add: nrows_le)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1663
  moreover have "n <= ?r" by (simp add:le_nrows, arith)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1664
  ultimately show "?r = n" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1665
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1666
20633
e98f59806244 renamed axclass_xxxx axclasses;
wenzelm
parents: 17915
diff changeset
  1667
lemma ncols_one_matrix[simp]: "ncols ((one_matrix n) :: ('a::zero_neq_one)matrix) = n" (is "?r = _")
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1668
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1669
  have "?r <= n" by (simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1670
  moreover have "n <= ?r" by (simp add: le_ncols, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1671
  ultimately show "?r = n" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1672
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1673
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1674
lemma one_matrix_mult_right[simp]: "ncols A <= n \<Longrightarrow> (A::('a::{semiring_1}) matrix) * (one_matrix n) = A"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1675
apply (subst Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1676
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1677
apply (simp add: times_matrix_def Rep_mult_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1678
apply (rule_tac j1="xa" in ssubst[OF foldseq_almostzero])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1679
apply (simp_all)
32440
153965be0f4b tuned proofs
nipkow
parents: 29700
diff changeset
  1680
by (simp add: ncols)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1681
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1682
lemma one_matrix_mult_left[simp]: "nrows A <= n \<Longrightarrow> (one_matrix n) * A = (A::('a::semiring_1) matrix)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1683
apply (subst Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1684
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1685
apply (simp add: times_matrix_def Rep_mult_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1686
apply (rule_tac j1="x" in ssubst[OF foldseq_almostzero])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1687
apply (simp_all)
32440
153965be0f4b tuned proofs
nipkow
parents: 29700
diff changeset
  1688
by (simp add: nrows)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1689
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1690
lemma transpose_matrix_mult: "transpose_matrix ((A::('a::comm_ring) matrix)*B) = (transpose_matrix B) * (transpose_matrix A)"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1691
apply (simp add: times_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1692
apply (subst transpose_mult_matrix)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 54864
diff changeset
  1693
apply (simp_all add: mult.commute)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1694
done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1695
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1696
lemma transpose_matrix_add: "transpose_matrix ((A::('a::monoid_add) matrix)+B) = transpose_matrix A + transpose_matrix B"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1697
by (simp add: plus_matrix_def transpose_combine_matrix)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1698
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1699
lemma transpose_matrix_diff: "transpose_matrix ((A::('a::group_add) matrix)-B) = transpose_matrix A - transpose_matrix B"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1700
by (simp add: diff_matrix_def transpose_combine_matrix)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1701
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1702
lemma transpose_matrix_minus: "transpose_matrix (-(A::('a::group_add) matrix)) = - transpose_matrix (A::'a matrix)"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1703
by (simp add: minus_matrix_def transpose_apply_matrix)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1704
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1705
definition right_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1706
  "right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X))) \<and> nrows X \<le> ncols A" 
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1707
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1708
definition left_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1709
  "left_inverse_matrix A X == (X * A = one_matrix (max(nrows X) (ncols A))) \<and> ncols X \<le> nrows A" 
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1710
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1711
definition inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1712
  "inverse_matrix A X == (right_inverse_matrix A X) \<and> (left_inverse_matrix A X)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1713
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1714
lemma right_inverse_matrix_dim: "right_inverse_matrix A X \<Longrightarrow> nrows A = ncols X"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1715
apply (insert ncols_mult[of A X], insert nrows_mult[of A X])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1716
by (simp add: right_inverse_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1717
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1718
lemma left_inverse_matrix_dim: "left_inverse_matrix A Y \<Longrightarrow> ncols A = nrows Y"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1719
apply (insert ncols_mult[of Y A], insert nrows_mult[of Y A]) 
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1720
by (simp add: left_inverse_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1721
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1722
lemma left_right_inverse_matrix_unique: 
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1723
  assumes "left_inverse_matrix A Y" "right_inverse_matrix A X"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1724
  shows "X = Y"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1725
proof -
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1726
  have "Y = Y * one_matrix (nrows A)" 
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1727
    apply (subst one_matrix_mult_right)
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1728
    using assms
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1729
    apply (simp_all add: left_inverse_matrix_def)
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1730
    done
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1731
  also have "\<dots> = Y * (A * X)" 
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1732
    apply (insert assms)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1733
    apply (frule right_inverse_matrix_dim)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1734
    by (simp add: right_inverse_matrix_def)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 54864
diff changeset
  1735
  also have "\<dots> = (Y * A) * X" by (simp add: mult.assoc)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1736
  also have "\<dots> = X" 
35612
0a9fb49a086d eliminated old-style prems;
wenzelm
parents: 35416
diff changeset
  1737
    apply (insert assms)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1738
    apply (frule left_inverse_matrix_dim)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1739
    apply (simp_all add:  left_inverse_matrix_def right_inverse_matrix_def one_matrix_mult_left)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1740
    done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1741
  ultimately show "X = Y" by (simp)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1742
qed
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1743
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1744
lemma inverse_matrix_inject: "\<lbrakk> inverse_matrix A X; inverse_matrix A Y \<rbrakk> \<Longrightarrow> X = Y"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1745
  by (auto simp add: inverse_matrix_def left_right_inverse_matrix_unique)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1746
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1747
lemma one_matrix_inverse: "inverse_matrix (one_matrix n) (one_matrix n)"
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1748
  by (simp add: inverse_matrix_def left_inverse_matrix_def right_inverse_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1749
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1750
lemma zero_imp_mult_zero: "(a::'a::semiring_0) = 0 | b = 0 \<Longrightarrow> a * b = 0"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1751
by auto
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1752
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1753
lemma Rep_matrix_zero_imp_mult_zero:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67399
diff changeset
  1754
  "\<forall>j i k. (Rep_matrix A j k = 0) | (Rep_matrix B k i) = 0  \<Longrightarrow> A * B = (0::('a::lattice_ring) matrix)"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1755
apply (subst Rep_matrix_inject[symmetric])
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1756
apply (rule ext)+
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1757
apply (auto simp add: Rep_matrix_mult foldseq_zero zero_imp_mult_zero)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1758
done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1759
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1760
lemma add_nrows: "nrows (A::('a::monoid_add) matrix) <= u \<Longrightarrow> nrows B <= u \<Longrightarrow> nrows (A + B) <= u"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1761
apply (simp add: plus_matrix_def)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1762
apply (rule combine_nrows)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1763
apply (simp_all)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1764
done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1765
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1766
lemma move_matrix_row_mult: "move_matrix ((A::('a::semiring_0) matrix) * B) j 0 = (move_matrix A j 0) * B"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1767
apply (subst Rep_matrix_inject[symmetric])
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1768
apply (rule ext)+
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1769
apply (auto simp add: Rep_matrix_mult foldseq_zero)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1770
apply (rule_tac foldseq_zerotail[symmetric])
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1771
apply (auto simp add: nrows zero_imp_mult_zero max2)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1772
apply (rule order_trans)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1773
apply (rule ncols_move_matrix_le)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1774
apply (simp add: max1)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1775
done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1776
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1777
lemma move_matrix_col_mult: "move_matrix ((A::('a::semiring_0) matrix) * B) 0 i = A * (move_matrix B 0 i)"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1778
apply (subst Rep_matrix_inject[symmetric])
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1779
apply (rule ext)+
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1780
apply (auto simp add: Rep_matrix_mult foldseq_zero)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1781
apply (rule_tac foldseq_zerotail[symmetric])
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1782
apply (auto simp add: ncols zero_imp_mult_zero max1)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1783
apply (rule order_trans)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1784
apply (rule nrows_move_matrix_le)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1785
apply (simp add: max2)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1786
done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1787
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1788
lemma move_matrix_add: "((move_matrix (A + B) j i)::(('a::monoid_add) matrix)) = (move_matrix A j i) + (move_matrix B j i)" 
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1789
apply (subst Rep_matrix_inject[symmetric])
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1790
apply (rule ext)+
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1791
apply (simp)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1792
done
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1793
34872
6ca970cfa873 Matrices form a semiring with 0
hoelzl
parents: 33657
diff changeset
  1794
lemma move_matrix_mult: "move_matrix ((A::('a::semiring_0) matrix)*B) j i = (move_matrix A j 0) * (move_matrix B 0 i)"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1795
by (simp add: move_matrix_ortho[of "A*B"] move_matrix_col_mult move_matrix_row_mult)
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1796
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35032
diff changeset
  1797
definition scalar_mult :: "('a::ring) \<Rightarrow> 'a matrix \<Rightarrow> 'a matrix" where
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67613
diff changeset
  1798
  "scalar_mult a m == apply_matrix ((*) a) m"
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1799
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1800
lemma scalar_mult_zero[simp]: "scalar_mult y 0 = 0" 
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
  1801
by (simp add: scalar_mult_def)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1802
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1803
lemma scalar_mult_add: "scalar_mult y (a+b) = (scalar_mult y a) + (scalar_mult y b)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28637
diff changeset
  1804
by (simp add: scalar_mult_def apply_matrix_add algebra_simps)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1805
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1806
lemma Rep_scalar_mult[simp]: "Rep_matrix (scalar_mult y a) j i = y * (Rep_matrix a j i)" 
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
  1807
by (simp add: scalar_mult_def)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1808
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1809
lemma scalar_mult_singleton[simp]: "scalar_mult y (singleton_matrix j i x) = singleton_matrix j i (y * x)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
  1810
apply (subst Rep_matrix_inject[symmetric])
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
  1811
apply (rule ext)+
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
  1812
apply (auto)
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
  1813
done
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1814
27653
180e28bab764 more class instantiations
haftmann
parents: 27580
diff changeset
  1815
lemma Rep_minus[simp]: "Rep_matrix (-(A::_::group_add)) x y = - (Rep_matrix A x y)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
  1816
by (simp add: minus_matrix_def)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1817
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61824
diff changeset
  1818
lemma Rep_abs[simp]: "Rep_matrix \<bar>A::_::lattice_ab_group_add\<bar> x y = \<bar>Rep_matrix A x y\<bar>"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 22452
diff changeset
  1819
by (simp add: abs_lattice sup_matrix_def)
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14738
diff changeset
  1820
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1821
end