--- a/src/HOL/Arith_Tools.thy Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/Arith_Tools.thy Wed Nov 28 09:01:34 2007 +0100
@@ -550,7 +550,7 @@
"add_Suc", "add_number_of_left", "mult_number_of_left",
"Suc_eq_add_numeral_1"])@
(map (fn s => thm s RS sym) ["numeral_1_eq_1", "numeral_0_eq_0"])
- @ arith_simps@ nat_arith @ rel_simps
+ @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps}
val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"},
@{thm "divide_Numeral1"},
@{thm "Ring_and_Field.divide_zero"}, @{thm "divide_Numeral0"},
--- a/src/HOL/Complex/ex/mirtac.ML Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/Complex/ex/mirtac.ML Wed Nov 28 09:01:34 2007 +0100
@@ -22,7 +22,7 @@
"add_Suc", "add_number_of_left", "mult_number_of_left",
"Suc_eq_add_numeral_1"])@
(map (fn s => thm s RS sym) ["numeral_1_eq_1", "numeral_0_eq_0"])
- @ arith_simps@ nat_arith @ rel_simps
+ @ @{thms arith_simps} @ nat_arith @ @{thms rel_simps}
val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"},
@{thm "real_of_nat_number_of"},
@{thm "real_of_nat_Suc"}, @{thm "real_of_nat_one"}, @{thm "real_of_one"},
--- a/src/HOL/IntArith.thy Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/IntArith.thy Wed Nov 28 09:01:34 2007 +0100
@@ -419,12 +419,4 @@
by (simp add:inj_on_def surj_def) (blast intro:sym)
qed
-subsection {* Legacy ML bindings *}
-
-ML {*
-val of_int_number_of_eq = @{thm of_int_number_of_eq};
-val nat_0 = @{thm nat_0};
-val nat_1 = @{thm nat_1};
-*}
-
end
--- a/src/HOL/NatBin.thy Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/NatBin.thy Wed Nov 28 09:01:34 2007 +0100
@@ -652,8 +652,7 @@
ML
{*
-val numerals = thms"numerals";
-val numeral_ss = simpset() addsimps numerals;
+val numeral_ss = simpset() addsimps @{thms numerals};
val nat_bin_arith_setup =
LinArith.map_data
@@ -662,8 +661,8 @@
inj_thms = inj_thms,
lessD = lessD, neqE = neqE,
simpset = simpset addsimps [Suc_nat_number_of, int_nat_number_of,
- not_neg_number_of_Pls,
- neg_number_of_Min,neg_number_of_BIT]})
+ @{thm not_neg_number_of_Pls}, @{thm neg_number_of_Min},
+ @{thm neg_number_of_BIT}]})
*}
declaration {* K nat_bin_arith_setup *}
@@ -834,79 +833,4 @@
"(k*m) div (k*n) = (if k = (0::nat) then 0 else m div n)"
by (simp add: nat_mult_div_cancel1)
-
-subsection {* legacy ML bindings *}
-
-ML
-{*
-val eq_nat_nat_iff = thm"eq_nat_nat_iff";
-val eq_nat_number_of = thm"eq_nat_number_of";
-val less_nat_number_of = thm"less_nat_number_of";
-val power2_eq_square = thm "power2_eq_square";
-val zero_le_power2 = thm "zero_le_power2";
-val zero_less_power2 = thm "zero_less_power2";
-val zero_eq_power2 = thm "zero_eq_power2";
-val abs_power2 = thm "abs_power2";
-val power2_abs = thm "power2_abs";
-val power2_minus = thm "power2_minus";
-val power_minus1_even = thm "power_minus1_even";
-val power_minus_even = thm "power_minus_even";
-val odd_power_less_zero = thm "odd_power_less_zero";
-val odd_0_le_power_imp_0_le = thm "odd_0_le_power_imp_0_le";
-
-val Suc_pred' = thm"Suc_pred'";
-val expand_Suc = thm"expand_Suc";
-val Suc_eq_add_numeral_1 = thm"Suc_eq_add_numeral_1";
-val Suc_eq_add_numeral_1_left = thm"Suc_eq_add_numeral_1_left";
-val add_eq_if = thm"add_eq_if";
-val mult_eq_if = thm"mult_eq_if";
-val power_eq_if = thm"power_eq_if";
-val eq_number_of_0 = thm"eq_number_of_0";
-val eq_0_number_of = thm"eq_0_number_of";
-val less_0_number_of = thm"less_0_number_of";
-val neg_imp_number_of_eq_0 = thm"neg_imp_number_of_eq_0";
-val eq_number_of_Suc = thm"eq_number_of_Suc";
-val Suc_eq_number_of = thm"Suc_eq_number_of";
-val less_number_of_Suc = thm"less_number_of_Suc";
-val less_Suc_number_of = thm"less_Suc_number_of";
-val le_number_of_Suc = thm"le_number_of_Suc";
-val le_Suc_number_of = thm"le_Suc_number_of";
-val eq_number_of_BIT_BIT = thm"eq_number_of_BIT_BIT";
-val eq_number_of_BIT_Pls = thm"eq_number_of_BIT_Pls";
-val eq_number_of_BIT_Min = thm"eq_number_of_BIT_Min";
-val eq_number_of_Pls_Min = thm"eq_number_of_Pls_Min";
-val of_nat_number_of_eq = thm"of_nat_number_of_eq";
-val nat_power_eq = thm"nat_power_eq";
-val power_nat_number_of = thm"power_nat_number_of";
-val zpower_number_of_even = thm"zpower_number_of_even";
-val zpower_number_of_odd = thm"zpower_number_of_odd";
-val nat_number_of_Pls = thm"nat_number_of_Pls";
-val nat_number_of_Min = thm"nat_number_of_Min";
-val Let_Suc = thm"Let_Suc";
-
-val nat_number = thms"nat_number";
-
-val nat_number_of_add_left = thm"nat_number_of_add_left";
-val nat_number_of_mult_left = thm"nat_number_of_mult_left";
-val left_add_mult_distrib = thm"left_add_mult_distrib";
-val nat_diff_add_eq1 = thm"nat_diff_add_eq1";
-val nat_diff_add_eq2 = thm"nat_diff_add_eq2";
-val nat_eq_add_iff1 = thm"nat_eq_add_iff1";
-val nat_eq_add_iff2 = thm"nat_eq_add_iff2";
-val nat_less_add_iff1 = thm"nat_less_add_iff1";
-val nat_less_add_iff2 = thm"nat_less_add_iff2";
-val nat_le_add_iff1 = thm"nat_le_add_iff1";
-val nat_le_add_iff2 = thm"nat_le_add_iff2";
-val nat_mult_le_cancel1 = thm"nat_mult_le_cancel1";
-val nat_mult_less_cancel1 = thm"nat_mult_less_cancel1";
-val nat_mult_eq_cancel1 = thm"nat_mult_eq_cancel1";
-val nat_mult_div_cancel1 = thm"nat_mult_div_cancel1";
-val nat_mult_le_cancel_disj = thm"nat_mult_le_cancel_disj";
-val nat_mult_less_cancel_disj = thm"nat_mult_less_cancel_disj";
-val nat_mult_eq_cancel_disj = thm"nat_mult_eq_cancel_disj";
-val nat_mult_div_cancel_disj = thm"nat_mult_div_cancel_disj";
-
-val power_minus_even = thm"power_minus_even";
-*}
-
end
--- a/src/HOL/Tools/Groebner_Basis/normalizer.ML Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/Tools/Groebner_Basis/normalizer.ML Wed Nov 28 09:01:34 2007 +0100
@@ -46,9 +46,9 @@
val is_numeral = can dest_numeral;
val numeral01_conv = Simplifier.rewrite
- (HOL_basic_ss addsimps [numeral_1_eq_1, numeral_0_eq_0]);
+ (HOL_basic_ss addsimps [@{thm numeral_1_eq_1}, @{thm numeral_0_eq_0}]);
val zero1_numeral_conv =
- Simplifier.rewrite (HOL_basic_ss addsimps [numeral_1_eq_1 RS sym, numeral_0_eq_0 RS sym]);
+ Simplifier.rewrite (HOL_basic_ss addsimps [@{thm numeral_1_eq_1} RS sym, @{thm numeral_0_eq_0} RS sym]);
fun zerone_conv cv = zero1_numeral_conv then_conv cv then_conv numeral01_conv;
val natarith = [@{thm "add_nat_number_of"}, @{thm "diff_nat_number_of"},
@{thm "mult_nat_number_of"}, @{thm "eq_nat_number_of"},
@@ -57,9 +57,10 @@
zerone_conv
(Simplifier.rewrite
(HOL_basic_ss
- addsimps arith_simps @ natarith @ rel_simps
- @ [if_False, if_True, @{thm add_0}, @{thm add_Suc}, add_number_of_left, Suc_eq_add_numeral_1]
- @ map (fn th => th RS sym) numerals));
+ addsimps @{thms arith_simps} @ natarith @ @{thms rel_simps}
+ @ [if_False, if_True, @{thm add_0}, @{thm add_Suc},
+ @{thm add_number_of_left}, @{thm Suc_eq_add_numeral_1}]
+ @ map (fn th => th RS sym) @{thms numerals}));
val nat_mul_conv = nat_add_conv;
val zeron_tm = @{cterm "0::nat"};
@@ -608,7 +609,7 @@
end;
val nat_arith = @{thms "nat_arith"};
-val nat_exp_ss = HOL_basic_ss addsimps (nat_number @ nat_arith @ arith_simps @ rel_simps)
+val nat_exp_ss = HOL_basic_ss addsimps (@{thms nat_number} @ nat_arith @ @{thms arith_simps} @ @{thms rel_simps})
addsimps [Let_def, if_False, if_True, @{thm add_0}, @{thm add_Suc}];
fun simple_cterm_ord t u = Term.term_ord (term_of t, term_of u) = LESS;
--- a/src/HOL/Wellfounded_Relations.thy Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/Wellfounded_Relations.thy Wed Nov 28 09:01:34 2007 +0100
@@ -266,33 +266,4 @@
apply (intro wf_trancl wf_pred_nat)
done
-
-ML
-{*
-val less_than_def = thm "less_than_def";
-val measure_def = thm "measure_def";
-val lex_prod_def = thm "lex_prod_def";
-val finite_psubset_def = thm "finite_psubset_def";
-
-val wf_less_than = thm "wf_less_than";
-val trans_less_than = thm "trans_less_than";
-val less_than_iff = thm "less_than_iff";
-val full_nat_induct = thm "full_nat_induct";
-val wf_inv_image = thm "wf_inv_image";
-val wf_measure = thm "wf_measure";
-val measure_induct = thm "measure_induct";
-val wf_lex_prod = thm "wf_lex_prod";
-val trans_lex_prod = thm "trans_lex_prod";
-val wf_finite_psubset = thm "wf_finite_psubset";
-val trans_finite_psubset = thm "trans_finite_psubset";
-val finite_acyclic_wf = thm "finite_acyclic_wf";
-val finite_acyclic_wf_converse = thm "finite_acyclic_wf_converse";
-val wf_iff_acyclic_if_finite = thm "wf_iff_acyclic_if_finite";
-val wf_weak_decr_stable = thm "wf_weak_decr_stable";
-val weak_decr_stable = thm "weak_decr_stable";
-val same_fstI = thm "same_fstI";
-val wf_same_fst = thm "wf_same_fst";
-*}
-
-
end
--- a/src/HOL/int_arith1.ML Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/int_arith1.ML Wed Nov 28 09:01:34 2007 +0100
@@ -5,91 +5,6 @@
Simprocs and decision procedure for linear arithmetic.
*)
-(** Misc ML bindings **)
-
-val succ_Pls = thm "succ_Pls";
-val succ_Min = thm "succ_Min";
-val succ_1 = thm "succ_1";
-val succ_0 = thm "succ_0";
-
-val pred_Pls = thm "pred_Pls";
-val pred_Min = thm "pred_Min";
-val pred_1 = thm "pred_1";
-val pred_0 = thm "pred_0";
-
-val minus_Pls = thm "minus_Pls";
-val minus_Min = thm "minus_Min";
-val minus_1 = thm "minus_1";
-val minus_0 = thm "minus_0";
-
-val add_Pls = thm "add_Pls";
-val add_Min = thm "add_Min";
-val add_BIT_11 = thm "add_BIT_11";
-val add_BIT_10 = thm "add_BIT_10";
-val add_BIT_0 = thm "add_BIT_0";
-val add_Pls_right = thm "add_Pls_right";
-val add_Min_right = thm "add_Min_right";
-
-val mult_Pls = thm "mult_Pls";
-val mult_Min = thm "mult_Min";
-val mult_num1 = thm "mult_num1";
-val mult_num0 = thm "mult_num0";
-
-val neg_def = thm "neg_def";
-val iszero_def = thm "iszero_def";
-
-val number_of_succ = thm "number_of_succ";
-val number_of_pred = thm "number_of_pred";
-val number_of_minus = thm "number_of_minus";
-val number_of_add = thm "number_of_add";
-val diff_number_of_eq = thm "diff_number_of_eq";
-val number_of_mult = thm "number_of_mult";
-val double_number_of_BIT = thm "double_number_of_BIT";
-val numeral_0_eq_0 = thm "numeral_0_eq_0";
-val numeral_1_eq_1 = thm "numeral_1_eq_1";
-val numeral_m1_eq_minus_1 = thm "numeral_m1_eq_minus_1";
-val mult_minus1 = thm "mult_minus1";
-val mult_minus1_right = thm "mult_minus1_right";
-val minus_number_of_mult = thm "minus_number_of_mult";
-val zero_less_nat_eq = thm "zero_less_nat_eq";
-val eq_number_of_eq = thm "eq_number_of_eq";
-val iszero_number_of_Pls = thm "iszero_number_of_Pls";
-val nonzero_number_of_Min = thm "nonzero_number_of_Min";
-val iszero_number_of_BIT = thm "iszero_number_of_BIT";
-val iszero_number_of_0 = thm "iszero_number_of_0";
-val iszero_number_of_1 = thm "iszero_number_of_1";
-val less_number_of_eq_neg = thm "less_number_of_eq_neg";
-val le_number_of_eq = thm "le_number_of_eq";
-val not_neg_number_of_Pls = thm "not_neg_number_of_Pls";
-val neg_number_of_Min = thm "neg_number_of_Min";
-val neg_number_of_BIT = thm "neg_number_of_BIT";
-val le_number_of_eq_not_less = thm "le_number_of_eq_not_less";
-val abs_number_of = thm "abs_number_of";
-val number_of_reorient = thm "number_of_reorient";
-val add_number_of_left = thm "add_number_of_left";
-val mult_number_of_left = thm "mult_number_of_left";
-val add_number_of_diff1 = thm "add_number_of_diff1";
-val add_number_of_diff2 = thm "add_number_of_diff2";
-val less_iff_diff_less_0 = thm "less_iff_diff_less_0";
-val eq_iff_diff_eq_0 = thm "eq_iff_diff_eq_0";
-val le_iff_diff_le_0 = thm "le_iff_diff_le_0";
-
-val arith_extra_simps = thms "arith_extra_simps";
-val arith_simps = thms "arith_simps";
-val rel_simps = thms "rel_simps";
-
-val zless_imp_add1_zle = thm "zless_imp_add1_zle";
-
-val combine_common_factor = thm "combine_common_factor";
-val eq_add_iff1 = thm "eq_add_iff1";
-val eq_add_iff2 = thm "eq_add_iff2";
-val less_add_iff1 = thm "less_add_iff1";
-val less_add_iff2 = thm "less_add_iff2";
-val le_add_iff1 = thm "le_add_iff1";
-val le_add_iff2 = thm "le_add_iff2";
-
-val arith_special = thms "arith_special";
-
structure Int_Numeral_Base_Simprocs =
struct
fun prove_conv tacs ctxt (_: thm list) (t, u) =
@@ -112,19 +27,19 @@
(*reorientation simprules using ==, for the following simproc*)
val meta_zero_reorient = @{thm zero_reorient} RS eq_reflection
val meta_one_reorient = @{thm one_reorient} RS eq_reflection
- val meta_number_of_reorient = number_of_reorient RS eq_reflection
+ val meta_number_of_reorient = @{thm number_of_reorient} RS eq_reflection
(*reorientation simplification procedure: reorients (polymorphic)
0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
fun reorient_proc sg _ (_ $ t $ u) =
case u of
- Const(@{const_name HOL.zero}, _) => NONE
+ Const(@{const_name HOL.zero}, _) => NONE
| Const(@{const_name HOL.one}, _) => NONE
| Const(@{const_name Numeral.number_of}, _) $ _ => NONE
| _ => SOME (case t of
- Const(@{const_name HOL.zero}, _) => meta_zero_reorient
- | Const(@{const_name HOL.one}, _) => meta_one_reorient
- | Const(@{const_name Numeral.number_of}, _) $ _ => meta_number_of_reorient)
+ Const(@{const_name HOL.zero}, _) => meta_zero_reorient
+ | Const(@{const_name HOL.one}, _) => meta_one_reorient
+ | Const(@{const_name Numeral.number_of}, _) $ _ => meta_number_of_reorient)
val reorient_simproc =
prep_simproc ("reorient_simproc", ["0=x", "1=x", "number_of w = x"], reorient_proc)
@@ -140,7 +55,7 @@
(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in Int_Numeral_Base_Simprocs
isn't complicated by the abstract 0 and 1.*)
-val numeral_syms = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym];
+val numeral_syms = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym];
(** New term ordering so that AC-rewriting brings numerals to the front **)
@@ -300,19 +215,19 @@
(*To perform binary arithmetic. The "left" rewriting handles patterns
created by the Int_Numeral_Base_Simprocs, such as 3 * (5 * x). *)
-val simps = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym,
- add_number_of_left, mult_number_of_left] @
- arith_simps @ rel_simps;
+val simps = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym,
+ @{thm add_number_of_left}, @{thm mult_number_of_left}] @
+ @{thms arith_simps} @ @{thms rel_simps};
(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms
during re-arrangement*)
val non_add_simps =
- subtract Thm.eq_thm [add_number_of_left, number_of_add RS sym] simps;
+ subtract Thm.eq_thm [@{thm add_number_of_left}, @{thm number_of_add} RS sym] simps;
(*To evaluate binary negations of coefficients*)
-val minus_simps = [numeral_m1_eq_minus_1 RS sym, number_of_minus RS sym,
- minus_1, minus_0, minus_Pls, minus_Min,
- pred_1, pred_0, pred_Pls, pred_Min];
+val minus_simps = [@{thm numeral_m1_eq_minus_1} RS sym, @{thm number_of_minus} RS sym,
+ @{thm minus_1}, @{thm minus_0}, @{thm minus_Pls}, @{thm minus_Min},
+ @{thm pred_1}, @{thm pred_0}, @{thm pred_Pls}, @{thm pred_Min}];
(*To let us treat subtraction as addition*)
val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}];
@@ -369,8 +284,8 @@
val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
val mk_bal = HOLogic.mk_eq
val dest_bal = HOLogic.dest_bin "op =" Term.dummyT
- val bal_add1 = eq_add_iff1 RS trans
- val bal_add2 = eq_add_iff2 RS trans
+ val bal_add1 = @{thm eq_add_iff1} RS trans
+ val bal_add2 = @{thm eq_add_iff2} RS trans
);
structure LessCancelNumerals = CancelNumeralsFun
@@ -378,8 +293,8 @@
val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
val mk_bal = HOLogic.mk_binrel @{const_name HOL.less}
val dest_bal = HOLogic.dest_bin @{const_name HOL.less} Term.dummyT
- val bal_add1 = less_add_iff1 RS trans
- val bal_add2 = less_add_iff2 RS trans
+ val bal_add1 = @{thm less_add_iff1} RS trans
+ val bal_add2 = @{thm less_add_iff2} RS trans
);
structure LeCancelNumerals = CancelNumeralsFun
@@ -387,8 +302,8 @@
val prove_conv = Int_Numeral_Base_Simprocs.prove_conv
val mk_bal = HOLogic.mk_binrel @{const_name HOL.less_eq}
val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} Term.dummyT
- val bal_add1 = le_add_iff1 RS trans
- val bal_add2 = le_add_iff2 RS trans
+ val bal_add1 = @{thm le_add_iff1} RS trans
+ val bal_add2 = @{thm le_add_iff2} RS trans
);
val cancel_numerals =
@@ -428,7 +343,7 @@
val dest_sum = dest_sum
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff 1
- val left_distrib = combine_common_factor RS trans
+ val left_distrib = @{thm combine_common_factor} RS trans
val prove_conv = Int_Numeral_Base_Simprocs.prove_conv_nohyps
val trans_tac = fn _ => trans_tac
@@ -458,7 +373,7 @@
val dest_sum = dest_sum
val mk_coeff = mk_fcoeff
val dest_coeff = dest_fcoeff 1
- val left_distrib = combine_common_factor RS trans
+ val left_distrib = @{thm combine_common_factor} RS trans
val prove_conv = Int_Numeral_Base_Simprocs.prove_conv_nohyps
val trans_tac = fn _ => trans_tac
@@ -631,7 +546,7 @@
end;
val add_rules =
- simp_thms @ arith_simps @ rel_simps @ arith_special @
+ simp_thms @ @{thms arith_simps} @ @{thms rel_simps} @ @{thms arith_special} @
[@{thm neg_le_iff_le}, @{thm numeral_0_eq_0}, @{thm numeral_1_eq_1},
@{thm minus_zero}, @{thm diff_minus}, @{thm left_minus}, @{thm right_minus},
@{thm mult_zero_left}, @{thm mult_zero_right}, @{thm mult_num1}, @{thm mult_1_right},
@@ -654,7 +569,7 @@
{add_mono_thms = add_mono_thms,
mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} :: mult_mono_thms,
inj_thms = nat_inj_thms @ inj_thms,
- lessD = lessD @ [zless_imp_add1_zle],
+ lessD = lessD @ [@{thm zless_imp_add1_zle}],
neqE = neqE,
simpset = simpset addsimps add_rules
addsimprocs Int_Numeral_Base_Simprocs
--- a/src/HOL/int_factor_simprocs.ML Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/int_factor_simprocs.ML Wed Nov 28 09:01:34 2007 +0100
@@ -19,7 +19,7 @@
and d = gcd(m,m') and n=m/d and n'=m'/d.
*)
-val rel_number_of = [eq_number_of_eq, less_number_of_eq_neg, le_number_of_eq];
+val rel_number_of = [@{thm eq_number_of_eq}, @{thm less_number_of_eq_neg}, @{thm le_number_of_eq}];
local
open Int_Numeral_Simprocs
@@ -216,7 +216,7 @@
(** Final simplification for the CancelFactor simprocs **)
val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq
- [@{thm mult_1_left}, @{thm mult_1_right}, @{thm zdiv_1}, numeral_1_eq_1];
+ [@{thm mult_1_left}, @{thm mult_1_right}, @{thm zdiv_1}, @{thm numeral_1_eq_1}];
fun cancel_simplify_meta_eq cancel_th ss th =
simplify_one ss (([th, cancel_th]) MRS trans);
--- a/src/HOL/nat_simprocs.ML Tue Nov 27 21:06:52 2007 +0100
+++ b/src/HOL/nat_simprocs.ML Wed Nov 28 09:01:34 2007 +0100
@@ -53,7 +53,7 @@
@{thm mult_nat_number_of}, @{thm nat_number_of_mult_left},
@{thm less_nat_number_of},
@{thm Let_number_of}, @{thm nat_number_of}] @
- arith_simps @ rel_simps;
+ @{thms arith_simps} @ @{thms rel_simps};
fun prep_simproc (name, pats, proc) =
Simplifier.simproc (the_context ()) name pats proc;