src/HOL/Matrix/MatrixGeneral.thy
changeset 15795 997884600e0a
parent 15580 900291ee0af8
child 16417 9bc16273c2d4
--- a/src/HOL/Matrix/MatrixGeneral.thy	Thu Apr 21 18:57:18 2005 +0200
+++ b/src/HOL/Matrix/MatrixGeneral.thy	Thu Apr 21 18:58:44 2005 +0200
@@ -1058,7 +1058,7 @@
     apply (induct n)
     apply (simp add: foldmatrix_def foldmatrix_transposed_def prems)+
     apply (auto)
-    by (drule_tac x1="(% j i. A j (Suc i))" in forall, simp)
+    by (drule_tac x3="(% j i. A j (Suc i))" in forall, simp)
   show "foldmatrix f g A m n = foldmatrix_transposed g f (transpose_infmatrix A) n m"
     apply (simp add: foldmatrix_def foldmatrix_transposed_def)
     apply (induct m, simp)