--- a/src/HOL/Matrix/ComputeNumeral.thy Sat Mar 17 11:59:59 2012 +0100
+++ b/src/HOL/Matrix/ComputeNumeral.thy Sat Mar 17 12:00:11 2012 +0100
@@ -50,7 +50,7 @@
(* multiplication for bit strings *)
lemma mult_Pls_right: "x * Int.Pls = Int.Pls" by (simp add: Pls_def)
-lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute, simp add: mult_Min)
+lemma mult_Min_right: "x * Int.Min = - x" by (subst mult_commute) simp
lemma multb0x: "(Int.Bit0 x) * y = Int.Bit0 (x * y)" by (rule mult_Bit0)
lemma multxb0: "x * (Int.Bit0 y) = Int.Bit0 (x * y)" unfolding Bit0_def by simp
lemma multb1: "(Int.Bit1 x) * (Int.Bit1 y) = Int.Bit1 (Int.Bit0 (x * y) + x + y)"
--- a/src/HOL/Matrix/Matrix.thy Sat Mar 17 11:59:59 2012 +0100
+++ b/src/HOL/Matrix/Matrix.thy Sat Mar 17 12:00:11 2012 +0100
@@ -64,7 +64,7 @@
lemma transpose_infmatrix: "transpose_infmatrix (% j i. P j i) = (% j i. P i j)"
apply (rule ext)+
- by (simp add: transpose_infmatrix_def)
+ by simp
lemma transpose_infmatrix_closed[simp]: "Rep_matrix (Abs_matrix (transpose_infmatrix (Rep_matrix x))) = transpose_infmatrix (Rep_matrix x)"
apply (rule Abs_matrix_inverse)
@@ -836,7 +836,7 @@
by (simp add: zero_matrix_def)
lemma transpose_matrix_zero[simp]: "transpose_matrix 0 = 0"
-apply (simp add: transpose_matrix_def transpose_infmatrix_def zero_matrix_def RepAbs_matrix)
+apply (simp add: transpose_matrix_def zero_matrix_def RepAbs_matrix)
apply (subst Rep_matrix_inject[symmetric], (rule ext)+)
apply (simp add: RepAbs_matrix)
done
@@ -1464,7 +1464,7 @@
definition "sup = combine_matrix sup"
instance
- by default (auto simp add: inf_le1 inf_le2 le_infI le_matrix_def inf_matrix_def sup_matrix_def)
+ by default (auto simp add: le_infI le_matrix_def inf_matrix_def sup_matrix_def)
end