--- a/Arith.ML Thu Nov 24 20:31:09 1994 +0100
+++ b/Arith.ML Fri Nov 25 09:12:16 1994 +0100
@@ -21,13 +21,13 @@
val diff_0 = diff_def RS def_nat_rec_0;
-val diff_0_eq_0 = prove_goalw Arith.thy [diff_def, pred_def]
+qed_goalw "diff_0_eq_0" Arith.thy [diff_def, pred_def]
"0 - n = 0"
(fn _ => [nat_ind_tac "n" 1, ALLGOALS(asm_simp_tac nat_ss)]);
(*Must simplify BEFORE the induction!! (Else we get a critical pair)
Suc(m) - Suc(n) rewrites to pred(Suc(m) - n) *)
-val diff_Suc_Suc = prove_goalw Arith.thy [diff_def, pred_def]
+qed_goalw "diff_Suc_Suc" Arith.thy [diff_def, pred_def]
"Suc(m) - Suc(n) = m - n"
(fn _ =>
[simp_tac nat_ss 1, nat_ind_tac "n" 1, ALLGOALS(asm_simp_tac nat_ss)]);
@@ -44,23 +44,23 @@
(*** Addition ***)
-val add_0_right = prove_goal Arith.thy "m + 0 = m"
+qed_goal "add_0_right" Arith.thy "m + 0 = m"
(fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]);
-val add_Suc_right = prove_goal Arith.thy "m + Suc(n) = Suc(m+n)"
+qed_goal "add_Suc_right" Arith.thy "m + Suc(n) = Suc(m+n)"
(fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]);
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)"
+qed_goal "add_assoc" 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)"
+qed_goal "add_commute" 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)"
+qed_goal "add_left_commute" 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]);
@@ -71,36 +71,36 @@
(*** Multiplication ***)
(*right annihilation in product*)
-val mult_0_right = prove_goal Arith.thy "m * 0 = 0"
+qed_goal "mult_0_right" Arith.thy "m * 0 = 0"
(fn _ => [nat_ind_tac "m" 1, ALLGOALS(asm_simp_tac arith_ss)]);
(*right Sucessor law for multiplication*)
-val mult_Suc_right = prove_goal Arith.thy "m * Suc(n) = m + (m * n)"
+qed_goal "mult_Suc_right" Arith.thy "m * Suc(n) = m + (m * n)"
(fn _ => [nat_ind_tac "m" 1,
ALLGOALS(asm_simp_tac (arith_ss addsimps add_ac))]);
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)"
+qed_goal "mult_commute" 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)"
+qed_goal "add_mult_distrib" 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)"
+qed_goal "add_mult_distrib2" 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)"
+qed_goal "mult_assoc" 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)"
+qed_goal "mult_left_commute" 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]);
@@ -108,7 +108,7 @@
(*** Difference ***)
-val diff_self_eq_0 = prove_goal Arith.thy "m - m = 0"
+qed_goal "diff_self_eq_0" Arith.thy "m - m = 0"
(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. *)