--- a/Arith.ML Fri Jun 24 15:11:39 1994 +0200
+++ b/Arith.ML Wed Jun 29 12:04:04 1994 +0200
@@ -53,14 +53,14 @@
val arith_ss = arith_ss addsimps [add_0_right,add_Suc_right];
(*Associative law for addition*)
-val add_assoc = prove_goal Arith.thy "(m + n) + k = m + (n + k)::nat"
+val add_assoc = prove_goal Arith.thy "(m + n) + k = m + ((n + k)::nat)"
(fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]);
(*Commutative law for addition*)
-val add_commute = prove_goal Arith.thy "m + n = n + m::nat"
+val add_commute = prove_goal Arith.thy "m + n = n + (m::nat)"
(fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]);
-val add_left_commute = prove_goal Arith.thy "x+(y+z)=y+(x+z)::nat"
+val add_left_commute = prove_goal Arith.thy "x+(y+z)=y+((x+z)::nat)"
(fn _ => [rtac (add_commute RS trans) 1, rtac (add_assoc RS trans) 1,
rtac (add_commute RS arg_cong) 1]);
@@ -82,25 +82,25 @@
val arith_ss = arith_ss addsimps [mult_0_right,mult_Suc_right];
(*Commutative law for multiplication*)
-val mult_commute = prove_goal Arith.thy "m * n = n * m::nat"
+val mult_commute = prove_goal Arith.thy "m * n = n * (m::nat)"
(fn _ => [nat_ind_tac "m" 1, ALLGOALS (asm_simp_tac arith_ss)]);
(*addition distributes over multiplication*)
-val add_mult_distrib = prove_goal Arith.thy "(m + n)*k = (m*k) + (n*k)::nat"
+val add_mult_distrib = prove_goal Arith.thy "(m + n)*k = (m*k) + ((n*k)::nat)"
(fn _ => [nat_ind_tac "m" 1,
ALLGOALS(asm_simp_tac (arith_ss addsimps add_ac))]);
-val add_mult_distrib2 = prove_goal Arith.thy "k*(m + n) = (k*m) + (k*n)::nat"
+val add_mult_distrib2 = prove_goal Arith.thy "k*(m + n) = (k*m) + ((k*n)::nat)"
(fn _ => [nat_ind_tac "m" 1,
ALLGOALS(asm_simp_tac (arith_ss addsimps add_ac))]);
val arith_ss = arith_ss addsimps [add_mult_distrib,add_mult_distrib2];
(*Associative law for multiplication*)
-val mult_assoc = prove_goal Arith.thy "(m * n) * k = m * (n * k)::nat"
+val mult_assoc = prove_goal Arith.thy "(m * n) * k = m * ((n * k)::nat)"
(fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]);
-val mult_left_commute = prove_goal Arith.thy "x*(y*z) = y*(x*z)::nat"
+val mult_left_commute = prove_goal Arith.thy "x*(y*z) = y*((x*z)::nat)"
(fn _ => [rtac trans 1, rtac mult_commute 1, rtac trans 1,
rtac mult_assoc 1, rtac (mult_commute RS arg_cong) 1]);
@@ -112,7 +112,7 @@
(fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]);
(*Addition is the inverse of subtraction: if n<=m then n+(m-n) = m. *)
-val [prem] = goal Arith.thy "[| ~ m<n |] ==> n+(m-n) = m::nat";
+val [prem] = goal Arith.thy "[| ~ m<n |] ==> n+(m-n) = (m::nat)";
by (rtac (prem RS rev_mp) 1);
by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
by (ALLGOALS(asm_simp_tac arith_ss));
@@ -223,14 +223,14 @@
(**** Additional theorems about "less than" ****)
-goal Arith.thy "n <= (m + n)::nat";
+goal Arith.thy "n <= ((m + n)::nat)";
by (nat_ind_tac "m" 1);
by (ALLGOALS(simp_tac arith_ss));
by (etac le_trans 1);
by (rtac (lessI RS less_imp_le) 1);
val le_add2 = result();
-goal Arith.thy "n <= (n + m)::nat";
+goal Arith.thy "n <= ((n + m)::nat)";
by (simp_tac (arith_ss addsimps add_ac) 1);
by (rtac le_add2 1);
val le_add1 = result();
@@ -238,7 +238,7 @@
val less_add_Suc1 = standard (lessI RS (le_add1 RS le_less_trans));
val less_add_Suc2 = standard (lessI RS (le_add2 RS le_less_trans));
-goal Arith.thy "m+k<=n --> m<=n::nat";
+goal Arith.thy "m+k<=n --> m<=(n::nat)";
by (nat_ind_tac "k" 1);
by (ALLGOALS (asm_simp_tac arith_ss));
by (fast_tac (HOL_cs addDs [Suc_leD]) 1);