| author | huffman | 
| Thu, 12 Apr 2007 03:37:30 +0200 | |
| changeset 22641 | a5dc96fad632 | 
| parent 22548 | 6ce4bddf3bcb | 
| child 22803 | 5129e02f4df2 | 
| permissions | -rw-r--r-- | 
| 10536 | 1 | (* Title: HOL/int_factor_simprocs.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 2000 University of Cambridge | |
| 5 | ||
| 14390 | 6 | Factor cancellation simprocs for the integers (and for fields). | 
| 10536 | 7 | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14113diff
changeset | 8 | This file can't be combined with int_arith1 because it requires IntDiv.thy. | 
| 10536 | 9 | *) | 
| 10 | ||
| 14390 | 11 | |
| 12 | (*To quote from Provers/Arith/cancel_numeral_factor.ML: | |
| 13 | ||
| 14 | Cancels common coefficients in balanced expressions: | |
| 15 | ||
| 16 | u*#m ~~ u'*#m' == #n*u ~~ #n'*u' | |
| 17 | ||
| 18 | where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /) | |
| 19 | and d = gcd(m,m') and n=m/d and n'=m'/d. | |
| 20 | *) | |
| 21 | ||
| 22 | val rel_number_of = [eq_number_of_eq, less_number_of_eq_neg, le_number_of_eq]; | |
| 23 | ||
| 24 | (** Factor cancellation theorems for integer division (div, not /) **) | |
| 10536 | 25 | |
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 26 | Goal "!!k::int. k~=0 ==> (k*m) div (k*n) = (m div n)"; | 
| 13462 | 27 | by (stac zdiv_zmult_zmult1 1); | 
| 28 | by Auto_tac; | |
| 10536 | 29 | qed "int_mult_div_cancel1"; | 
| 30 | ||
| 10703 | 31 | (*For ExtractCommonTermFun, cancelling common factors*) | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 32 | Goal "(k*m) div (k*n) = (if k = (0::int) then 0 else m div n)"; | 
| 13462 | 33 | by (simp_tac (simpset() addsimps [int_mult_div_cancel1]) 1); | 
| 10703 | 34 | qed "int_mult_div_cancel_disj"; | 
| 35 | ||
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14113diff
changeset | 36 | |
| 10536 | 37 | local | 
| 38 | open Int_Numeral_Simprocs | |
| 39 | in | |
| 40 | ||
| 41 | structure CancelNumeralFactorCommon = | |
| 42 | struct | |
| 13462 | 43 | val mk_coeff = mk_coeff | 
| 44 | val dest_coeff = dest_coeff 1 | |
| 16973 | 45 | val trans_tac = fn _ => trans_tac | 
| 18328 | 46 | |
| 47 | val norm_ss1 = HOL_ss addsimps minus_from_mult_simps @ mult_1s | |
| 20485 | 48 | val norm_ss2 = HOL_ss addsimps simps @ mult_minus_simps | 
| 18328 | 49 | val norm_ss3 = HOL_ss addsimps mult_ac | 
| 16973 | 50 | fun norm_tac ss = | 
| 18328 | 51 | ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) | 
| 52 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) | |
| 53 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) | |
| 54 | ||
| 20485 | 55 | val numeral_simp_ss = HOL_ss addsimps rel_number_of @ simps | 
| 18328 | 56 | fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) | 
| 22548 | 57 | val simplify_meta_eq = Int_Numeral_Simprocs.simplify_meta_eq | 
| 58 |     [@{thm add_0}, @{thm add_0_right}, @{thm mult_zero_left},
 | |
| 59 |       @{thm mult_zero_right}, @{thm mult_num1}, @{thm mult_1_right}];
 | |
| 10536 | 60 | end | 
| 61 | ||
| 14390 | 62 | (*Version for integer division*) | 
| 10536 | 63 | structure DivCancelNumeralFactor = CancelNumeralFactorFun | 
| 64 | (open CancelNumeralFactorCommon | |
| 20485 | 65 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 21416 | 66 | val mk_bal = HOLogic.mk_binop "Divides.div" | 
| 67 | val dest_bal = HOLogic.dest_bin "Divides.div" HOLogic.intT | |
| 10536 | 68 | val cancel = int_mult_div_cancel1 RS trans | 
| 69 | val neg_exchanges = false | |
| 70 | ) | |
| 71 | ||
| 14390 | 72 | (*Version for fields*) | 
| 73 | structure FieldDivCancelNumeralFactor = CancelNumeralFactorFun | |
| 74 | (open CancelNumeralFactorCommon | |
| 20485 | 75 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 14390 | 76 | val mk_bal = HOLogic.mk_binop "HOL.divide" | 
| 77 | val dest_bal = HOLogic.dest_bin "HOL.divide" Term.dummyT | |
| 22548 | 78 |   val cancel = @{thm mult_divide_cancel_left} RS trans
 | 
| 14390 | 79 | val neg_exchanges = false | 
| 80 | ) | |
| 81 | ||
| 82 | (*Version for ordered rings: mult_cancel_left requires an ordering*) | |
| 10536 | 83 | structure EqCancelNumeralFactor = CancelNumeralFactorFun | 
| 84 | (open CancelNumeralFactorCommon | |
| 20485 | 85 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 10536 | 86 | val mk_bal = HOLogic.mk_eq | 
| 14390 | 87 | val dest_bal = HOLogic.dest_bin "op =" Term.dummyT | 
| 22548 | 88 |   val cancel = @{thm mult_cancel_left} RS trans
 | 
| 10536 | 89 | val neg_exchanges = false | 
| 90 | ) | |
| 91 | ||
| 14390 | 92 | (*Version for (unordered) fields*) | 
| 93 | structure FieldEqCancelNumeralFactor = CancelNumeralFactorFun | |
| 94 | (open CancelNumeralFactorCommon | |
| 20485 | 95 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 14390 | 96 | val mk_bal = HOLogic.mk_eq | 
| 97 | val dest_bal = HOLogic.dest_bin "op =" Term.dummyT | |
| 22548 | 98 |   val cancel = @{thm field_mult_cancel_left} RS trans
 | 
| 14390 | 99 | val neg_exchanges = false | 
| 100 | ) | |
| 101 | ||
| 10536 | 102 | structure LessCancelNumeralFactor = CancelNumeralFactorFun | 
| 103 | (open CancelNumeralFactorCommon | |
| 20485 | 104 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 19277 | 105 | val mk_bal = HOLogic.mk_binrel "Orderings.less" | 
| 106 | val dest_bal = HOLogic.dest_bin "Orderings.less" Term.dummyT | |
| 22548 | 107 |   val cancel = @{thm mult_less_cancel_left} RS trans
 | 
| 10536 | 108 | val neg_exchanges = true | 
| 109 | ) | |
| 110 | ||
| 111 | structure LeCancelNumeralFactor = CancelNumeralFactorFun | |
| 112 | (open CancelNumeralFactorCommon | |
| 20485 | 113 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 19277 | 114 | val mk_bal = HOLogic.mk_binrel "Orderings.less_eq" | 
| 115 | val dest_bal = HOLogic.dest_bin "Orderings.less_eq" Term.dummyT | |
| 22548 | 116 |   val cancel = @{thm mult_le_cancel_left} RS trans
 | 
| 10536 | 117 | val neg_exchanges = true | 
| 118 | ) | |
| 119 | ||
| 14390 | 120 | val ring_cancel_numeral_factors = | 
| 20485 | 121 | map Int_Numeral_Base_Simprocs.prep_simproc | 
| 14390 | 122 |    [("ring_eq_cancel_numeral_factor",
 | 
| 14738 | 123 |      ["(l::'a::{ordered_idom,number_ring}) * m = n",
 | 
| 124 |       "(l::'a::{ordered_idom,number_ring}) = m * n"],
 | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 125 | K EqCancelNumeralFactor.proc), | 
| 14390 | 126 |     ("ring_less_cancel_numeral_factor",
 | 
| 14738 | 127 |      ["(l::'a::{ordered_idom,number_ring}) * m < n",
 | 
| 128 |       "(l::'a::{ordered_idom,number_ring}) < m * n"],
 | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 129 | K LessCancelNumeralFactor.proc), | 
| 14390 | 130 |     ("ring_le_cancel_numeral_factor",
 | 
| 14738 | 131 |      ["(l::'a::{ordered_idom,number_ring}) * m <= n",
 | 
| 132 |       "(l::'a::{ordered_idom,number_ring}) <= m * n"],
 | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 133 | K LeCancelNumeralFactor.proc), | 
| 14390 | 134 |     ("int_div_cancel_numeral_factors",
 | 
| 135 | ["((l::int) * m) div n", "(l::int) div (m * n)"], | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 136 | K DivCancelNumeralFactor.proc)]; | 
| 10536 | 137 | |
| 14390 | 138 | |
| 139 | val field_cancel_numeral_factors = | |
| 20485 | 140 | map Int_Numeral_Base_Simprocs.prep_simproc | 
| 14390 | 141 |    [("field_eq_cancel_numeral_factor",
 | 
| 142 |      ["(l::'a::{field,number_ring}) * m = n",
 | |
| 143 |       "(l::'a::{field,number_ring}) = m * n"],
 | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 144 | K FieldEqCancelNumeralFactor.proc), | 
| 14390 | 145 |     ("field_cancel_numeral_factor",
 | 
| 14426 | 146 |      ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
 | 
| 147 |       "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
 | |
| 148 |       "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
 | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 149 | K FieldDivCancelNumeralFactor.proc)] | 
| 14390 | 150 | |
| 10536 | 151 | end; | 
| 152 | ||
| 14390 | 153 | Addsimprocs ring_cancel_numeral_factors; | 
| 154 | Addsimprocs field_cancel_numeral_factors; | |
| 10536 | 155 | |
| 156 | (*examples: | |
| 157 | print_depth 22; | |
| 158 | set timing; | |
| 159 | set trace_simp; | |
| 13462 | 160 | fun test s = (Goal s; by (Simp_tac 1)); | 
| 10536 | 161 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 162 | test "9*x = 12 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 163 | test "(9*x) div (12 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 164 | test "9*x < 12 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 165 | test "9*x <= 12 * (y::int)"; | 
| 10536 | 166 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 167 | test "-99*x = 132 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 168 | test "(-99*x) div (132 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 169 | test "-99*x < 132 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 170 | test "-99*x <= 132 * (y::int)"; | 
| 10536 | 171 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 172 | test "999*x = -396 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 173 | test "(999*x) div (-396 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 174 | test "999*x < -396 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 175 | test "999*x <= -396 * (y::int)"; | 
| 10536 | 176 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 177 | test "-99*x = -81 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 178 | test "(-99*x) div (-81 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 179 | test "-99*x <= -81 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 180 | test "-99*x < -81 * (y::int)"; | 
| 10536 | 181 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 182 | test "-2 * x = -1 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 183 | test "-2 * x = -(y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 184 | test "(-2 * x) div (-1 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 185 | test "-2 * x < -(y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 186 | test "-2 * x <= -1 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 187 | test "-x < -23 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 188 | test "-x <= -23 * (y::int)"; | 
| 10536 | 189 | *) | 
| 190 | ||
| 14390 | 191 | (*And the same examples for fields such as rat or real: | 
| 192 | test "0 <= (y::rat) * -2"; | |
| 193 | test "9*x = 12 * (y::rat)"; | |
| 194 | test "(9*x) / (12 * (y::rat)) = z"; | |
| 195 | test "9*x < 12 * (y::rat)"; | |
| 196 | test "9*x <= 12 * (y::rat)"; | |
| 197 | ||
| 198 | test "-99*x = 132 * (y::rat)"; | |
| 199 | test "(-99*x) / (132 * (y::rat)) = z"; | |
| 200 | test "-99*x < 132 * (y::rat)"; | |
| 201 | test "-99*x <= 132 * (y::rat)"; | |
| 202 | ||
| 203 | test "999*x = -396 * (y::rat)"; | |
| 204 | test "(999*x) / (-396 * (y::rat)) = z"; | |
| 205 | test "999*x < -396 * (y::rat)"; | |
| 206 | test "999*x <= -396 * (y::rat)"; | |
| 207 | ||
| 208 | test "(- ((2::rat) * x) <= 2 * y)"; | |
| 209 | test "-99*x = -81 * (y::rat)"; | |
| 210 | test "(-99*x) / (-81 * (y::rat)) = z"; | |
| 211 | test "-99*x <= -81 * (y::rat)"; | |
| 212 | test "-99*x < -81 * (y::rat)"; | |
| 213 | ||
| 214 | test "-2 * x = -1 * (y::rat)"; | |
| 215 | test "-2 * x = -(y::rat)"; | |
| 216 | test "(-2 * x) / (-1 * (y::rat)) = z"; | |
| 217 | test "-2 * x < -(y::rat)"; | |
| 218 | test "-2 * x <= -1 * (y::rat)"; | |
| 219 | test "-x < -23 * (y::rat)"; | |
| 220 | test "-x <= -23 * (y::rat)"; | |
| 221 | *) | |
| 222 | ||
| 10703 | 223 | |
| 224 | (** Declarations for ExtractCommonTerm **) | |
| 225 | ||
| 226 | local | |
| 227 | open Int_Numeral_Simprocs | |
| 228 | in | |
| 229 | ||
| 230 | (*Find first term that matches u*) | |
| 18442 | 231 | fun find_first_t past u []         = raise TERM ("find_first_t", [])
 | 
| 232 | | find_first_t past u (t::terms) = | |
| 13462 | 233 | if u aconv t then (rev past @ terms) | 
| 18442 | 234 | else find_first_t (t::past) u terms | 
| 235 | handle TERM _ => find_first_t (t::past) u terms; | |
| 10703 | 236 | |
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 237 | (** Final simplification for the CancelFactor simprocs **) | 
| 22548 | 238 | val simplify_one = Int_Numeral_Simprocs.simplify_meta_eq | 
| 239 |   [@{thm mult_1_left}, mult_1_right, zdiv_1, numeral_1_eq_1];
 | |
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 240 | |
| 16973 | 241 | fun cancel_simplify_meta_eq cancel_th ss th = | 
| 242 | simplify_one ss (([th, cancel_th]) MRS trans); | |
| 10703 | 243 | |
| 14426 | 244 | (*At present, long_mk_prod creates Numeral1, so this requires the axclass | 
| 245 | number_ring*) | |
| 10703 | 246 | structure CancelFactorCommon = | 
| 247 | struct | |
| 13462 | 248 | val mk_sum = long_mk_prod | 
| 249 | val dest_sum = dest_prod | |
| 250 | val mk_coeff = mk_coeff | |
| 251 | val dest_coeff = dest_coeff | |
| 18442 | 252 | val find_first = find_first_t [] | 
| 16973 | 253 | val trans_tac = fn _ => trans_tac | 
| 18328 | 254 | val norm_ss = HOL_ss addsimps mult_1s @ mult_ac | 
| 255 | fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) | |
| 10703 | 256 | end; | 
| 257 | ||
| 14738 | 258 | (*mult_cancel_left requires an ordered comm_ring_1, such as int. The version in | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 259 | rat_arith.ML works for all fields.*) | 
| 10703 | 260 | structure EqCancelFactor = ExtractCommonTermFun | 
| 261 | (open CancelFactorCommon | |
| 20485 | 262 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 10703 | 263 | val mk_bal = HOLogic.mk_eq | 
| 264 | val dest_bal = HOLogic.dest_bin "op =" HOLogic.intT | |
| 22548 | 265 |   val simplify_meta_eq  = cancel_simplify_meta_eq @{thm mult_cancel_left}
 | 
| 10703 | 266 | ); | 
| 267 | ||
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 268 | (*int_mult_div_cancel_disj is for integer division (div). The version in | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 269 | rat_arith.ML works for all fields, using real division (/).*) | 
| 10703 | 270 | structure DivideCancelFactor = ExtractCommonTermFun | 
| 271 | (open CancelFactorCommon | |
| 20485 | 272 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 21416 | 273 | val mk_bal = HOLogic.mk_binop "Divides.div" | 
| 274 | val dest_bal = HOLogic.dest_bin "Divides.div" HOLogic.intT | |
| 10703 | 275 | val simplify_meta_eq = cancel_simplify_meta_eq int_mult_div_cancel_disj | 
| 276 | ); | |
| 277 | ||
| 13462 | 278 | val int_cancel_factor = | 
| 20485 | 279 | map Int_Numeral_Base_Simprocs.prep_simproc | 
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 280 |    [("ring_eq_cancel_factor", ["(l::int) * m = n", "(l::int) = m * n"],
 | 
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 281 | K EqCancelFactor.proc), | 
| 14390 | 282 |     ("int_divide_cancel_factor", ["((l::int) * m) div n", "(l::int) div (m*n)"],
 | 
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 283 | K DivideCancelFactor.proc)]; | 
| 10703 | 284 | |
| 14390 | 285 | (** Versions for all fields, including unordered ones (type complex).*) | 
| 286 | ||
| 287 | structure FieldEqCancelFactor = ExtractCommonTermFun | |
| 288 | (open CancelFactorCommon | |
| 20485 | 289 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 14390 | 290 | val mk_bal = HOLogic.mk_eq | 
| 291 | val dest_bal = HOLogic.dest_bin "op =" Term.dummyT | |
| 22548 | 292 |   val simplify_meta_eq  = cancel_simplify_meta_eq @{thm field_mult_cancel_left}
 | 
| 14390 | 293 | ); | 
| 294 | ||
| 295 | structure FieldDivideCancelFactor = ExtractCommonTermFun | |
| 296 | (open CancelFactorCommon | |
| 20485 | 297 | val prove_conv = Int_Numeral_Base_Simprocs.prove_conv | 
| 14390 | 298 | val mk_bal = HOLogic.mk_binop "HOL.divide" | 
| 299 | val dest_bal = HOLogic.dest_bin "HOL.divide" Term.dummyT | |
| 22548 | 300 |   val simplify_meta_eq  = cancel_simplify_meta_eq @{thm mult_divide_cancel_eq_if}
 | 
| 14390 | 301 | ); | 
| 302 | ||
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 303 | (*The number_ring class is necessary because the simprocs refer to the | 
| 14426 | 304 | binary number 1. FIXME: probably they could use 1 instead.*) | 
| 14390 | 305 | val field_cancel_factor = | 
| 20485 | 306 | map Int_Numeral_Base_Simprocs.prep_simproc | 
| 14390 | 307 |    [("field_eq_cancel_factor",
 | 
| 14426 | 308 |      ["(l::'a::{field,number_ring}) * m = n",
 | 
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 309 |       "(l::'a::{field,number_ring}) = m * n"],
 | 
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 310 | K FieldEqCancelFactor.proc), | 
| 14390 | 311 |     ("field_divide_cancel_factor",
 | 
| 14426 | 312 |      ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
 | 
| 313 |       "(l::'a::{division_by_zero,field,number_ring}) / (m * n)"],
 | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19277diff
changeset | 314 | K FieldDivideCancelFactor.proc)]; | 
| 14390 | 315 | |
| 10703 | 316 | end; | 
| 317 | ||
| 318 | Addsimprocs int_cancel_factor; | |
| 14390 | 319 | Addsimprocs field_cancel_factor; | 
| 10703 | 320 | |
| 321 | ||
| 322 | (*examples: | |
| 323 | print_depth 22; | |
| 324 | set timing; | |
| 325 | set trace_simp; | |
| 13462 | 326 | fun test s = (Goal s; by (Asm_simp_tac 1)); | 
| 10703 | 327 | |
| 328 | test "x*k = k*(y::int)"; | |
| 13462 | 329 | test "k = k*(y::int)"; | 
| 10703 | 330 | test "a*(b*c) = (b::int)"; | 
| 331 | test "a*(b*c) = d*(b::int)*(x*a)"; | |
| 332 | ||
| 333 | test "(x*k) div (k*(y::int)) = (uu::int)"; | |
| 13462 | 334 | test "(k) div (k*(y::int)) = (uu::int)"; | 
| 10703 | 335 | test "(a*(b*c)) div ((b::int)) = (uu::int)"; | 
| 336 | test "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)"; | |
| 337 | *) | |
| 338 | ||
| 14390 | 339 | (*And the same examples for fields such as rat or real: | 
| 340 | print_depth 22; | |
| 341 | set timing; | |
| 342 | set trace_simp; | |
| 343 | fun test s = (Goal s; by (Asm_simp_tac 1)); | |
| 344 | ||
| 345 | test "x*k = k*(y::rat)"; | |
| 346 | test "k = k*(y::rat)"; | |
| 347 | test "a*(b*c) = (b::rat)"; | |
| 348 | test "a*(b*c) = d*(b::rat)*(x*a)"; | |
| 349 | ||
| 350 | ||
| 351 | test "(x*k) / (k*(y::rat)) = (uu::rat)"; | |
| 352 | test "(k) / (k*(y::rat)) = (uu::rat)"; | |
| 353 | test "(a*(b*c)) / ((b::rat)) = (uu::rat)"; | |
| 354 | test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)"; | |
| 355 | ||
| 356 | (*FIXME: what do we do about this?*) | |
| 357 | test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z"; | |
| 358 | *) |