--- a/src/HOL/Integ/IntArith.ML Fri Apr 21 11:36:00 2000 +0200
+++ b/src/HOL/Integ/IntArith.ML Sun Apr 23 11:33:41 2000 +0200
@@ -5,7 +5,6 @@
Simprocs and decision procedure for linear arithmetic.
*)
-
(*** Simprocs for numeric literals ***)
(** Combining of literal coefficients in sums of products **)
@@ -23,6 +22,275 @@
qed "zle_iff_zdiff_zle_0";
+(** For cancel_numerals **)
+
+Goal "!!i::int. ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)";
+by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1);
+qed "diff_add_eq1";
+
+Goal "!!i::int. ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))";
+by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]) 1);
+qed "diff_add_eq2";
+
+val rel_iff_rel_0_rls = map (inst "y" "?u+?v")
+ [zless_iff_zdiff_zless_0, eq_iff_zdiff_eq_0,
+ zle_iff_zdiff_zle_0] @
+ map (inst "y" "n")
+ [zless_iff_zdiff_zless_0, eq_iff_zdiff_eq_0,
+ zle_iff_zdiff_zle_0];
+
+Goal "!!i::int. (i*u + m = j*u + n) = ((i-j)*u + m = n)";
+by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@
+ zadd_ac@rel_iff_rel_0_rls) 1);
+qed "eq_add_iff1";
+
+Goal "!!i::int. (i*u + m = j*u + n) = (m = (j-i)*u + n)";
+by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@
+ zadd_ac@rel_iff_rel_0_rls) 1);
+qed "eq_add_iff2";
+
+Goal "!!i::int. (i*u + m < j*u + n) = ((i-j)*u + m < n)";
+by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@
+ zadd_ac@rel_iff_rel_0_rls) 1);
+qed "less_add_iff1";
+
+Goal "!!i::int. (i*u + m < j*u + n) = (m < (j-i)*u + n)";
+by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@
+ zadd_ac@rel_iff_rel_0_rls) 1);
+qed "less_add_iff2";
+
+Goal "!!i::int. (i*u + m <= j*u + n) = ((i-j)*u + m <= n)";
+by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]@
+ zadd_ac@rel_iff_rel_0_rls) 1);
+qed "le_add_iff1";
+
+Goal "!!i::int. (i*u + m <= j*u + n) = (m <= (j-i)*u + n)";
+by (asm_simp_tac (simpset() addsimps [zdiff_def, zadd_zmult_distrib]
+ @zadd_ac@rel_iff_rel_0_rls) 1);
+qed "le_add_iff2";
+
+
+structure Int_Numeral_Simprocs =
+struct
+
+(*Utilities*)
+
+fun mk_numeral n = HOLogic.number_of_const HOLogic.intT $
+ NumeralSyntax.mk_bin n;
+
+(*Decodes a binary INTEGER*)
+fun dest_numeral (Const("Numeral.number_of", _) $ w) = NumeralSyntax.dest_bin w
+ | dest_numeral t = raise TERM("dest_numeral", [t]);
+
+fun find_first_numeral past (t::terms) =
+ ((dest_numeral t, rev past @ terms)
+ handle TERM _ => find_first_numeral (t::past) terms)
+ | find_first_numeral past [] = raise TERM("find_first_numeral", []);
+
+val zero = mk_numeral 0;
+val mk_plus = HOLogic.mk_binop "op +";
+
+val uminus_const = Const ("uminus", HOLogic.intT --> HOLogic.intT);
+
+(*Thus mk_sum[t] yields t+#0; longer sums don't have a trailing zero*)
+fun mk_sum [] = zero
+ | mk_sum [t,u] = mk_plus (t, u)
+ | mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
+
+val dest_plus = HOLogic.dest_bin "op +" HOLogic.intT;
+
+(*decompose additions AND subtractions as a sum*)
+fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
+ dest_summing (pos, t, dest_summing (pos, u, ts))
+ | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
+ dest_summing (pos, t, dest_summing (not pos, u, ts))
+ | dest_summing (pos, t, ts) =
+ if pos then t::ts else uminus_const$t :: ts;
+
+fun dest_sum t = dest_summing (true, t, []);
+
+val mk_diff = HOLogic.mk_binop "op -";
+val dest_diff = HOLogic.dest_bin "op -" HOLogic.intT;
+
+val one = mk_numeral 1;
+val mk_times = HOLogic.mk_binop "op *";
+
+fun mk_prod [] = one
+ | mk_prod [t] = t
+ | mk_prod (t :: ts) = if t = one then mk_prod ts
+ else mk_times (t, mk_prod ts);
+
+val dest_times = HOLogic.dest_bin "op *" HOLogic.intT;
+
+fun dest_prod t =
+ let val (t,u) = dest_times t
+ in dest_prod t @ dest_prod u end
+ handle TERM _ => [t];
+
+(*DON'T do the obvious simplifications; that would create special cases*)
+fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts);
+
+(*Express t as a product of (possibly) a numeral with other sorted terms*)
+fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
+ | dest_coeff sign t =
+ let val ts = sort Term.term_ord (dest_prod t)
+ val (n, ts') = find_first_numeral [] ts
+ handle TERM _ => (1, ts)
+ in (sign*n, mk_prod ts') end;
+
+(*Find first coefficient-term THAT MATCHES u*)
+fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
+ | find_first_coeff past u (t::terms) =
+ let val (n,u') = dest_coeff 1 t
+ in if u aconv u' then (n, rev past @ terms)
+ else find_first_coeff (t::past) u terms
+ end
+ handle TERM _ => find_first_coeff (t::past) u terms;
+
+
+(*Simplify #1*n and n*#1 to n*)
+val add_0s = [zadd_0, zadd_0_right];
+val mult_1s = [zmult_1, zmult_1_right, zmult_minus1, zmult_minus1_right];
+
+(*To perform binary arithmetic*)
+val bin_simps = [number_of_add RS sym, add_number_of_left] @
+ bin_arith_simps @ bin_rel_simps;
+
+(*To let us treat subtraction as addition*)
+val diff_simps = [zdiff_def, zminus_zadd_distrib, zminus_zminus];
+
+val mk_eqv = HOLogic.mk_Trueprop o HOLogic.mk_eq;
+
+fun prove_conv tacs sg (t, u) =
+ if t aconv u then None
+ else
+ Some
+ (mk_meta_eq (prove_goalw_cterm [] (cterm_of sg (mk_eqv (t, u)))
+ (K tacs))
+ handle ERROR => error
+ ("The error(s) above occurred while trying to prove " ^
+ (string_of_cterm (cterm_of sg (mk_eqv (t, u))))));
+
+fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc;
+fun prep_pat s = Thm.read_cterm (Theory.sign_of Int.thy) (s, HOLogic.termT);
+val prep_pats = map prep_pat;
+
+structure CancelNumeralsCommon =
+ struct
+ val mk_sum = mk_sum
+ val dest_sum = dest_sum
+ val mk_coeff = mk_coeff
+ val dest_coeff = dest_coeff 1
+ val find_first_coeff = find_first_coeff []
+ val prove_conv = prove_conv
+ val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@zadd_ac))
+ THEN ALLGOALS
+ (simp_tac (HOL_ss addsimps [zmult_zminus_right RS sym]@bin_simps@zmult_ac))
+ val numeral_simp_tac = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps))
+ end;
+
+
+(* int eq *)
+structure EqCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+ val mk_bal = HOLogic.mk_eq
+ val dest_bal = HOLogic.dest_bin "op =" HOLogic.intT
+ val bal_add1 = eq_add_iff1 RS trans
+ val bal_add2 = eq_add_iff2 RS trans
+);
+
+(* int less *)
+structure LessCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+ val mk_bal = HOLogic.mk_binrel "op <"
+ val dest_bal = HOLogic.dest_bin "op <" HOLogic.intT
+ val bal_add1 = less_add_iff1 RS trans
+ val bal_add2 = less_add_iff2 RS trans
+);
+
+(* int le *)
+structure LeCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+ val mk_bal = HOLogic.mk_binrel "op <="
+ val dest_bal = HOLogic.dest_bin "op <=" HOLogic.intT
+ val bal_add1 = le_add_iff1 RS trans
+ val bal_add2 = le_add_iff2 RS trans
+);
+
+(* int diff *)
+structure DiffCancelNumerals = CancelNumeralsFun
+ (open CancelNumeralsCommon
+ val mk_bal = HOLogic.mk_binop "op -"
+ val dest_bal = HOLogic.dest_bin "op -" HOLogic.intT
+ val bal_add1 = diff_add_eq1 RS trans
+ val bal_add2 = diff_add_eq2 RS trans
+);
+
+
+val cancel_numerals =
+ map prep_simproc
+ [("inteq_cancel_numerals",
+ prep_pats ["(l::int) + m = n", "(l::int) = m + n",
+ "(l::int) - m = n", "(l::int) = m - n",
+ "(l::int) * m = n", "(l::int) = m * n"],
+ EqCancelNumerals.proc),
+ ("intless_cancel_numerals",
+ prep_pats ["(l::int) + m < n", "(l::int) < m + n",
+ "(l::int) - m < n", "(l::int) < m - n",
+ "(l::int) * m < n", "(l::int) < m * n"],
+ LessCancelNumerals.proc),
+ ("intle_cancel_numerals",
+ prep_pats ["(l::int) + m <= n", "(l::int) <= m + n",
+ "(l::int) - m <= n", "(l::int) <= m - n",
+ "(l::int) * m <= n", "(l::int) <= m * n"],
+ LeCancelNumerals.proc),
+ ("intdiff_cancel_numerals",
+ prep_pats ["((l::int) + m) - n", "(l::int) - (m + n)",
+ "((l::int) - m) - n", "(l::int) - (m - n)",
+ "(l::int) * m - n", "(l::int) - m * n"],
+ DiffCancelNumerals.proc)];
+
+end;
+
+
+Addsimprocs Int_Numeral_Simprocs.cancel_numerals;
+
+(*examples:
+print_depth 22;
+set proof_timing;
+set trace_simp;
+fun test s = (Goal s; by (Simp_tac 1));
+
+test "#2*u = (u::int)";
+test "(i + j + #12 + (k::int)) - #15 = y";
+test "(i + j + #12 + (k::int)) - #5 = y";
+
+test "y - b < (b::int)";
+test "y - (#3*b + c) < (b::int) - #2*c";
+
+test "(#2*x + (u*v) + y) - v*#3*u = (w::int)";
+test "(#2*x*u*v + (u*v)*#4 + y) - v*u*#4 = (w::int)";
+test "(#2*x*u*v + (u*v)*#4 + y) - v*u = (w::int)";
+
+test "(i + j + #12 + (k::int)) = u + #15 + y";
+test "(i + j*#2 + #12 + (k::int)) = j + #5 + y";
+
+test "#2*y + #3*z + #6*w + #2*y + #3*z + #2*u = #2*y' + #3*z' + #6*w' + #2*y' + #3*z' + u + (vv::int)";
+
+(*negative numerals*)
+test "(i + j + #-2 + (k::int)) - (u + #5 + y) = zz";
+test "(i + j + #-3 + (k::int)) < u + #5 + y";
+test "(i + j + #3 + (k::int)) < u + #-6 + y";
+test "(i + j + #-12 + (k::int)) - #15 = y";
+test "(i + j + #12 + (k::int)) - #-15 = y";
+test "(i + j + #-12 + (k::int)) - #-15 = y";
+*)
+
+
+
+(****************************************************************************************************************************************************************************************************************************************************************
+
+
structure Int_CC_Data : COMBINE_COEFF_DATA =
struct
val ss = HOL_ss
@@ -52,6 +320,7 @@
structure Int_CC = Combine_Coeff (Int_CC_Data);
Addsimprocs [Int_CC.sum_conv, Int_CC.rel_conv];
+****************************************************************)
(** Constant folding for integer plus and times **)
@@ -122,8 +391,9 @@
val add_rules = simp_thms @ bin_arith_simps @ bin_rel_simps @
[int_0,zmult_0,zmult_0_right];
-val simprocs = [Int_Cancel.sum_conv, Int_Cancel.rel_conv,
- Int_CC.sum_conv, Int_CC.rel_conv];
+val simprocs = Int_Numeral_Simprocs.cancel_numerals@
+ [Int_Cancel.sum_conv, Int_Cancel.rel_conv (*****,
+ Int_CC.sum_conv, Int_CC.rel_conv***)];
val add_mono_thms =
map (fn s => prove_goal Int.thy s
@@ -192,18 +462,6 @@
(* End of linear arithmetic *)
(*---------------------------------------------------------------------------*)
-(** Simplification of arithmetic when nested to the right **)
-
-Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::int)";
-by (simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1);
-qed "add_number_of_left";
-
-Goal "number_of v * (number_of w * z) = (number_of(bin_mult v w) * z::int)";
-by (simp_tac (simpset() addsimps [zmult_assoc RS sym]) 1);
-qed "mult_number_of_left";
-
-Addsimps [add_number_of_left, mult_number_of_left];
-
(** Simplification of inequalities involving numerical constants **)
Goal "(w <= z + (#1::int)) = (w<=z | w = z + (#1::int))";