src/HOL/Matrix/MatrixGeneral.thy
author schirmer
Mon, 03 May 2004 23:22:17 +0200
changeset 14700 2f885b7e5ba7
parent 14691 e1eedc8cad37
child 14940 b9ab8babd8b3
permissions -rw-r--r--
reimplementation of HOL records; only one type is created for each record extension, instead of one type for each field. See NEWS.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     1
(*  Title:      HOL/Matrix/MatrixGeneral.thy
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     2
    ID:         $Id$
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     3
    Author:     Steven Obua
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
     4
    License:    2004 Technische Universitaet Muenchen
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     5
*)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     6
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     7
theory MatrixGeneral = Main:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     8
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
     9
types 'a infmatrix = "[nat, nat] \<Rightarrow> 'a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    10
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    11
constdefs
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    12
  nonzero_positions :: "('a::zero) infmatrix \<Rightarrow> (nat*nat) set"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    13
  "nonzero_positions A == {pos. A (fst pos) (snd pos) ~= 0}"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    14
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    15
typedef 'a matrix = "{(f::(('a::zero) infmatrix)). finite (nonzero_positions f)}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    16
apply (rule_tac x="(% j i. 0)" in exI)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    17
by (simp add: nonzero_positions_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    18
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    19
declare Rep_matrix_inverse[simp]
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    20
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    21
lemma finite_nonzero_positions : "finite (nonzero_positions (Rep_matrix A))"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    22
apply (rule Abs_matrix_induct)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    23
by (simp add: Abs_matrix_inverse matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    24
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    25
constdefs
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    26
  nrows :: "('a::zero) matrix \<Rightarrow> nat"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    27
  "nrows A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image fst) (nonzero_positions (Rep_matrix A))))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    28
  ncols :: "('a::zero) matrix \<Rightarrow> nat"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    29
  "ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    30
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    31
lemma nrows:
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    32
  assumes hyp: "nrows A \<le> m"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    33
  shows "(Rep_matrix A m n) = 0" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    34
proof cases
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    35
  assume "nonzero_positions(Rep_matrix A) = {}"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    36
  then show "(Rep_matrix A m n) = 0" by (simp add: nonzero_positions_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    37
next
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    38
  assume a: "nonzero_positions(Rep_matrix A) \<noteq> {}"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    39
  let ?S = "fst`(nonzero_positions(Rep_matrix A))"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    40
  from a have b: "?S \<noteq> {}" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    41
  have c: "finite (?S)" by (simp add: finite_nonzero_positions)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    42
  from hyp have d: "Max (?S) < m" by (simp add: a nrows_def)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    43
  have "m \<notin> ?S"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    44
    proof -
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    45
      have "m \<in> ?S \<Longrightarrow> m <= Max(?S)" by (simp add: Max[OF c b])
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    46
      moreover from d have "~(m <= Max ?S)" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    47
      ultimately show "m \<notin> ?S" by (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    48
    qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    49
  thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    50
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    51
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    52
constdefs
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    53
  transpose_infmatrix :: "'a infmatrix \<Rightarrow> 'a infmatrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    54
  "transpose_infmatrix A j i == A i j"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    55
  transpose_matrix :: "('a::zero) matrix \<Rightarrow> 'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    56
  "transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    57
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    58
declare transpose_infmatrix_def[simp]
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    59
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    60
lemma transpose_infmatrix_twice[simp]: "transpose_infmatrix (transpose_infmatrix A) = A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    61
by ((rule ext)+, simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    62
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    63
lemma transpose_infmatrix: "transpose_infmatrix (% j i. P j i) = (% j i. P i j)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    64
  apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    65
  by (simp add: transpose_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    66
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    67
lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    68
apply (rule Abs_matrix_inverse)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    69
apply (simp add: matrix_def nonzero_positions_def image_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    70
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    71
  let ?A = "{pos. Rep_matrix x (snd pos) (fst pos) \<noteq> 0}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    72
  let ?swap = "% pos. (snd pos, fst pos)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    73
  let ?B = "{pos. Rep_matrix x (fst pos) (snd pos) \<noteq> 0}"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    74
  have swap_image: "?swap`?A = ?B"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    75
    apply (simp add: image_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    76
    apply (rule set_ext)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    77
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    78
    proof
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    79
      fix y
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    80
      assume hyp: "\<exists>a b. Rep_matrix x b a \<noteq> 0 \<and> y = (b, a)"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    81
      thus "Rep_matrix x (fst y) (snd y) \<noteq> 0"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    82
        proof -
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    83
          from hyp obtain a b where "(Rep_matrix x b a \<noteq> 0 & y = (b,a))" by blast
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    84
          then show "Rep_matrix x (fst y) (snd y) \<noteq> 0" by (simp)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    85
        qed
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    86
    next
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    87
      fix y
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    88
      assume hyp: "Rep_matrix x (fst y) (snd y) \<noteq> 0"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    89
      show "\<exists> a b. (Rep_matrix x b a \<noteq> 0 & y = (b,a))" by (rule exI[of _ "snd y"], rule exI[of _ "fst y"], simp)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    90
    qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    91
  then have "finite (?swap`?A)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    92
    proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    93
      have "finite (nonzero_positions (Rep_matrix x))" by (simp add: finite_nonzero_positions)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    94
      then have "finite ?B" by (simp add: nonzero_positions_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    95
      with swap_image show "finite (?swap`?A)" by (simp)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    96
    qed
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    97
  moreover
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
    98
  have "inj_on ?swap ?A" by (simp add: inj_on_def)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
    99
  ultimately show "finite ?A"by (rule finite_imageD[of ?swap ?A])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   100
qed
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   101
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   102
lemma transpose_matrix[simp]: "Rep_matrix(transpose_matrix A) j i = Rep_matrix A i j"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   103
by (simp add: transpose_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   104
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   105
lemma transpose_transpose_id[simp]: "transpose_matrix (transpose_matrix A) = A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   106
by (simp add: transpose_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   107
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   108
lemma nrows_transpose[simp]: "nrows (transpose_matrix A) = ncols A"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   109
by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   110
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   111
lemma ncols_transpose[simp]: "ncols (transpose_matrix A) = nrows A"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   112
by (simp add: nrows_def ncols_def nonzero_positions_def transpose_matrix_def image_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   113
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   114
lemma ncols: "ncols A <= n \<Longrightarrow> Rep_matrix A m n = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   115
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   116
  assume "ncols A <= n"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   117
  then have "nrows (transpose_matrix A) <= n" by (simp)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   118
  then have "Rep_matrix (transpose_matrix A) n m = 0" by (rule nrows)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   119
  thus "Rep_matrix A m n = 0" by (simp add: transpose_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   120
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   121
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   122
lemma ncols_le: "(ncols A <= n) = (! j i. n <= i \<longrightarrow> (Rep_matrix A j i) = 0)" (is "_ = ?st")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   123
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   124
apply (simp add: ncols)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   125
proof (simp add: ncols_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   126
  let ?P = "nonzero_positions (Rep_matrix A)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   127
  let ?p = "snd`?P"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   128
  have a:"finite ?p" by (simp add: finite_nonzero_positions)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   129
  let ?m = "Max ?p"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   130
  assume "~(Suc (?m) <= n)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   131
  then have b:"n <= ?m" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   132
  fix a b
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   133
  assume "(a,b) \<in> ?P"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   134
  then have "?p \<noteq> {}" by (auto)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   135
  with a have "?m \<in>  ?p" by (simp)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   136
  moreover have "!x. (x \<in> ?p \<longrightarrow> (? y. (Rep_matrix A y x) \<noteq> 0))" by (simp add: nonzero_positions_def image_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   137
  ultimately have "? y. (Rep_matrix A y ?m) \<noteq> 0" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   138
  moreover assume ?st
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   139
  ultimately show "False" using b by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   140
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   141
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   142
lemma less_ncols: "(n < ncols A) = (? j i. n <= i & (Rep_matrix A j i) \<noteq> 0)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   143
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   144
  have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   145
  show ?concl by (simp add: a ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   146
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   147
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   148
lemma le_ncols: "(n <= ncols A) = (\<forall> m. (\<forall> j i. m <= i \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n <= m)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   149
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   150
apply (subgoal_tac "ncols A <= m")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   151
apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   152
apply (simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   153
apply (drule_tac x="ncols A" in spec)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   154
by (simp add: ncols)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   155
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   156
lemma nrows_le: "(nrows A <= n) = (! j i. n <= j \<longrightarrow> (Rep_matrix A j i) = 0)" (is ?s)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   157
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   158
  have "(nrows A <= n) = (ncols (transpose_matrix A) <= n)" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   159
  also have "\<dots> = (! j i. n <= i \<longrightarrow> (Rep_matrix (transpose_matrix A) j i = 0))" by (rule ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   160
  also have "\<dots> = (! j i. n <= i \<longrightarrow> (Rep_matrix A i j) = 0)" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   161
  finally show "(nrows A <= n) = (! j i. n <= j \<longrightarrow> (Rep_matrix A j i) = 0)" by (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   162
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   163
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   164
lemma less_nrows: "(m < nrows A) = (? j i. m <= j & (Rep_matrix A j i) \<noteq> 0)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   165
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   166
  have a: "!! (a::nat) b. (a < b) = (~(b <= a))" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   167
  show ?concl by (simp add: a nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   168
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   169
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   170
lemma le_nrows: "(n <= nrows A) = (\<forall> m. (\<forall> j i. m <= j \<longrightarrow> (Rep_matrix A j i) = 0) \<longrightarrow> n <= m)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   171
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   172
apply (subgoal_tac "nrows A <= m")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   173
apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   174
apply (simp add: nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   175
apply (drule_tac x="nrows A" in spec)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   176
by (simp add: nrows)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   177
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   178
lemma finite_natarray1: "finite {x. x < (n::nat)}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   179
apply (induct n)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   180
apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   181
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   182
  fix n
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   183
  have "{x. x < Suc n} = insert n {x. x < n}"  by (rule set_ext, simp, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   184
  moreover assume "finite {x. x < n}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   185
  ultimately show "finite {x. x < Suc n}" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   186
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   187
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   188
lemma finite_natarray2: "finite {pos. (fst pos) < (m::nat) & (snd pos) < (n::nat)}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   189
  apply (induct m)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   190
  apply (simp+)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   191
  proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   192
    fix m::nat
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   193
    let ?s0 = "{pos. fst pos < m & snd pos < n}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   194
    let ?s1 = "{pos. fst pos < (Suc m) & snd pos < n}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   195
    let ?sd = "{pos. fst pos = m & snd pos < n}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   196
    assume f0: "finite ?s0"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   197
    have f1: "finite ?sd"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   198
    proof -
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   199
      let ?f = "% x. (m, x)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   200
      have "{pos. fst pos = m & snd pos < n} = ?f ` {x. x < n}" by (rule set_ext, simp add: image_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   201
      moreover have "finite {x. x < n}" by (simp add: finite_natarray1)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   202
      ultimately show "finite {pos. fst pos = m & snd pos < n}" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   203
    qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   204
    have su: "?s0 \<union> ?sd = ?s1" by (rule set_ext, simp, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   205
    from f0 f1 have "finite (?s0 \<union> ?sd)" by (rule finite_UnI)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   206
    with su show "finite ?s1" by (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   207
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   208
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   209
lemma RepAbs_matrix:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   210
  assumes aem: "? m. ! j i. m <= j \<longrightarrow> x j i = 0" (is ?em) and aen:"? n. ! j i. (n <= i \<longrightarrow> x j i = 0)" (is ?en)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   211
  shows "(Rep_matrix (Abs_matrix x)) = x"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   212
apply (rule Abs_matrix_inverse)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   213
apply (simp add: matrix_def nonzero_positions_def)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   214
proof -
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   215
  from aem obtain m where a: "! j i. m <= j \<longrightarrow> x j i = 0" by (blast)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   216
  from aen obtain n where b: "! j i. n <= i \<longrightarrow> x j i = 0" by (blast)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   217
  let ?u = "{pos. x (fst pos) (snd pos) \<noteq> 0}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   218
  let ?v = "{pos. fst pos < m & snd pos < n}"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   219
  have c: "!! (m::nat) a. ~(m <= a) \<Longrightarrow> a < m" by (arith)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   220
  from a b have "(?u \<inter> (-?v)) = {}"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   221
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   222
    apply (rule set_ext)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   223
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   224
    apply auto
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   225
    by (rule c, auto)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   226
  then have d: "?u \<subseteq> ?v" by blast
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   227
  moreover have "finite ?v" by (simp add: finite_natarray2)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   228
  ultimately show "finite ?u" by (rule finite_subset)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   229
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   230
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   231
constdefs
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   232
  apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   233
  "apply_infmatrix f == % A. (% j i. f (A j i))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   234
  apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   235
  "apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   236
  combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   237
  "combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   238
  combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('c::zero) matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   239
  "combine_matrix f == % A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   240
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   241
lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   242
by (simp add: apply_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   243
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   244
lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   245
by (simp add: combine_infmatrix_def)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   246
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   247
constdefs
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   248
commutative :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   249
"commutative f == ! x y. f x y = f y x"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   250
associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   251
"associative f == ! x y z. f (f x y) z = f x (f y z)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   252
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   253
text{*
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   254
To reason about associativity and commutativity of operations on matrices,
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   255
let's take a step back and look at the general situtation: Assume that we have
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   256
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.
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   257
Each function $f: A \times A \rightarrow A$ now induces a function $f': B \times B \rightarrow B$ by $f' = u \circ f$.
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   258
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.$
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   259
*}
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   260
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   261
lemma combine_infmatrix_commute:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   262
  "commutative f \<Longrightarrow> commutative (combine_infmatrix f)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   263
by (simp add: commutative_def combine_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   264
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   265
lemma combine_matrix_commute:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   266
"commutative f \<Longrightarrow> commutative (combine_matrix f)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   267
by (simp add: combine_matrix_def commutative_def combine_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   268
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   269
text{*
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   270
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\}$,
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   271
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
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   272
\[ f' (f' 1 1) -1 = u(f (u (f 1 1)) -1) = u(f (u 2) -1) = u (f 0 -1) = -1, \]
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   273
but on the other hand we have
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   274
\[ f' 1 (f' 1 -1) = u (f 1 (u (f 1 -1))) = u (f 1 0) = 1.\]
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   275
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:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   276
*}
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   277
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   278
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)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   279
by (rule subsetI, simp add: nonzero_positions_def combine_infmatrix_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   280
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   281
lemma finite_nonzero_positions_Rep[simp]: "finite (nonzero_positions (Rep_matrix A))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   282
by (insert Rep_matrix [of A], simp add: matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   283
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   284
lemma combine_infmatrix_closed [simp]:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   285
  "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)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   286
apply (rule Abs_matrix_inverse)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   287
apply (simp add: matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   288
apply (rule finite_subset[of _ "(nonzero_positions (Rep_matrix A)) \<union> (nonzero_positions (Rep_matrix B))"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   289
by (simp_all)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   290
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   291
text {* We need the next two lemmas only later, but it is analog to the above one, so we prove them now: *}
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   292
lemma nonzero_positions_apply_infmatrix[simp]: "f 0 = 0 \<Longrightarrow> nonzero_positions (apply_infmatrix f A) \<subseteq> nonzero_positions A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   293
by (rule subsetI, simp add: nonzero_positions_def apply_infmatrix_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   294
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   295
lemma apply_infmatrix_closed [simp]:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   296
  "f 0 = 0 \<Longrightarrow> Rep_matrix (Abs_matrix (apply_infmatrix f (Rep_matrix A))) = apply_infmatrix f (Rep_matrix A)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   297
apply (rule Abs_matrix_inverse)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   298
apply (simp add: matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   299
apply (rule finite_subset[of _ "nonzero_positions (Rep_matrix A)"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   300
by (simp_all)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   301
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   302
lemma combine_infmatrix_assoc[simp]: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_infmatrix f)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   303
by (simp add: associative_def combine_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   304
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   305
lemma comb: "f = g \<Longrightarrow> x = y \<Longrightarrow> f x = g y"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   306
by (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   307
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   308
lemma combine_matrix_assoc: "f 0 0 = 0 \<Longrightarrow> associative f \<Longrightarrow> associative (combine_matrix f)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   309
apply (simp(no_asm) add: associative_def combine_matrix_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   310
apply (rule comb [of Abs_matrix Abs_matrix])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   311
by (auto, insert combine_infmatrix_assoc[of f], simp add: associative_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   312
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   313
lemma Rep_apply_matrix[simp]: "f 0 = 0 \<Longrightarrow> Rep_matrix (apply_matrix f A) j i = f (Rep_matrix A j i)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   314
by (simp add: apply_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   315
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   316
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)"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   317
  by(simp add: combine_matrix_def)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   318
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   319
lemma combine_nrows: "f 0 0 = 0  \<Longrightarrow> nrows (combine_matrix f A B) <= max (nrows A) (nrows B)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   320
by (simp add: nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   321
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   322
lemma combine_ncols: "f 0 0 = 0 \<Longrightarrow> ncols (combine_matrix f A B) <= max (ncols A) (ncols B)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   323
by (simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   324
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   325
lemma combine_nrows: "f 0 0 = 0 \<Longrightarrow> nrows A <= q \<Longrightarrow> nrows B <= q \<Longrightarrow> nrows(combine_matrix f A B) <= q"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   326
  by (simp add: nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   327
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   328
lemma combine_ncols: "f 0 0 = 0 \<Longrightarrow> ncols A <= q \<Longrightarrow> ncols B <= q \<Longrightarrow> ncols(combine_matrix f A B) <= q"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   329
  by (simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   330
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   331
constdefs
14654
aad262a36014 fixed constdefs typing;
wenzelm
parents: 14593
diff changeset
   332
  zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   333
  "zero_r_neutral f == ! a. f a 0 = a"
14654
aad262a36014 fixed constdefs typing;
wenzelm
parents: 14593
diff changeset
   334
  zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   335
  "zero_l_neutral f == ! a. f 0 a = a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   336
  zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   337
  "zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   338
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   339
consts foldseq :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   340
primrec
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   341
  "foldseq f s 0 = s 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   342
  "foldseq f s (Suc n) = f (s 0) (foldseq f (% k. s(Suc k)) n)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   343
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   344
consts foldseq_transposed ::  "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   345
primrec
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   346
  "foldseq_transposed f s 0 = s 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   347
  "foldseq_transposed f s (Suc n) = f (foldseq_transposed f s n) (s (Suc n))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   348
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   349
lemma foldseq_assoc : "associative f \<Longrightarrow> foldseq f = foldseq_transposed f"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   350
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   351
  assume a:"associative f"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   352
  then have sublemma: "!! n. ! N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   353
  proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   354
    fix n
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   355
    show "!N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   356
    proof (induct n)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   357
      show "!N s. N <= 0 \<longrightarrow> foldseq f s N = foldseq_transposed f s N" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   358
    next
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   359
      fix n
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   360
      assume b:"! N s. N <= n \<longrightarrow> foldseq f s N = foldseq_transposed f s N"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   361
      have c:"!!N s. N <= n \<Longrightarrow> foldseq f s N = foldseq_transposed f s N" by (simp add: b)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   362
      show "! N t. N <= Suc n \<longrightarrow> foldseq f t N = foldseq_transposed f t N"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   363
      proof (auto)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   364
        fix N t
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   365
        assume Nsuc: "N <= Suc n"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   366
        show "foldseq f t N = foldseq_transposed f t N"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   367
        proof cases
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   368
          assume "N <= n"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   369
          then show "foldseq f t N = foldseq_transposed f t N" by (simp add: b)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   370
        next
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   371
          assume "~(N <= n)"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   372
          with Nsuc have Nsuceq: "N = Suc n" by simp
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   373
          have neqz: "n \<noteq> 0 \<Longrightarrow> ? m. n = Suc m & Suc m <= n" by arith
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   374
          have assocf: "!! x y z. f x (f y z) = f (f x y) z" by (insert a, simp add: associative_def)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   375
          show "foldseq f t N = foldseq_transposed f t N"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   376
            apply (simp add: Nsuceq)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   377
            apply (subst c)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   378
            apply (simp)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   379
            apply (case_tac "n = 0")
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   380
            apply (simp)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   381
            apply (drule neqz)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   382
            apply (erule exE)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   383
            apply (simp)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   384
            apply (subst assocf)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   385
            proof -
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   386
              fix m
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   387
              assume "n = Suc m & Suc m <= n"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   388
              then have mless: "Suc m <= n" by arith
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   389
              then have step1: "foldseq_transposed f (% k. t (Suc k)) m = foldseq f (% k. t (Suc k)) m" (is "?T1 = ?T2")
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   390
                apply (subst c)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   391
                by simp+
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   392
              have step2: "f (t 0) ?T2 = foldseq f t (Suc m)" (is "_ = ?T3") by simp
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   393
              have step3: "?T3 = foldseq_transposed f t (Suc m)" (is "_ = ?T4")
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   394
                apply (subst c)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   395
                by (simp add: mless)+
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   396
              have step4: "?T4 = f (foldseq_transposed f t m) (t (Suc m))" (is "_=?T5") by simp
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   397
              from step1 step2 step3 step4 show sowhat: "f (f (t 0) ?T1) (t (Suc (Suc m))) = f ?T5 (t (Suc (Suc m)))" by simp
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   398
            qed
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   399
          qed
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   400
        qed
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   401
      qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   402
    qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   403
    show "foldseq f = foldseq_transposed f" by ((rule ext)+, insert sublemma, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   404
  qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   405
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   406
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)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   407
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   408
  assume assoc: "associative f"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   409
  assume comm: "commutative f"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   410
  from assoc have a:"!! x y z. f (f x y) z = f x (f y z)" by (simp add: associative_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   411
  from comm have b: "!! x y. f x y = f y x" by (simp add: commutative_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   412
  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)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   413
  have "!! n. (! u v. foldseq f (%k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   414
    apply (induct_tac n)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   415
    apply (simp+, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   416
    by (simp add: a b c)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   417
  then show "foldseq f (% k. f (u k) (v k)) n = f (foldseq f u n) (foldseq f v n)" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   418
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   419
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   420
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); ? x y. (f x) \<noteq> (f y); ? x y. (g x) \<noteq> (g y); f x x = x; g x x = x\<rbrakk> \<Longrightarrow> f=g | (! y. f y x = y) | (! y. g y x = y)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   421
oops
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   422
(* Model found
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   423
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   424
Trying to find a model that refutes: \<lbrakk>associative f; associative g;
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   425
 \<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;
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   426
 \<exists>x y. g x \<noteq> g y; f x x = x; g x x = x\<rbrakk>
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   427
\<Longrightarrow> f = g \<or> (\<forall>y. f y x = y) \<or> (\<forall>y. g y x = y)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   428
Searching for a model of size 1, translating term... invoking SAT solver... no model found.
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   429
Searching for a model of size 2, translating term... invoking SAT solver... no model found.
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   430
Searching for a model of size 3, translating term... invoking SAT solver...
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   431
Model found:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   432
Size of types: 'a: 3
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   433
x: a1
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   434
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))
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   435
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))
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   436
*)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   437
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   438
lemma foldseq_zero:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   439
assumes fz: "f 0 0 = 0" and sz: "! i. i <= n \<longrightarrow> s i = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   440
shows "foldseq f s n = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   441
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   442
  have "!! n. ! s. (! i. i <= n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s n = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   443
    apply (induct_tac n)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   444
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   445
    by (simp add: fz)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   446
  then show "foldseq f s n = 0" by (simp add: sz)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   447
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   448
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   449
lemma foldseq_significant_positions:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   450
  assumes p: "! i. i <= N \<longrightarrow> S i = T i"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   451
  shows "foldseq f S N = foldseq f T N" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   452
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   453
  have "!! m . ! s t. (! i. i<=m \<longrightarrow> s i = t i) \<longrightarrow> foldseq f s m = foldseq f t m"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   454
    apply (induct_tac m)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   455
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   456
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   457
    apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   458
    proof -
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   459
      fix n
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   460
      fix s::"nat\<Rightarrow>'a"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   461
      fix t::"nat\<Rightarrow>'a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   462
      assume a: "\<forall>s t. (\<forall>i\<le>n. s i = t i) \<longrightarrow> foldseq f s n = foldseq f t n"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   463
      assume b: "\<forall>i\<le>Suc n. s i = t i"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   464
      have c:"!! a b. a = b \<Longrightarrow> f (t 0) a = f (t 0) b" by blast
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   465
      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)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   466
      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)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   467
    qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   468
  with p show ?concl by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   469
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   470
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   471
lemma foldseq_tail: "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" (is "?p \<Longrightarrow> ?concl")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   472
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   473
  have suc: "!! a b. \<lbrakk>a <= Suc b; a \<noteq> Suc b\<rbrakk> \<Longrightarrow> a <= b" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   474
  have suc2: "!! a b. \<lbrakk>a <= Suc b; ~ (a <= b)\<rbrakk> \<Longrightarrow> a = (Suc b)" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   475
  have eq: "!! (a::nat) b. \<lbrakk>~(a < b); a <= b\<rbrakk> \<Longrightarrow> a = b" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   476
  have a:"!! a b c . a = b \<Longrightarrow> f c a = f c b" by blast
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   477
  have "!! n. ! 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"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   478
    apply (induct_tac n)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   479
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   480
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   481
    apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   482
    apply (case_tac "m = Suc na")
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   483
    apply (simp)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   484
    apply (rule a)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   485
    apply (rule foldseq_significant_positions)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   486
    apply (simp, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   487
    apply (drule eq, simp+)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   488
    apply (drule suc, simp+)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   489
    proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   490
      fix na m s
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   491
      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"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   492
      assume subb:"m <= na"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   493
      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
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   494
      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 =
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   495
        foldseq f (% k. s(Suc k)) na"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   496
        by (rule subc[of m "% k. s(Suc k)", THEN sym], simp add: subb)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   497
      from subb have sube: "m \<noteq> 0 \<Longrightarrow> ? mm. m = Suc mm & mm <= na" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   498
      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) =
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   499
        foldseq f (\<lambda>k. if k < m then s k else foldseq f (\<lambda>k. s (k + m)) (Suc na - m)) m"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   500
        apply (simp add: subd)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   501
        apply (case_tac "m=0")
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   502
        apply (simp)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   503
        apply (drule sube)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   504
        apply (auto)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   505
        apply (rule a)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   506
        by (simp add: subc if_def)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   507
    qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   508
  then show "?p \<Longrightarrow> ?concl" by simp
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   509
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   510
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   511
lemma foldseq_zerotail:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   512
  assumes
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   513
  fz: "f 0 0 = 0"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   514
  and sz: "! i.  n <= i \<longrightarrow> s i = 0"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   515
  and nm: "n <= m"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   516
  shows
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   517
  "foldseq f s n = foldseq f s m"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   518
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   519
  have a: "!! a b. ~(a < b) \<Longrightarrow> a <= b \<Longrightarrow> (a::nat) = b" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   520
  show "foldseq f s n = foldseq f s m"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   521
    apply (simp add: foldseq_tail[OF nm, of f s])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   522
    apply (rule foldseq_significant_positions)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   523
    apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   524
    apply (drule a)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   525
    apply (simp+)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   526
    apply (subst foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   527
    by (simp add: fz sz)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   528
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   529
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   530
lemma foldseq_zerotail2:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   531
  assumes "! x. f x 0 = x"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   532
  and "! i. n < i \<longrightarrow> s i = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   533
  and nm: "n <= m"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   534
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   535
  "foldseq f s n = foldseq f s m" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   536
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   537
  have "f 0 0 = 0" by (simp add: prems)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   538
  have a:"!! (i::nat) m. ~(i < m) \<Longrightarrow> i <= m \<Longrightarrow> i = m" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   539
  have b:"!! m n. n <= m \<Longrightarrow> m \<noteq> n \<Longrightarrow> ? k. m-n = Suc k" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   540
  have c: "0 <= m" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   541
  have d: "!! k. k \<noteq> 0 \<Longrightarrow> ? l. k = Suc l" by arith
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   542
  show ?concl
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   543
    apply (subst foldseq_tail[OF nm])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   544
    apply (rule foldseq_significant_positions)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   545
    apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   546
    apply (case_tac "m=n")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   547
    apply (drule a, simp+)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   548
    apply (drule b[OF nm])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   549
    apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   550
    apply (drule a)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   551
    apply simp+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   552
    apply (case_tac "k=0")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   553
    apply (simp add: prems)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   554
    apply (drule d)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   555
    apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   556
    by (simp add: prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   557
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   558
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   559
lemma foldseq_zerostart:
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   560
  "! x. f 0 (f 0 x) = f 0 x \<Longrightarrow>  ! i. i <= n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   561
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   562
  assume f00x: "! x. f 0 (f 0 x) = f 0 x"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   563
  have "! s. (! i. i<=n \<longrightarrow> s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   564
    apply (induct n)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   565
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   566
    apply (rule allI, rule impI)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   567
    proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   568
      fix n
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   569
      fix s
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   570
      have a:"foldseq f s (Suc (Suc n)) = f (s 0) (foldseq f (% k. s(Suc k)) (Suc n))" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   571
      assume b: "! s. ((\<forall>i\<le>n. s i = 0) \<longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n)))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   572
      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
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   573
      assume d: "! i. i <= Suc n \<longrightarrow> s i = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   574
      show "foldseq f s (Suc (Suc n)) = f 0 (s (Suc (Suc n)))"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   575
        apply (subst a)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   576
        apply (subst c)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   577
        by (simp add: d f00x)+
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   578
    qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   579
  then show "! i. i <= n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s (Suc n) = f 0 (s (Suc n))" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   580
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   581
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   582
lemma foldseq_zerostart2:
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   583
  "! x. f 0 x = x \<Longrightarrow>  ! i. i < n \<longrightarrow> s i = 0 \<Longrightarrow> foldseq f s n = s n"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   584
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   585
  assume a:"! i. i<n \<longrightarrow> s i = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   586
  assume x:"! x. f 0 x = x"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   587
  from x have f00x: "! x. f 0 (f 0 x) = f 0 x" by blast
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   588
  have b: "!! i l. i < Suc l = (i <= l)" by arith
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   589
  have d: "!! k. k \<noteq> 0 \<Longrightarrow> ? l. k = Suc l" by arith
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   590
  show "foldseq f s n = s n"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   591
  apply (case_tac "n=0")
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   592
  apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   593
  apply (insert a)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   594
  apply (drule d)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   595
  apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   596
  apply (simp add: b)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   597
  apply (insert f00x)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   598
  apply (drule foldseq_zerostart)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   599
  by (simp add: x)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   600
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   601
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   602
lemma foldseq_almostzero:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   603
  assumes f0x:"! x. f 0 x = x" and fx0: "! x. f x 0 = x" and s0:"! i. i \<noteq> j \<longrightarrow> s i = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   604
  shows "foldseq f s n = (if (j <= n) then (s j) else 0)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   605
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   606
  from s0 have a: "! i. i < j \<longrightarrow> s i = 0" by simp
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   607
  from s0 have b: "! i. j < i \<longrightarrow> s i = 0" by simp
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   608
  show ?concl
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   609
    apply auto
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   610
    apply (subst foldseq_zerotail2[of f, OF fx0, of j, OF b, of n, THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   611
    apply simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   612
    apply (subst foldseq_zerostart2)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   613
    apply (simp add: f0x a)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   614
    apply (subst foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   615
    by (simp add: s0 f0x)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   616
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   617
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   618
lemma foldseq_distr_unary:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   619
  assumes "!! a b. g (f a b) = f (g a) (g b)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   620
  shows "g(foldseq f s n) = foldseq f (% x. g(s x)) n" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   621
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   622
  have "! s. g(foldseq f s n) = foldseq f (% x. g(s x)) n"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   623
    apply (induct_tac n)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   624
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   625
    apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   626
    apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   627
    apply (drule_tac x="% k. s (Suc k)" in spec)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   628
    by (simp add: prems)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   629
  then show ?concl by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   630
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   631
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   632
constdefs
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   633
  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"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   634
  "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)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   635
  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"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   636
  "mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   637
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   638
lemma mult_matrix_n:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   639
  assumes prems: "ncols A \<le>  n" (is ?An) "nrows B \<le> n" (is ?Bn) "fadd 0 0 = 0" "fmul 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   640
  shows c:"mult_matrix fmul fadd A B = mult_matrix_n n fmul fadd A B" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   641
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   642
  show ?concl using prems
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   643
    apply (simp add: mult_matrix_def mult_matrix_n_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   644
    apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   645
    by (rule foldseq_zerotail, simp_all add: nrows_le ncols_le prems)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   646
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   647
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   648
lemma mult_matrix_nm:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   649
  assumes prems: "ncols A <= n" "nrows B <= n" "ncols A <= m" "nrows B <= m" "fadd 0 0 = 0" "fmul 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   650
  shows "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   651
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   652
  from prems have "mult_matrix_n n fmul fadd A B = mult_matrix fmul fadd A B" by (simp add: mult_matrix_n)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   653
  also from prems have "\<dots> = mult_matrix_n m fmul fadd A B" by (simp add: mult_matrix_n[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   654
  finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   655
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   656
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   657
constdefs
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   658
  r_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   659
  "r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   660
  l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   661
  "l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   662
  distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   663
  "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   664
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   665
lemma max1: "!! a x y. (a::nat) <= x \<Longrightarrow> a <= max x y" by (arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   666
lemma max2: "!! b x y. (b::nat) <= y \<Longrightarrow> b <= max x y" by (arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   667
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   668
lemma r_distributive_matrix:
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   669
 assumes prems:
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   670
  "r_distributive fmul fadd"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   671
  "associative fadd"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   672
  "commutative fadd"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   673
  "fadd 0 0 = 0"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   674
  "! a. fmul a 0 = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   675
  "! a. fmul 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   676
 shows "r_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   677
proof -
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   678
  from prems show ?concl
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   679
    apply (simp add: r_distributive_def mult_matrix_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   680
    proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   681
      fix a::"'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   682
      fix u::"'b matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   683
      fix v::"'b matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   684
      let ?mx = "max (ncols a) (max (nrows u) (nrows v))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   685
      from prems show "mult_matrix_n (max (ncols a) (nrows (combine_matrix fadd u v))) fmul fadd a (combine_matrix fadd u v) =
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   686
        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)"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   687
        apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   688
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   689
        apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   690
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   691
        apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   692
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   693
        apply (simp add: mult_matrix_n_def r_distributive_def foldseq_distr[of fadd])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   694
        apply (simp add: combine_matrix_def combine_infmatrix_def)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   695
        apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   696
        apply (subst RepAbs_matrix)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   697
        apply (simp, auto)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   698
        apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   699
        apply (rule exI[of _ "ncols v"], simp add: ncols_le foldseq_zero)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   700
        apply (subst RepAbs_matrix)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   701
        apply (simp, auto)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   702
        apply (rule exI[of _ "nrows a"], simp add: nrows_le foldseq_zero)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   703
        apply (rule exI[of _ "ncols u"], simp add: ncols_le foldseq_zero)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   704
        done
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   705
    qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   706
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   707
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   708
lemma l_distributive_matrix:
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   709
 assumes prems:
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   710
  "l_distributive fmul fadd"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   711
  "associative fadd"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   712
  "commutative fadd"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   713
  "fadd 0 0 = 0"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   714
  "! a. fmul a 0 = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   715
  "! a. fmul 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   716
 shows "l_distributive (mult_matrix fmul fadd) (combine_matrix fadd)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   717
proof -
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   718
  from prems show ?concl
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   719
    apply (simp add: l_distributive_def mult_matrix_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   720
    proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   721
      fix a::"'b matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   722
      fix u::"'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   723
      fix v::"'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   724
      let ?mx = "max (nrows a) (max (ncols u) (ncols v))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   725
      from prems show "mult_matrix_n (max (ncols (combine_matrix fadd u v)) (nrows a)) fmul fadd (combine_matrix fadd u v) a =
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   726
               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)"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   727
        apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   728
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   729
        apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   730
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   731
        apply (subst mult_matrix_nm[of _ _ _ ?mx fadd fmul])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   732
        apply (simp add: max1 max2 combine_nrows combine_ncols)+
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   733
        apply (simp add: mult_matrix_n_def l_distributive_def foldseq_distr[of fadd])
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   734
        apply (simp add: combine_matrix_def combine_infmatrix_def)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   735
        apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   736
        apply (subst RepAbs_matrix)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   737
        apply (simp, auto)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   738
        apply (rule exI[of _ "nrows v"], simp add: nrows_le foldseq_zero)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   739
        apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   740
        apply (subst RepAbs_matrix)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   741
        apply (simp, auto)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   742
        apply (rule exI[of _ "nrows u"], simp add: nrows_le foldseq_zero)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   743
        apply (rule exI[of _ "ncols a"], simp add: ncols_le foldseq_zero)
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   744
        done
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   745
    qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   746
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   747
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   748
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14662
diff changeset
   749
instance matrix :: (zero) zero ..
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   750
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   751
defs(overloaded)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   752
  zero_matrix_def: "(0::('a::zero) matrix) == Abs_matrix(% j i. 0)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   753
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   754
lemma Rep_zero_matrix_def[simp]: "Rep_matrix 0 j i = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   755
  apply (simp add: zero_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   756
  apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   757
  by (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   758
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   759
lemma zero_matrix_def_nrows[simp]: "nrows 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   760
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   761
  have a:"!! (x::nat). x <= 0 \<Longrightarrow> x = 0" by (arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   762
  show "nrows 0 = 0" by (rule a, subst nrows_le, simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   763
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   764
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   765
lemma zero_matrix_def_ncols[simp]: "ncols 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   766
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   767
  have a:"!! (x::nat). x <= 0 \<Longrightarrow> x = 0" by (arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   768
  show "ncols 0 = 0" by (rule a, subst ncols_le, simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   769
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   770
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   771
lemma combine_matrix_zero_r_neutral: "zero_r_neutral f \<Longrightarrow> zero_r_neutral (combine_matrix f)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   772
  by (simp add: zero_r_neutral_def combine_matrix_def combine_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   773
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   774
lemma mult_matrix_zero_closed: "\<lbrakk>fadd 0 0 = 0; zero_closed fmul\<rbrakk> \<Longrightarrow> zero_closed (mult_matrix fmul fadd)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   775
  apply (simp add: zero_closed_def mult_matrix_def mult_matrix_n_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   776
  apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   777
  by (subst foldseq_zero, (simp add: zero_matrix_def)+)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   778
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   779
lemma mult_matrix_n_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; !a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd A 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   780
  apply (simp add: mult_matrix_n_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   781
  apply (subst foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   782
  by (simp_all add: zero_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   783
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   784
lemma mult_matrix_n_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; !a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix_n n fmul fadd 0 A = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   785
  apply (simp add: mult_matrix_n_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   786
  apply (subst foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   787
  by (simp_all add: zero_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   788
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   789
lemma mult_matrix_zero_left[simp]: "\<lbrakk>fadd 0 0 = 0; !a. fmul 0 a = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd 0 A = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   790
by (simp add: mult_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   791
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   792
lemma mult_matrix_zero_right[simp]: "\<lbrakk>fadd 0 0 = 0; !a. fmul a 0 = 0\<rbrakk> \<Longrightarrow> mult_matrix fmul fadd A 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   793
by (simp add: mult_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   794
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   795
lemma apply_matrix_zero[simp]: "f 0 = 0 \<Longrightarrow> apply_matrix f 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   796
  apply (simp add: apply_matrix_def apply_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   797
  by (simp add: zero_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   798
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   799
lemma combine_matrix_zero: "f 0 0 = 0 \<Longrightarrow> combine_matrix f 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   800
  apply (simp add: combine_matrix_def combine_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   801
  by (simp add: zero_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   802
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   803
lemma apply_zero_matrix_def[simp]: "apply_matrix (% x. 0) A = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   804
  apply (simp add: apply_matrix_def apply_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   805
  by (simp add: zero_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   806
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   807
constdefs
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   808
  singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   809
  "singleton_matrix j i a == Abs_matrix(% m n. if j = m & i = n then a else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   810
  move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   811
  "move_matrix A y x == Abs_matrix(% j i. if (neg ((int j)-y)) | (neg ((int i)-x)) then 0 else Rep_matrix A (nat ((int j)-y)) (nat ((int i)-x)))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   812
  take_rows :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   813
  "take_rows A r == Abs_matrix(% j i. if (j < r) then (Rep_matrix A j i) else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   814
  take_columns :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   815
  "take_columns A c == Abs_matrix(% j i. if (i < c) then (Rep_matrix A j i) else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   816
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   817
constdefs
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   818
  column_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   819
  "column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   820
  row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   821
  "row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   822
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   823
lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   824
apply (simp add: singleton_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   825
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   826
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   827
apply (rule exI[of _ "Suc m"], simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   828
apply (rule exI[of _ "Suc n"], simp+)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   829
by (subst RepAbs_matrix, rule exI[of _ "Suc j"], simp, rule exI[of _ "Suc i"], simp+)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   830
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   831
lemma singleton_matrix_zero[simp]: "singleton_matrix j i 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   832
  by (simp add: singleton_matrix_def zero_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   833
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   834
lemma nrows_singleton[simp]: "nrows(singleton_matrix j i e) = (if e = 0 then 0 else Suc j)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   835
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   836
apply (rule le_anti_sym)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   837
apply (subst nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   838
apply (simp add: singleton_matrix_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   839
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   840
apply (simp, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   841
apply (simp, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   842
apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   843
apply (simp add: Suc_le_eq)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   844
apply (rule not_leE)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   845
apply (subst nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   846
by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   847
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   848
lemma ncols_singleton[simp]: "ncols(singleton_matrix j i e) = (if e = 0 then 0 else Suc i)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   849
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   850
apply (rule le_anti_sym)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   851
apply (subst ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   852
apply (simp add: singleton_matrix_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   853
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   854
apply (simp, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   855
apply (simp, arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   856
apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   857
apply (simp add: Suc_le_eq)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   858
apply (rule not_leE)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   859
apply (subst ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   860
by simp
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   861
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   862
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)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   863
apply (simp add: singleton_matrix_def combine_matrix_def combine_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   864
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   865
apply (rule exI[of _ "Suc j"], simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   866
apply (rule exI[of _ "Suc i"], simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   867
apply (rule comb[of "Abs_matrix" "Abs_matrix"], simp, (rule ext)+)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   868
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   869
apply (rule exI[of _ "Suc j"], simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   870
apply (rule exI[of _ "Suc i"], simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   871
by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   872
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   873
lemma Rep_move_matrix[simp]:
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   874
  "Rep_matrix (move_matrix A y x) j i =
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   875
  (if (neg ((int j)-y)) | (neg ((int i)-x)) then 0 else Rep_matrix A (nat((int j)-y)) (nat((int i)-x)))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   876
apply (simp add: move_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   877
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   878
by (subst RepAbs_matrix,
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   879
  rule exI[of _ "(nrows A)+(nat (abs y))"], auto, rule nrows, arith,
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   880
  rule exI[of _ "(ncols A)+(nat (abs x))"], auto, rule ncols, arith)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   881
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   882
lemma Rep_take_columns[simp]:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   883
  "Rep_matrix (take_columns A c) j i =
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   884
  (if i < c then (Rep_matrix A j i) else 0)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   885
apply (simp add: take_columns_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   886
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   887
apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   888
apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   889
done
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   890
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   891
lemma Rep_take_rows[simp]:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   892
  "Rep_matrix (take_rows A r) j i =
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   893
  (if j < r then (Rep_matrix A j i) else 0)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   894
apply (simp add: take_rows_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   895
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   896
apply (rule exI[of _ "nrows A"], auto, simp add: nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   897
apply (rule exI[of _ "ncols A"], auto, simp add: ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   898
done
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   899
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   900
lemma Rep_column_of_matrix[simp]:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   901
  "Rep_matrix (column_of_matrix A c) j i = (if i = 0 then (Rep_matrix A j c) else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   902
  by (simp add: column_of_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   903
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   904
lemma Rep_row_of_matrix[simp]:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   905
  "Rep_matrix (row_of_matrix A r) j i = (if j = 0 then (Rep_matrix A r i) else 0)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   906
  by (simp add: row_of_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   907
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   908
lemma mult_matrix_singleton_right[simp]:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   909
  assumes prems:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   910
  "! x. fmul x 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   911
  "! x. fmul 0 x = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   912
  "! x. fadd 0 x = x"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   913
  "! x. fadd x 0 = x"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   914
  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))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   915
  apply (simp add: mult_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   916
  apply (subst mult_matrix_nm[of _ _ _ "max (ncols A) (Suc j)"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   917
  apply (simp add: max_def)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   918
  apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   919
  apply (simp add: prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   920
  apply (simp add: mult_matrix_n_def apply_matrix_def apply_infmatrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   921
  apply (rule comb[of "Abs_matrix" "Abs_matrix"], auto, (rule ext)+)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   922
  apply (subst foldseq_almostzero[of _ j])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   923
  apply (simp add: prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   924
  apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   925
  apply (insert ncols_le[of A j])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   926
  apply (arith)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   927
  proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   928
    fix k
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   929
    fix l
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   930
    assume a:"~neg(int l - int i)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   931
    assume b:"nat (int l - int i) = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   932
    from a b have a: "l = i" by(insert not_neg_nat[of "int l - int i"], simp add: a b)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   933
    assume c: "i \<noteq> l"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   934
    from c a show "fmul (Rep_matrix A k j) e = 0" by blast
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   935
  qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   936
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   937
lemma mult_matrix_ext:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   938
  assumes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   939
  eprem:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   940
  "? e. (! a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   941
  and fprems:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   942
  "! a. fmul 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   943
  "! a. fmul a 0 = 0"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   944
  "! a. fadd a 0 = a"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   945
  "! a. fadd 0 a = a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   946
  and contraprems:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   947
  "mult_matrix fmul fadd A = mult_matrix fmul fadd B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   948
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   949
  "A = B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   950
proof(rule contrapos_np[of "False"], simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   951
  assume a: "A \<noteq> B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   952
  have b: "!! f g. (! x y. f x y = g x y) \<Longrightarrow> f = g" by ((rule ext)+, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   953
  have "? j i. (Rep_matrix A j i) \<noteq> (Rep_matrix B j i)"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   954
    apply (rule contrapos_np[of "False"], simp+)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   955
    apply (insert b[of "Rep_matrix A" "Rep_matrix B"], simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   956
    by (simp add: Rep_matrix_inject a)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   957
  then obtain J I where c:"(Rep_matrix A J I) \<noteq> (Rep_matrix B J I)" by blast
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   958
  from eprem obtain e where eprops:"(! a b. a \<noteq> b \<longrightarrow> fmul a e \<noteq> fmul b e)" by blast
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   959
  let ?S = "singleton_matrix I 0 e"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   960
  let ?comp = "mult_matrix fmul fadd"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   961
  have d: "!!x f g. f = g \<Longrightarrow> f x = g x" by blast
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   962
  have e: "(% x. fmul x e) 0 = 0" by (simp add: prems)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   963
  have "~(?comp A ?S = ?comp B ?S)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   964
    apply (rule notI)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   965
    apply (simp add: fprems eprops)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   966
    apply (simp add: Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   967
    apply (drule d[of _ _ "J"], drule d[of _ _ "0"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   968
    by (simp add: e c eprops)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   969
  with contraprems show "False" by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   970
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   971
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   972
constdefs
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   973
  foldmatrix :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   974
  "foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   975
  foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   976
  "foldmatrix_transposed f g A m n == foldseq g (% j. foldseq_transposed f (A j) n) m"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   977
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   978
lemma foldmatrix_transpose:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   979
  assumes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   980
  "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   981
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   982
  "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   983
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   984
  have forall:"!! P x. (! x. P x) \<Longrightarrow> P x" by auto
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   985
  have tworows:"! A. foldmatrix f g A 1 n = foldmatrix_transposed g f (transpose_infmatrix A) n 1"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   986
    apply (induct n)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   987
    apply (simp add: foldmatrix_def foldmatrix_transposed_def prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   988
    apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   989
    by (drule_tac x1="(% j i. A j (Suc i))" in forall, simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   990
  show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   991
    apply (simp add: foldmatrix_def foldmatrix_transposed_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   992
    apply (induct m, simp)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
   993
    apply (simp)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   994
    apply (insert tworows)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   995
    apply (drule_tac x="% j i. (if j = 0 then (foldseq_transposed g (\<lambda>u. A u i) na) else (A (Suc na) i))" in spec)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   996
    by (simp add: foldmatrix_def foldmatrix_transposed_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   997
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   998
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
   999
lemma foldseq_foldseq:
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1000
assumes
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1001
  "associative f"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1002
  "associative g"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1003
  "! a b c d. g(f a b) (f c d) = f (g a c) (g b d)"
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1004
shows
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1005
  "foldseq g (% j. foldseq f (A j) n) m = foldseq f (% j. foldseq g ((transpose_infmatrix A) j) m) n"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1006
  apply (insert foldmatrix_transpose[of g f A m n])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1007
  by (simp add: foldmatrix_def foldmatrix_transposed_def foldseq_assoc[THEN sym] prems)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1008
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1009
lemma mult_n_nrows:
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1010
assumes
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1011
"! a. fmul 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1012
"! a. fmul a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1013
"fadd 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1014
shows "nrows (mult_matrix_n n fmul fadd A B) \<le> nrows A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1015
apply (subst nrows_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1016
apply (simp add: mult_matrix_n_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1017
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1018
apply (rule_tac x="nrows A" in exI)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1019
apply (simp add: nrows prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1020
apply (rule_tac x="ncols B" in exI)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1021
apply (simp add: ncols prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1022
by (simp add: nrows prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1023
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1024
lemma mult_n_ncols:
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1025
assumes
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1026
"! a. fmul 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1027
"! a. fmul a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1028
"fadd 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1029
shows "ncols (mult_matrix_n n fmul fadd A B) \<le> ncols B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1030
apply (subst ncols_le)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1031
apply (simp add: mult_matrix_n_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1032
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1033
apply (rule_tac x="nrows A" in exI)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1034
apply (simp add: nrows prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1035
apply (rule_tac x="ncols B" in exI)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1036
apply (simp add: ncols prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1037
by (simp add: ncols prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1038
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1039
lemma mult_nrows:
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1040
assumes
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1041
"! a. fmul 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1042
"! a. fmul a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1043
"fadd 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1044
shows "nrows (mult_matrix fmul fadd A B) \<le> nrows A"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1045
by (simp add: mult_matrix_def mult_n_nrows prems)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1046
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1047
lemma mult_ncols:
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1048
assumes
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1049
"! a. fmul 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1050
"! a. fmul a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1051
"fadd 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1052
shows "ncols (mult_matrix fmul fadd A B) \<le> ncols B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1053
by (simp add: mult_matrix_def mult_n_ncols prems)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1054
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1055
lemma mult_matrix_assoc:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1056
  assumes prems:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1057
  "! a. fmul1 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1058
  "! a. fmul1 a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1059
  "! a. fmul2 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1060
  "! a. fmul2 a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1061
  "fadd1 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1062
  "fadd2 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1063
  "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1064
  "associative fadd1"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1065
  "associative fadd2"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1066
  "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1067
  "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1068
  "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1069
  shows "mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B) C = mult_matrix fmul1 fadd1 A (mult_matrix fmul2 fadd2 B C)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1070
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1071
  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
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1072
  have fmul2fadd1fold: "!! x s n. fmul2 (foldseq fadd1 s n)  x = foldseq fadd1 (% k. fmul2 (s k) x) n"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1073
    by (rule_tac g1 = "% y. fmul2 y x" in ssubst [OF foldseq_distr_unary], simp_all!)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1074
  have fmul1fadd2fold: "!! x s n. fmul1 x (foldseq fadd2 s n) = foldseq fadd2 (% k. fmul1 x (s k)) n"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1075
      by (rule_tac g1 = "% y. fmul1 x y" in ssubst [OF foldseq_distr_unary], simp_all!)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1076
  let ?N = "max (ncols A) (max (ncols B) (max (nrows B) (nrows C)))"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1077
  show ?concl
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1078
    apply (simp add: Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1079
    apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1080
    apply (simp add: mult_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1081
    apply (subst 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)"])
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1082
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1083
    apply (subst 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)"])     apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1084
    apply (subst mult_matrix_nm[of _ _ _ "?N"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1085
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1086
    apply (subst mult_matrix_nm[of _ _ _ "?N"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1087
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1088
    apply (subst mult_matrix_nm[of _ _ _ "?N"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1089
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1090
    apply (subst mult_matrix_nm[of _ _ _ "?N"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1091
    apply (simp add: max1 max2 mult_n_ncols mult_n_nrows prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1092
    apply (simp add: mult_matrix_n_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1093
    apply (rule comb_left)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1094
    apply ((rule ext)+, simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1095
    apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1096
    apply (rule exI[of _ "nrows B"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1097
    apply (simp add: nrows prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1098
    apply (rule exI[of _ "ncols C"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1099
    apply (simp add: prems ncols foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1100
    apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1101
    apply (rule exI[of _ "nrows A"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1102
    apply (simp add: nrows prems foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1103
    apply (rule exI[of _ "ncols B"])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1104
    apply (simp add: prems ncols foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1105
    apply (simp add: fmul2fadd1fold fmul1fadd2fold prems)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1106
    apply (subst foldseq_foldseq)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1107
    apply (simp add: prems)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1108
    by (simp add: transpose_infmatrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1109
qed
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1110
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1111
lemma
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1112
  assumes prems:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1113
  "! a. fmul1 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1114
  "! a. fmul1 a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1115
  "! a. fmul2 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1116
  "! a. fmul2 a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1117
  "fadd1 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1118
  "fadd2 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1119
  "! a b c d. fadd2 (fadd1 a b) (fadd1 c d) = fadd1 (fadd2 a c) (fadd2 b d)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1120
  "associative fadd1"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1121
  "associative fadd2"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1122
  "! a b c. fmul2 (fmul1 a b) c = fmul1 a (fmul2 b c)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1123
  "! a b c. fmul2 (fadd1 a b) c = fadd1 (fmul2 a c) (fmul2 b c)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1124
  "! a b c. fmul1 c (fadd2 a b) = fadd2 (fmul1 c a) (fmul1 c b)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1125
  shows
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1126
  "(mult_matrix fmul1 fadd1 A) o (mult_matrix fmul2 fadd2 B) = mult_matrix fmul2 fadd2 (mult_matrix fmul1 fadd1 A B)"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1127
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1128
apply (simp add: comp_def )
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1129
by (simp add: mult_matrix_assoc prems)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1130
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1131
lemma mult_matrix_assoc_simple:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1132
  assumes prems:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1133
  "! a. fmul 0 a = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1134
  "! a. fmul a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1135
  "fadd 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1136
  "associative fadd"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1137
  "commutative fadd"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1138
  "associative fmul"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1139
  "distributive fmul fadd"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1140
  shows "mult_matrix fmul fadd (mult_matrix fmul fadd A B) C = mult_matrix fmul fadd A (mult_matrix fmul fadd B C)" (is ?concl)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1141
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1142
  have "!! a b c d. fadd (fadd a b) (fadd c d) = fadd (fadd a c) (fadd b d)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1143
    by (simp! add: associative_def commutative_def)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1144
  then show ?concl
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1145
    apply (subst mult_matrix_assoc)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1146
    apply (simp_all!)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1147
    by (simp add: associative_def distributive_def l_distributive_def r_distributive_def)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1148
qed
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1149
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1150
lemma transpose_apply_matrix: "f 0 = 0 \<Longrightarrow> transpose_matrix (apply_matrix f A) = apply_matrix f (transpose_matrix A)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1151
apply (simp add: Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1152
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1153
by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1154
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1155
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)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1156
apply (simp add: Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1157
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1158
by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1159
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1160
lemma Rep_mult_matrix:
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1161
  assumes
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1162
  "! a. fmul 0 a = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1163
  "! a. fmul a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1164
  "fadd 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1165
  shows
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1166
  "Rep_matrix(mult_matrix fmul fadd A B) j i =
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1167
  foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) (max (ncols A) (nrows B))"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1168
apply (simp add: mult_matrix_def mult_matrix_n_def)
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1169
apply (subst RepAbs_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1170
apply (rule exI[of _ "nrows A"], simp! add: nrows foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1171
apply (rule exI[of _ "ncols B"], simp! add: ncols foldseq_zero)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1172
by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1173
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1174
lemma transpose_mult_matrix:
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1175
  assumes
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1176
  "! a. fmul 0 a = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1177
  "! a. fmul a 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1178
  "fadd 0 0 = 0"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1179
  "! x y. fmul y x = fmul x y"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1180
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1181
  "transpose_matrix (mult_matrix fmul fadd A B) = mult_matrix fmul fadd (transpose_matrix B) (transpose_matrix A)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1182
  apply (simp add: Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1183
  apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1184
  by (simp! add: Rep_mult_matrix max_ac)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1185
14691
e1eedc8cad37 tuned instance statements;
wenzelm
parents: 14662
diff changeset
  1186
instance matrix :: ("{ord, zero}") ord ..
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1187
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1188
defs (overloaded)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1189
  le_matrix_def: "(A::('a::{ord,zero}) matrix) <= B == ! j i. (Rep_matrix A j i) <= (Rep_matrix B j i)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1190
  less_def: "(A::('a::{ord,zero}) matrix) < B == (A <= B) & (A \<noteq> B)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1191
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1192
instance matrix :: ("{order, zero}") order
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1193
apply intro_classes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1194
apply (simp_all add: le_matrix_def less_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1195
apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1196
apply (drule_tac x=j in spec, drule_tac x=j in spec)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1197
apply (drule_tac x=i in spec, drule_tac x=i in spec)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1198
apply (simp)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1199
apply (simp add: Rep_matrix_inject[THEN sym])
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1200
apply (rule ext)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1201
apply (drule_tac x=xa in spec, drule_tac x=xa in spec)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1202
apply (drule_tac x=xb in spec, drule_tac x=xb in spec)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1203
by simp
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1204
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1205
lemma le_apply_matrix:
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1206
  assumes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1207
  "f 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1208
  "! x y. x <= y \<longrightarrow> f x <= f y"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1209
  "(a::('a::{ord, zero}) matrix) <= b"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1210
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1211
  "apply_matrix f a <= apply_matrix f b"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1212
  by (simp! add: le_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1213
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1214
lemma le_combine_matrix:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1215
  assumes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1216
  "f 0 0 = 0"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1217
  "! a b c d. a <= b & c <= d \<longrightarrow> f a c <= f b d"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1218
  "A <= B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1219
  "C <= D"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1220
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1221
  "combine_matrix f A C <= combine_matrix f B D"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1222
by (simp! add: le_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1223
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1224
lemma le_left_combine_matrix:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1225
  assumes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1226
  "f 0 0 = 0"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1227
  "! a b c. 0 <= c & a <= b \<longrightarrow> f c a <= f c b"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1228
  "0 <= C"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1229
  "A <= B"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1230
  shows
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1231
  "combine_matrix f C A <= combine_matrix f C B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1232
  by (simp! add: le_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1233
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1234
lemma le_right_combine_matrix:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1235
  assumes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1236
  "f 0 0 = 0"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1237
  "! a b c. 0 <= c & a <= b \<longrightarrow> f a c <= f b c"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1238
  "0 <= C"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1239
  "A <= B"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1240
  shows
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1241
  "combine_matrix f A C <= combine_matrix f B C"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1242
  by (simp! add: le_matrix_def)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1243
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1244
lemma le_transpose_matrix: "(A <= B) = (transpose_matrix A <= transpose_matrix B)"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1245
  by (simp add: le_matrix_def, auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1246
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1247
lemma le_foldseq:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1248
  assumes
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1249
  "! a b c d . a <= b & c <= d \<longrightarrow> f a c <= f b d"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1250
  "! i. i <= n \<longrightarrow> s i <= t i"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1251
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1252
  "foldseq f s n <= foldseq f t n"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1253
proof -
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1254
  have "! s t. (! i. i<=n \<longrightarrow> s i <= t i) \<longrightarrow> foldseq f s n <= foldseq f t n" by (induct_tac n, simp_all!)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1255
  then show "foldseq f s n <= foldseq f t n" by (simp!)
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1256
qed
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1257
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1258
lemma le_left_mult:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1259
  assumes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1260
  "! a b c d. a <= b & c <= d \<longrightarrow> fadd a c <= fadd b d"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1261
  "! c a b. 0 <= c & a <= b \<longrightarrow> fmul c a <= fmul c b"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1262
  "! a. fmul 0 a = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1263
  "! a. fmul a 0 = 0"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1264
  "fadd 0 0 = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1265
  "0 <= C"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1266
  "A <= B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1267
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1268
  "mult_matrix fmul fadd C A <= mult_matrix fmul fadd C B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1269
  apply (simp! add: le_matrix_def Rep_mult_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1270
  apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1271
  apply (subst foldseq_zerotail[of _ _ _ "max (ncols C) (max (nrows A) (nrows B))"], simp_all add: nrows ncols max1 max2)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1272
  apply (rule le_foldseq)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1273
  by (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1274
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1275
lemma le_right_mult:
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1276
  assumes
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1277
  "! a b c d. a <= b & c <= d \<longrightarrow> fadd a c <= fadd b d"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1278
  "! c a b. 0 <= c & a <= b \<longrightarrow> fmul a c <= fmul b c"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1279
  "! a. fmul 0 a = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1280
  "! a. fmul a 0 = 0"
14662
d2c6a0f030ab proper document setup;
wenzelm
parents: 14654
diff changeset
  1281
  "fadd 0 0 = 0"
14593
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1282
  "0 <= C"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1283
  "A <= B"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1284
  shows
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1285
  "mult_matrix fmul fadd A C <= mult_matrix fmul fadd B C"
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1286
  apply (simp! add: le_matrix_def Rep_mult_matrix)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1287
  apply (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1288
  apply (subst foldseq_zerotail[of _ _ _ "max (nrows C) (max (ncols A) (ncols B))"], simp_all add: nrows ncols max1 max2)+
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1289
  apply (rule le_foldseq)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1290
  by (auto)
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1291
90c88e7ef62d first version of matrices for HOL/Isabelle
obua
parents:
diff changeset
  1292
end