| author | wenzelm | 
| Fri, 10 Feb 2006 02:22:48 +0100 | |
| changeset 19001 | 64e4b5bc6443 | 
| parent 18708 | 4b3dadb4fe33 | 
| child 19044 | d4bc0ee9383a | 
| permissions | -rw-r--r-- | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 1 | (* Title: HOL/Integ/int_arith1.ML | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 3 | Authors: Larry Paulson and Tobias Nipkow | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 4 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 5 | Simprocs and decision procedure for linear arithmetic. | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 6 | *) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 7 | |
| 14329 | 8 | (** Misc ML bindings **) | 
| 9 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 10 | val bin_succ_Pls = thm"bin_succ_Pls"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 11 | val bin_succ_Min = thm"bin_succ_Min"; | 
| 15013 | 12 | val bin_succ_1 = thm"bin_succ_1"; | 
| 13 | val bin_succ_0 = thm"bin_succ_0"; | |
| 14 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 15 | val bin_pred_Pls = thm"bin_pred_Pls"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 16 | val bin_pred_Min = thm"bin_pred_Min"; | 
| 15013 | 17 | val bin_pred_1 = thm"bin_pred_1"; | 
| 18 | val bin_pred_0 = thm"bin_pred_0"; | |
| 19 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 20 | val bin_minus_Pls = thm"bin_minus_Pls"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 21 | val bin_minus_Min = thm"bin_minus_Min"; | 
| 15013 | 22 | val bin_minus_1 = thm"bin_minus_1"; | 
| 23 | val bin_minus_0 = thm"bin_minus_0"; | |
| 24 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 25 | val bin_add_Pls = thm"bin_add_Pls"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 26 | val bin_add_Min = thm"bin_add_Min"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 27 | val bin_add_BIT_11 = thm"bin_add_BIT_11"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 28 | val bin_add_BIT_10 = thm"bin_add_BIT_10"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 29 | val bin_add_BIT_0 = thm"bin_add_BIT_0"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 30 | val bin_add_Pls_right = thm"bin_add_Pls_right"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 31 | val bin_add_Min_right = thm"bin_add_Min_right"; | 
| 15013 | 32 | |
| 33 | val bin_mult_Pls = thm"bin_mult_Pls"; | |
| 34 | val bin_mult_Min = thm"bin_mult_Min"; | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 35 | val bin_mult_1 = thm"bin_mult_1"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 36 | val bin_mult_0 = thm"bin_mult_0"; | 
| 15013 | 37 | |
| 38 | val neg_def = thm "neg_def"; | |
| 39 | val iszero_def = thm "iszero_def"; | |
| 40 | ||
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 41 | val number_of_succ = thm"number_of_succ"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 42 | val number_of_pred = thm"number_of_pred"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 43 | val number_of_minus = thm"number_of_minus"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 44 | val number_of_add = thm"number_of_add"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 45 | val diff_number_of_eq = thm"diff_number_of_eq"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 46 | val number_of_mult = thm"number_of_mult"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 47 | val double_number_of_BIT = thm"double_number_of_BIT"; | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 48 | val numeral_0_eq_0 = thm"numeral_0_eq_0"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 49 | val numeral_1_eq_1 = thm"numeral_1_eq_1"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 50 | val numeral_m1_eq_minus_1 = thm"numeral_m1_eq_minus_1"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 51 | val mult_minus1 = thm"mult_minus1"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 52 | val mult_minus1_right = thm"mult_minus1_right"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 53 | val minus_number_of_mult = thm"minus_number_of_mult"; | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 54 | val zero_less_nat_eq = thm"zero_less_nat_eq"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 55 | val eq_number_of_eq = thm"eq_number_of_eq"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 56 | val iszero_number_of_Pls = thm"iszero_number_of_Pls"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 57 | val nonzero_number_of_Min = thm"nonzero_number_of_Min"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 58 | val iszero_number_of_BIT = thm"iszero_number_of_BIT"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 59 | val iszero_number_of_0 = thm"iszero_number_of_0"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 60 | val iszero_number_of_1 = thm"iszero_number_of_1"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 61 | val less_number_of_eq_neg = thm"less_number_of_eq_neg"; | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 62 | val le_number_of_eq = thm"le_number_of_eq"; | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 63 | val not_neg_number_of_Pls = thm"not_neg_number_of_Pls"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 64 | val neg_number_of_Min = thm"neg_number_of_Min"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 65 | val neg_number_of_BIT = thm"neg_number_of_BIT"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 66 | val le_number_of_eq_not_less = thm"le_number_of_eq_not_less"; | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 67 | val abs_number_of = thm"abs_number_of"; | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 68 | val number_of_reorient = thm"number_of_reorient"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 69 | val add_number_of_left = thm"add_number_of_left"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 70 | val mult_number_of_left = thm"mult_number_of_left"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 71 | val add_number_of_diff1 = thm"add_number_of_diff1"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 72 | val add_number_of_diff2 = thm"add_number_of_diff2"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 73 | val less_iff_diff_less_0 = thm"less_iff_diff_less_0"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 74 | val eq_iff_diff_eq_0 = thm"eq_iff_diff_eq_0"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 75 | val le_iff_diff_le_0 = thm"le_iff_diff_le_0"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 76 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 77 | val bin_arith_extra_simps = thms"bin_arith_extra_simps"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 78 | val bin_arith_simps = thms"bin_arith_simps"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 79 | val bin_rel_simps = thms"bin_rel_simps"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 80 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 81 | val zless_imp_add1_zle = thm "zless_imp_add1_zle"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 82 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 83 | val combine_common_factor = thm"combine_common_factor"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 84 | val eq_add_iff1 = thm"eq_add_iff1"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 85 | val eq_add_iff2 = thm"eq_add_iff2"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 86 | val less_add_iff1 = thm"less_add_iff1"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 87 | val less_add_iff2 = thm"less_add_iff2"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 88 | val le_add_iff1 = thm"le_add_iff1"; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 89 | val le_add_iff2 = thm"le_add_iff2"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 90 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 91 | val arith_special = thms"arith_special"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 92 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 93 | structure Bin_Simprocs = | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 94 | struct | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 95 | fun prove_conv tacs sg (hyps: thm list) xs (t, u) = | 
| 15531 | 96 | if t aconv u then NONE | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 97 | else | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 98 | let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)) | 
| 17956 | 99 | in SOME (Goal.prove sg xs [] eq (K (EVERY tacs))) end | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 100 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 101 | fun prove_conv_nohyps tacs sg = prove_conv tacs sg []; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 102 | fun prove_conv_nohyps_novars tacs sg = prove_conv tacs sg [] []; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 103 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 104 | fun prep_simproc (name, pats, proc) = | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 105 | Simplifier.simproc (Theory.sign_of (the_context())) name pats proc; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 106 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 107 |   fun is_numeral (Const("Numeral.number_of", _) $ w) = true
 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 108 | | is_numeral _ = false | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 109 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 110 | fun simplify_meta_eq f_number_of_eq f_eq = | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 111 | mk_meta_eq ([f_eq, f_number_of_eq] MRS trans) | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 112 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 113 | (*reorientation simprules using ==, for the following simproc*) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 114 | val meta_zero_reorient = zero_reorient RS eq_reflection | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 115 | val meta_one_reorient = one_reorient RS eq_reflection | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 116 | val meta_number_of_reorient = number_of_reorient RS eq_reflection | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 117 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 118 | (*reorientation simplification procedure: reorients (polymorphic) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 119 | 0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*) | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 120 | fun reorient_proc sg _ (_ $ t $ u) = | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 121 | case u of | 
| 15531 | 122 | 	Const("0", _) => NONE
 | 
| 123 |       | Const("1", _) => NONE
 | |
| 124 |       | Const("Numeral.number_of", _) $ _ => NONE
 | |
| 125 | | _ => SOME (case t of | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 126 | 		  Const("0", _) => meta_zero_reorient
 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 127 | 		| Const("1", _) => meta_one_reorient
 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 128 | 		| Const("Numeral.number_of", _) $ _ => meta_number_of_reorient)
 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 129 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 130 | val reorient_simproc = | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 131 |       prep_simproc ("reorient_simproc", ["0=x", "1=x", "number_of w = x"], reorient_proc)
 | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 132 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 133 | end; | 
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 134 | |
| 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 135 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 136 | Addsimps arith_special; | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 137 | Addsimprocs [Bin_Simprocs.reorient_simproc]; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 138 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 139 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 140 | structure Int_Numeral_Simprocs = | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 141 | struct | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 142 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 143 | (*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 144 | isn't complicated by the abstract 0 and 1.*) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 145 | val numeral_syms = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym]; | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 146 | |
| 14474 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 147 | (** New term ordering so that AC-rewriting brings numerals to the front **) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 148 | |
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 149 | (*Order integers by absolute value and then by sign. The standard integer | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 150 | ordering is not well-founded.*) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 151 | fun num_ord (i,j) = | 
| 15965 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 paulson parents: 
15921diff
changeset | 152 | (case IntInf.compare (IntInf.abs i, IntInf.abs j) of | 
| 16973 | 153 | EQUAL => int_ord (IntInf.sign i, IntInf.sign j) | 
| 14474 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 154 | | ord => ord); | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 155 | |
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 156 | (*This resembles Term.term_ord, but it puts binary numerals before other | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 157 | non-atomic terms.*) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 158 | local open Term | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 159 | in | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 160 | fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) = | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 161 | (case numterm_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 162 | | numterm_ord | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 163 |      (Const("Numeral.number_of", _) $ v, Const("Numeral.number_of", _) $ w) =
 | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 164 | num_ord (HOLogic.dest_binum v, HOLogic.dest_binum w) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 165 |   | numterm_ord (Const("Numeral.number_of", _) $ _, _) = LESS
 | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 166 |   | numterm_ord (_, Const("Numeral.number_of", _) $ _) = GREATER
 | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 167 | | numterm_ord (t, u) = | 
| 16973 | 168 | (case int_ord (size_of_term t, size_of_term u) of | 
| 14474 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 169 | EQUAL => | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 170 | let val (f, ts) = strip_comb t and (g, us) = strip_comb u in | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 171 | (case hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 172 | end | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 173 | | ord => ord) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 174 | and numterms_ord (ts, us) = list_ord numterm_ord (ts, us) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 175 | end; | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 176 | |
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 177 | fun numtermless tu = (numterm_ord tu = LESS); | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 178 | |
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 179 | (*Defined in this file, but perhaps needed only for simprocs of type nat.*) | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 180 | val num_ss = HOL_ss settermless numtermless; | 
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 181 | |
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 182 | |
| 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 paulson parents: 
14390diff
changeset | 183 | (** Utilities **) | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 184 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 185 | fun mk_numeral T n = HOLogic.number_of_const T $ HOLogic.mk_bin n; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 186 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 187 | (*Decodes a binary INTEGER*) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 188 | fun dest_numeral (Const("0", _)) = 0
 | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 189 |   | dest_numeral (Const("1", _)) = 1
 | 
| 13462 | 190 |   | dest_numeral (Const("Numeral.number_of", _) $ w) =
 | 
| 10890 | 191 | (HOLogic.dest_binum w | 
| 192 |       handle TERM _ => raise TERM("Int_Numeral_Simprocs.dest_numeral:1", [w]))
 | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 193 |   | dest_numeral t = raise TERM("Int_Numeral_Simprocs.dest_numeral:2", [t]);
 | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 194 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 195 | fun find_first_numeral past (t::terms) = | 
| 13462 | 196 | ((dest_numeral t, rev past @ terms) | 
| 197 | handle TERM _ => find_first_numeral (t::past) terms) | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 198 |   | find_first_numeral past [] = raise TERM("find_first_numeral", []);
 | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 199 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 200 | val mk_plus = HOLogic.mk_binop "op +"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 201 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 202 | fun mk_minus t = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 203 | let val T = Term.fastype_of t | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 204 |   in Const ("uminus", T --> T) $ t
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 205 | end; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 206 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 207 | (*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 208 | fun mk_sum T [] = mk_numeral T 0 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 209 | | mk_sum T [t,u] = mk_plus (t, u) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 210 | | mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 211 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 212 | (*this version ALWAYS includes a trailing zero*) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 213 | fun long_mk_sum T [] = mk_numeral T 0 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 214 | | long_mk_sum T (t :: ts) = mk_plus (t, mk_sum T ts); | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 215 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 216 | val dest_plus = HOLogic.dest_bin "op +" Term.dummyT; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 217 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 218 | (*decompose additions AND subtractions as a sum*) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 219 | fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) =
 | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 220 | dest_summing (pos, t, dest_summing (pos, u, ts)) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 221 |   | dest_summing (pos, Const ("op -", _) $ t $ u, ts) =
 | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 222 | dest_summing (pos, t, dest_summing (not pos, u, ts)) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 223 | | dest_summing (pos, t, ts) = | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 224 | if pos then t::ts else mk_minus t :: ts; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 225 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 226 | fun dest_sum t = dest_summing (true, t, []); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 227 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 228 | val mk_diff = HOLogic.mk_binop "op -"; | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 229 | val dest_diff = HOLogic.dest_bin "op -" Term.dummyT; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 230 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 231 | val mk_times = HOLogic.mk_binop "op *"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 232 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 233 | fun mk_prod T = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 234 | let val one = mk_numeral T 1 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 235 | fun mk [] = one | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 236 | | mk [t] = t | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 237 | | mk (t :: ts) = if t = one then mk ts else mk_times (t, mk ts) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 238 | in mk end; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 239 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 240 | (*This version ALWAYS includes a trailing one*) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 241 | fun long_mk_prod T [] = mk_numeral T 1 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 242 | | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 243 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 244 | val dest_times = HOLogic.dest_bin "op *" Term.dummyT; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 245 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 246 | fun dest_prod t = | 
| 13462 | 247 | let val (t,u) = dest_times t | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 248 | in dest_prod t @ dest_prod u end | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 249 | handle TERM _ => [t]; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 250 | |
| 13462 | 251 | (*DON'T do the obvious simplifications; that would create special cases*) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 252 | fun mk_coeff (k, t) = mk_times (mk_numeral (Term.fastype_of t) k, t); | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 253 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 254 | (*Express t as a product of (possibly) a numeral with other sorted terms*) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 255 | fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t
 | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 256 | | dest_coeff sign t = | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 257 | let val ts = sort Term.term_ord (dest_prod t) | 
| 13462 | 258 | val (n, ts') = find_first_numeral [] ts | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 259 | handle TERM _ => (1, ts) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 260 | in (sign*n, mk_prod (Term.fastype_of t) ts') end; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 261 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 262 | (*Find first coefficient-term THAT MATCHES u*) | 
| 13462 | 263 | fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
 | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 264 | | find_first_coeff past u (t::terms) = | 
| 13462 | 265 | let val (n,u') = dest_coeff 1 t | 
| 266 | in if u aconv u' then (n, rev past @ terms) | |
| 267 | else find_first_coeff (t::past) u terms | |
| 268 | end | |
| 269 | handle TERM _ => find_first_coeff (t::past) u terms; | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 270 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 271 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 272 | (*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 273 | val add_0s = thms "add_0s"; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 274 | val mult_1s = thms "mult_1s"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 275 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 276 | (*To perform binary arithmetic. The "left" rewriting handles patterns | 
| 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 277 | created by the simprocs, such as 3 * (5 * x). *) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 278 | val bin_simps = [numeral_0_eq_0 RS sym, numeral_1_eq_1 RS sym, | 
| 13462 | 279 | add_number_of_left, mult_number_of_left] @ | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 280 | bin_arith_simps @ bin_rel_simps; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 281 | |
| 14113 | 282 | (*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms | 
| 283 | during re-arrangement*) | |
| 284 | val non_add_bin_simps = | |
| 285 | bin_simps \\ [add_number_of_left, number_of_add RS sym]; | |
| 286 | ||
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 287 | (*To evaluate binary negations of coefficients*) | 
| 15013 | 288 | val minus_simps = [numeral_m1_eq_minus_1 RS sym, number_of_minus RS sym, | 
| 289 | bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min, | |
| 290 | bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min]; | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 291 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 292 | (*To let us treat subtraction as addition*) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 293 | val diff_simps = [diff_minus, minus_add_distrib, minus_minus]; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 294 | |
| 10713 
69c9fc1d11f2
simproc bug fix: negative literals and large terms
 paulson parents: 
10693diff
changeset | 295 | (*push the unary minus down: - x * y = x * - y *) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 296 | val minus_mult_eq_1_to_2 = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 297 | [minus_mult_left RS sym, minus_mult_right] MRS trans |> standard; | 
| 10713 
69c9fc1d11f2
simproc bug fix: negative literals and large terms
 paulson parents: 
10693diff
changeset | 298 | |
| 
69c9fc1d11f2
simproc bug fix: negative literals and large terms
 paulson parents: 
10693diff
changeset | 299 | (*to extract again any uncancelled minuses*) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 300 | val minus_from_mult_simps = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 301 | [minus_minus, minus_mult_left RS sym, minus_mult_right RS sym]; | 
| 10713 
69c9fc1d11f2
simproc bug fix: negative literals and large terms
 paulson parents: 
10693diff
changeset | 302 | |
| 
69c9fc1d11f2
simproc bug fix: negative literals and large terms
 paulson parents: 
10693diff
changeset | 303 | (*combine unary minus with numeric literals, however nested within a product*) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 304 | val mult_minus_simps = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 305 | [mult_assoc, minus_mult_left, minus_mult_eq_1_to_2]; | 
| 10713 
69c9fc1d11f2
simproc bug fix: negative literals and large terms
 paulson parents: 
10693diff
changeset | 306 | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 307 | (*Apply the given rewrite (if present) just once*) | 
| 15531 | 308 | fun trans_tac NONE = all_tac | 
| 309 | | trans_tac (SOME th) = ALLGOALS (rtac (th RS trans)); | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 310 | |
| 18328 | 311 | fun simplify_meta_eq rules = | 
| 312 | let val ss0 = HOL_basic_ss addeqcongs [eq_cong2] addsimps rules | |
| 313 | in fn ss => simplify (Simplifier.inherit_context ss ss0) o mk_meta_eq end | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 314 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 315 | structure CancelNumeralsCommon = | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 316 | struct | 
| 13462 | 317 | val mk_sum = mk_sum | 
| 318 | val dest_sum = dest_sum | |
| 319 | val mk_coeff = mk_coeff | |
| 320 | val dest_coeff = dest_coeff 1 | |
| 321 | val find_first_coeff = find_first_coeff [] | |
| 16973 | 322 | val trans_tac = fn _ => trans_tac | 
| 18328 | 323 | |
| 324 | val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ | |
| 325 | diff_simps @ minus_simps @ add_ac | |
| 326 | val norm_ss2 = num_ss addsimps non_add_bin_simps @ mult_minus_simps | |
| 327 | val norm_ss3 = num_ss addsimps minus_from_mult_simps @ add_ac @ mult_ac | |
| 16973 | 328 | fun norm_tac ss = | 
| 18328 | 329 | ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) | 
| 330 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) | |
| 331 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) | |
| 332 | ||
| 333 | val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps | |
| 334 | fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) | |
| 16973 | 335 | val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s) | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 336 | end; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 337 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 338 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 339 | structure EqCancelNumerals = CancelNumeralsFun | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 340 | (open CancelNumeralsCommon | 
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 341 | val prove_conv = Bin_Simprocs.prove_conv | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 342 | val mk_bal = HOLogic.mk_eq | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 343 | val dest_bal = HOLogic.dest_bin "op =" Term.dummyT | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 344 | val bal_add1 = eq_add_iff1 RS trans | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 345 | val bal_add2 = eq_add_iff2 RS trans | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 346 | ); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 347 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 348 | structure LessCancelNumerals = CancelNumeralsFun | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 349 | (open CancelNumeralsCommon | 
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 350 | val prove_conv = Bin_Simprocs.prove_conv | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 351 | val mk_bal = HOLogic.mk_binrel "op <" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 352 | val dest_bal = HOLogic.dest_bin "op <" Term.dummyT | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 353 | val bal_add1 = less_add_iff1 RS trans | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 354 | val bal_add2 = less_add_iff2 RS trans | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 355 | ); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 356 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 357 | structure LeCancelNumerals = CancelNumeralsFun | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 358 | (open CancelNumeralsCommon | 
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 359 | val prove_conv = Bin_Simprocs.prove_conv | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 360 | val mk_bal = HOLogic.mk_binrel "op <=" | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 361 | val dest_bal = HOLogic.dest_bin "op <=" Term.dummyT | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 362 | val bal_add1 = le_add_iff1 RS trans | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 363 | val bal_add2 = le_add_iff2 RS trans | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 364 | ); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 365 | |
| 13462 | 366 | val cancel_numerals = | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 367 | map Bin_Simprocs.prep_simproc | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 368 |    [("inteq_cancel_numerals",
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 369 | ["(l::'a::number_ring) + m = n", | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 370 | "(l::'a::number_ring) = m + n", | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 371 | "(l::'a::number_ring) - m = n", | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 372 | "(l::'a::number_ring) = m - n", | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 373 | "(l::'a::number_ring) * m = n", | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 374 | "(l::'a::number_ring) = m * n"], | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 375 | EqCancelNumerals.proc), | 
| 13462 | 376 |     ("intless_cancel_numerals",
 | 
| 14738 | 377 |      ["(l::'a::{ordered_idom,number_ring}) + m < n",
 | 
| 378 |       "(l::'a::{ordered_idom,number_ring}) < m + n",
 | |
| 379 |       "(l::'a::{ordered_idom,number_ring}) - m < n",
 | |
| 380 |       "(l::'a::{ordered_idom,number_ring}) < m - n",
 | |
| 381 |       "(l::'a::{ordered_idom,number_ring}) * m < n",
 | |
| 382 |       "(l::'a::{ordered_idom,number_ring}) < m * n"],
 | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 383 | LessCancelNumerals.proc), | 
| 13462 | 384 |     ("intle_cancel_numerals",
 | 
| 14738 | 385 |      ["(l::'a::{ordered_idom,number_ring}) + m <= n",
 | 
| 386 |       "(l::'a::{ordered_idom,number_ring}) <= m + n",
 | |
| 387 |       "(l::'a::{ordered_idom,number_ring}) - m <= n",
 | |
| 388 |       "(l::'a::{ordered_idom,number_ring}) <= m - n",
 | |
| 389 |       "(l::'a::{ordered_idom,number_ring}) * m <= n",
 | |
| 390 |       "(l::'a::{ordered_idom,number_ring}) <= m * n"],
 | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 391 | LeCancelNumerals.proc)]; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 392 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 393 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 394 | structure CombineNumeralsData = | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 395 | struct | 
| 15965 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 paulson parents: 
15921diff
changeset | 396 | val add = IntInf.+ | 
| 13462 | 397 | val mk_sum = long_mk_sum (*to work for e.g. 2*x + 3*x *) | 
| 398 | val dest_sum = dest_sum | |
| 399 | val mk_coeff = mk_coeff | |
| 400 | val dest_coeff = dest_coeff 1 | |
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 401 | val left_distrib = combine_common_factor RS trans | 
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 402 | val prove_conv = Bin_Simprocs.prove_conv_nohyps | 
| 16973 | 403 | val trans_tac = fn _ => trans_tac | 
| 18328 | 404 | |
| 405 | val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ | |
| 406 | diff_simps @ minus_simps @ add_ac | |
| 407 | val norm_ss2 = num_ss addsimps non_add_bin_simps @ mult_minus_simps | |
| 408 | val norm_ss3 = num_ss addsimps minus_from_mult_simps @ add_ac @ mult_ac | |
| 16973 | 409 | fun norm_tac ss = | 
| 18328 | 410 | ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) | 
| 411 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) | |
| 412 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) | |
| 413 | ||
| 414 | val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps | |
| 415 | fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) | |
| 16973 | 416 | val simplify_meta_eq = simplify_meta_eq (add_0s @ mult_1s) | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 417 | end; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 418 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 419 | structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); | 
| 13462 | 420 | |
| 421 | val combine_numerals = | |
| 422 | Bin_Simprocs.prep_simproc | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 423 |     ("int_combine_numerals", 
 | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 424 | ["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 425 | CombineNumerals.proc); | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 426 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 427 | end; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 428 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 429 | Addsimprocs Int_Numeral_Simprocs.cancel_numerals; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 430 | Addsimprocs [Int_Numeral_Simprocs.combine_numerals]; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 431 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 432 | (*examples: | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 433 | print_depth 22; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 434 | set timing; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 435 | set trace_simp; | 
| 13462 | 436 | fun test s = (Goal s, by (Simp_tac 1)); | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 437 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 438 | test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 439 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 440 | test "2*u = (u::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 441 | test "(i + j + 12 + (k::int)) - 15 = y"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 442 | test "(i + j + 12 + (k::int)) - 5 = y"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 443 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 444 | test "y - b < (b::int)"; | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 445 | test "y - (3*b + c) < (b::int) - 2*c"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 446 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 447 | test "(2*x - (u*v) + y) - v*3*u = (w::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 448 | test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 449 | test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 450 | test "u*v - (x*u*v + (u*v)*4 + y) = (w::int)"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 451 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 452 | test "(i + j + 12 + (k::int)) = u + 15 + y"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 453 | test "(i + j*2 + 12 + (k::int)) = j + 5 + y"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 454 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 455 | 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)"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 456 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 457 | test "a + -(b+c) + b = (d::int)"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 458 | test "a + -(b+c) - b = (d::int)"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 459 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 460 | (*negative numerals*) | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 461 | test "(i + j + -2 + (k::int)) - (u + 5 + y) = zz"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 462 | test "(i + j + -3 + (k::int)) < u + 5 + y"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 463 | test "(i + j + 3 + (k::int)) < u + -6 + y"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 464 | test "(i + j + -12 + (k::int)) - 15 = y"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 465 | test "(i + j + 12 + (k::int)) - -15 = y"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 466 | test "(i + j + -12 + (k::int)) - -15 = y"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 467 | *) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 468 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 469 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 470 | (** Constant folding for multiplication in semirings **) | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 471 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 472 | (*We do not need folding for addition: combine_numerals does the same thing*) | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 473 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 474 | structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 475 | struct | 
| 13462 | 476 | val ss = HOL_ss | 
| 477 | val eq_reflection = eq_reflection | |
| 16423 | 478 | val thy_ref = Theory.self_ref (the_context ()) | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 479 | val add_ac = mult_ac | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 480 | end; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 481 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 482 | structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 483 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 484 | val assoc_fold_simproc = | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 485 | Bin_Simprocs.prep_simproc | 
| 14738 | 486 |    ("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"],
 | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 487 | Semiring_Times_Assoc.proc); | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 488 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 489 | Addsimprocs [assoc_fold_simproc]; | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 490 | |
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 491 | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 492 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 493 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 494 | (*** decision procedure for linear arithmetic ***) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 495 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 496 | (*---------------------------------------------------------------------------*) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 497 | (* Linear arithmetic *) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 498 | (*---------------------------------------------------------------------------*) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 499 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 500 | (* | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 501 | Instantiation of the generic linear arithmetic package for int. | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 502 | *) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 503 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 504 | (* Update parameters of arithmetic prover *) | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 505 | local | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 506 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 507 | (* reduce contradictory <= to False *) | 
| 13462 | 508 | val add_rules = | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 509 | simp_thms @ bin_arith_simps @ bin_rel_simps @ arith_special @ | 
| 14390 | 510 | [neg_le_iff_le, numeral_0_eq_0, numeral_1_eq_1, | 
| 14369 | 511 | minus_zero, diff_minus, left_minus, right_minus, | 
| 512 | mult_zero_left, mult_zero_right, mult_1, mult_1_right, | |
| 513 | minus_mult_left RS sym, minus_mult_right RS sym, | |
| 514 | minus_add_distrib, minus_minus, mult_assoc, | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 515 | of_nat_0, of_nat_1, of_nat_Suc, of_nat_add, of_nat_mult, | 
| 16473 
b24c820a0b85
moving some generic inequalities from integer arith to nat arith
 paulson parents: 
16423diff
changeset | 516 | of_int_0, of_int_1, of_int_add, of_int_mult, int_eq_of_nat]; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 517 | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 518 | val simprocs = [assoc_fold_simproc, Int_Numeral_Simprocs.combine_numerals]@ | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 519 | Int_Numeral_Simprocs.cancel_numerals; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 520 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 521 | in | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 522 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 523 | val int_arith_setup = | 
| 18708 | 524 |   Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
 | 
| 14368 
2763da611ad9
converted Real/Lubs to Isar script. Converting arithmetic setup
 paulson parents: 
14331diff
changeset | 525 |    {add_mono_thms = add_mono_thms,
 | 
| 15184 | 526 | mult_mono_thms = [mult_strict_left_mono,mult_left_mono] @ mult_mono_thms, | 
| 10574 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 nipkow parents: 
9571diff
changeset | 527 | inj_thms = [zle_int RS iffD2,int_int_eq RS iffD2] @ inj_thms, | 
| 14272 
5efbb548107d
Tidying of the integer development; towards removing the
 paulson parents: 
14271diff
changeset | 528 | lessD = lessD @ [zless_imp_add1_zle], | 
| 15921 | 529 | neqE = thm "linorder_neqE_int" :: neqE, | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 530 | simpset = simpset addsimps add_rules | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 531 | addsimprocs simprocs | 
| 18708 | 532 | addcongs [if_weak_cong]}) #> | 
| 533 |   arith_inj_const ("IntDef.of_nat", HOLogic.natT --> HOLogic.intT) #>
 | |
| 534 |   arith_inj_const ("IntDef.int", HOLogic.natT --> HOLogic.intT) #>
 | |
| 535 | arith_discrete "IntDef.int"; | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 536 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 537 | end; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 538 | |
| 13462 | 539 | val fast_int_arith_simproc = | 
| 540 | Simplifier.simproc (Theory.sign_of (the_context())) | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 541 | "fast_int_arith" | 
| 14738 | 542 |      ["(m::'a::{ordered_idom,number_ring}) < n",
 | 
| 543 |       "(m::'a::{ordered_idom,number_ring}) <= n",
 | |
| 544 |       "(m::'a::{ordered_idom,number_ring}) = n"] Fast_Arith.lin_arith_prover;
 | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 545 | |
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 546 | Addsimprocs [fast_int_arith_simproc] | 
| 13462 | 547 | |
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 548 | |
| 15661 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 wenzelm parents: 
15531diff
changeset | 549 | (* Some test data | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 550 | Goal "!!a::int. [| a <= b; c <= d; x+y<z |] ==> a+c <= b+d"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 551 | by (fast_arith_tac 1); | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 552 | Goal "!!a::int. [| a < b; c < d |] ==> a-d+ 2 <= b+(-c)"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 553 | by (fast_arith_tac 1); | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11713diff
changeset | 554 | Goal "!!a::int. [| a < b; c < d |] ==> a+c+ 1 < b+d"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 555 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 556 | Goal "!!a::int. [| a <= b; b+b <= c |] ==> a+a <= c"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 557 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 558 | Goal "!!a::int. [| a+b <= i+j; a<=b; i<=j |] \ | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 559 | \ ==> a+a <= j+j"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 560 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 561 | Goal "!!a::int. [| a+b < i+j; a<b; i<j |] \ | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 562 | \ ==> a+a - - -1 < j+j - 3"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 563 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 564 | Goal "!!a::int. a+b+c <= i+j+k & a<=b & b<=c & i<=j & j<=k --> a+a+a <= k+k+k"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 565 | by (arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 566 | Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 567 | \ ==> a <= l"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 568 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 569 | Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 570 | \ ==> a+a+a+a <= l+l+l+l"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 571 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 572 | Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 573 | \ ==> a+a+a+a+a <= l+l+l+l+i"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 574 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 575 | Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 576 | \ ==> a+a+a+a+a+a <= l+l+l+l+i+l"; | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 577 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 578 | Goal "!!a::int. [| a+b+c+d <= i+j+k+l; a<=b; b<=c; c<=d; i<=j; j<=k; k<=l |] \ | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 579 | \ ==> 6*a <= 5*l+i"; | 
| 9436 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 580 | by (fast_arith_tac 1); | 
| 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 wenzelm parents: diff
changeset | 581 | *) |