| author | wenzelm | 
| Fri, 15 Jul 2005 15:44:18 +0200 | |
| changeset 16863 | 79b9a6481ae4 | 
| parent 15271 | 3c32a26510c4 | 
| child 16973 | b2a894562b8f | 
| 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 | |
| 10536 | 45 | val trans_tac = trans_tac | 
| 13462 | 46 | val norm_tac = | 
| 14390 | 47 | ALLGOALS (simp_tac (HOL_ss addsimps minus_from_mult_simps @ mult_1s)) | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 48 | THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@mult_minus_simps)) | 
| 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 49 | THEN ALLGOALS (simp_tac (HOL_ss addsimps mult_ac)) | 
| 14390 | 50 | val numeral_simp_tac = | 
| 51 | ALLGOALS (simp_tac (HOL_ss addsimps rel_number_of@bin_simps)) | |
| 52 | val simplify_meta_eq = | |
| 53 | Int_Numeral_Simprocs.simplify_meta_eq | |
| 54 | [add_0, add_0_right, | |
| 55 | mult_zero_left, mult_zero_right, mult_1, mult_1_right]; | |
| 10536 | 56 | end | 
| 57 | ||
| 14390 | 58 | (*Version for integer division*) | 
| 10536 | 59 | structure DivCancelNumeralFactor = CancelNumeralFactorFun | 
| 60 | (open CancelNumeralFactorCommon | |
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 61 | val prove_conv = Bin_Simprocs.prove_conv | 
| 10536 | 62 | val mk_bal = HOLogic.mk_binop "Divides.op div" | 
| 63 | val dest_bal = HOLogic.dest_bin "Divides.op div" HOLogic.intT | |
| 64 | val cancel = int_mult_div_cancel1 RS trans | |
| 65 | val neg_exchanges = false | |
| 66 | ) | |
| 67 | ||
| 14390 | 68 | (*Version for fields*) | 
| 69 | structure FieldDivCancelNumeralFactor = CancelNumeralFactorFun | |
| 70 | (open CancelNumeralFactorCommon | |
| 71 | val prove_conv = Bin_Simprocs.prove_conv | |
| 72 | val mk_bal = HOLogic.mk_binop "HOL.divide" | |
| 73 | val dest_bal = HOLogic.dest_bin "HOL.divide" Term.dummyT | |
| 74 | val cancel = mult_divide_cancel_left RS trans | |
| 75 | val neg_exchanges = false | |
| 76 | ) | |
| 77 | ||
| 78 | (*Version for ordered rings: mult_cancel_left requires an ordering*) | |
| 10536 | 79 | structure EqCancelNumeralFactor = CancelNumeralFactorFun | 
| 80 | (open CancelNumeralFactorCommon | |
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 81 | val prove_conv = Bin_Simprocs.prove_conv | 
| 10536 | 82 | val mk_bal = HOLogic.mk_eq | 
| 14390 | 83 | val dest_bal = HOLogic.dest_bin "op =" Term.dummyT | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14113diff
changeset | 84 | val cancel = mult_cancel_left RS trans | 
| 10536 | 85 | val neg_exchanges = false | 
| 86 | ) | |
| 87 | ||
| 14390 | 88 | (*Version for (unordered) fields*) | 
| 89 | structure FieldEqCancelNumeralFactor = CancelNumeralFactorFun | |
| 90 | (open CancelNumeralFactorCommon | |
| 91 | val prove_conv = Bin_Simprocs.prove_conv | |
| 92 | val mk_bal = HOLogic.mk_eq | |
| 93 | val dest_bal = HOLogic.dest_bin "op =" Term.dummyT | |
| 94 | val cancel = field_mult_cancel_left RS trans | |
| 95 | val neg_exchanges = false | |
| 96 | ) | |
| 97 | ||
| 10536 | 98 | structure LessCancelNumeralFactor = CancelNumeralFactorFun | 
| 99 | (open CancelNumeralFactorCommon | |
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 100 | val prove_conv = Bin_Simprocs.prove_conv | 
| 10536 | 101 | val mk_bal = HOLogic.mk_binrel "op <" | 
| 14390 | 102 | val dest_bal = HOLogic.dest_bin "op <" Term.dummyT | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14113diff
changeset | 103 | val cancel = mult_less_cancel_left RS trans | 
| 10536 | 104 | val neg_exchanges = true | 
| 105 | ) | |
| 106 | ||
| 107 | structure LeCancelNumeralFactor = CancelNumeralFactorFun | |
| 108 | (open CancelNumeralFactorCommon | |
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 109 | val prove_conv = Bin_Simprocs.prove_conv | 
| 10536 | 110 | val mk_bal = HOLogic.mk_binrel "op <=" | 
| 14390 | 111 | val dest_bal = HOLogic.dest_bin "op <=" Term.dummyT | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14113diff
changeset | 112 | val cancel = mult_le_cancel_left RS trans | 
| 10536 | 113 | val neg_exchanges = true | 
| 114 | ) | |
| 115 | ||
| 14390 | 116 | val ring_cancel_numeral_factors = | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 117 | map Bin_Simprocs.prep_simproc | 
| 14390 | 118 |    [("ring_eq_cancel_numeral_factor",
 | 
| 14738 | 119 |      ["(l::'a::{ordered_idom,number_ring}) * m = n",
 | 
| 120 |       "(l::'a::{ordered_idom,number_ring}) = m * n"],
 | |
| 10536 | 121 | EqCancelNumeralFactor.proc), | 
| 14390 | 122 |     ("ring_less_cancel_numeral_factor",
 | 
| 14738 | 123 |      ["(l::'a::{ordered_idom,number_ring}) * m < n",
 | 
| 124 |       "(l::'a::{ordered_idom,number_ring}) < m * n"],
 | |
| 10536 | 125 | LessCancelNumeralFactor.proc), | 
| 14390 | 126 |     ("ring_le_cancel_numeral_factor",
 | 
| 14738 | 127 |      ["(l::'a::{ordered_idom,number_ring}) * m <= n",
 | 
| 128 |       "(l::'a::{ordered_idom,number_ring}) <= m * n"],
 | |
| 10536 | 129 | LeCancelNumeralFactor.proc), | 
| 14390 | 130 |     ("int_div_cancel_numeral_factors",
 | 
| 131 | ["((l::int) * m) div n", "(l::int) div (m * n)"], | |
| 10536 | 132 | DivCancelNumeralFactor.proc)]; | 
| 133 | ||
| 14390 | 134 | |
| 135 | val field_cancel_numeral_factors = | |
| 136 | map Bin_Simprocs.prep_simproc | |
| 137 |    [("field_eq_cancel_numeral_factor",
 | |
| 138 |      ["(l::'a::{field,number_ring}) * m = n",
 | |
| 139 |       "(l::'a::{field,number_ring}) = m * n"],
 | |
| 140 | FieldEqCancelNumeralFactor.proc), | |
| 141 |     ("field_cancel_numeral_factor",
 | |
| 14426 | 142 |      ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
 | 
| 143 |       "(l::'a::{division_by_zero,field,number_ring}) / (m * n)",
 | |
| 144 |       "((number_of v)::'a::{division_by_zero,field,number_ring}) / (number_of w)"],
 | |
| 14390 | 145 | FieldDivCancelNumeralFactor.proc)] | 
| 146 | ||
| 10536 | 147 | end; | 
| 148 | ||
| 14390 | 149 | Addsimprocs ring_cancel_numeral_factors; | 
| 150 | Addsimprocs field_cancel_numeral_factors; | |
| 10536 | 151 | |
| 152 | (*examples: | |
| 153 | print_depth 22; | |
| 154 | set timing; | |
| 155 | set trace_simp; | |
| 13462 | 156 | fun test s = (Goal s; by (Simp_tac 1)); | 
| 10536 | 157 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 158 | test "9*x = 12 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 159 | test "(9*x) div (12 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 160 | test "9*x < 12 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 161 | test "9*x <= 12 * (y::int)"; | 
| 10536 | 162 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 163 | test "-99*x = 132 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 164 | test "(-99*x) div (132 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 165 | test "-99*x < 132 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 166 | test "-99*x <= 132 * (y::int)"; | 
| 10536 | 167 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 168 | test "999*x = -396 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 169 | test "(999*x) div (-396 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 170 | test "999*x < -396 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 171 | test "999*x <= -396 * (y::int)"; | 
| 10536 | 172 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 173 | test "-99*x = -81 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 174 | test "(-99*x) div (-81 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 175 | test "-99*x <= -81 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 176 | test "-99*x < -81 * (y::int)"; | 
| 10536 | 177 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 178 | test "-2 * x = -1 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 179 | test "-2 * x = -(y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 180 | test "(-2 * x) div (-1 * (y::int)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 181 | test "-2 * x < -(y::int)"; | 
| 
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 "-x < -23 * (y::int)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 184 | test "-x <= -23 * (y::int)"; | 
| 10536 | 185 | *) | 
| 186 | ||
| 14390 | 187 | (*And the same examples for fields such as rat or real: | 
| 188 | test "0 <= (y::rat) * -2"; | |
| 189 | test "9*x = 12 * (y::rat)"; | |
| 190 | test "(9*x) / (12 * (y::rat)) = z"; | |
| 191 | test "9*x < 12 * (y::rat)"; | |
| 192 | test "9*x <= 12 * (y::rat)"; | |
| 193 | ||
| 194 | test "-99*x = 132 * (y::rat)"; | |
| 195 | test "(-99*x) / (132 * (y::rat)) = z"; | |
| 196 | test "-99*x < 132 * (y::rat)"; | |
| 197 | test "-99*x <= 132 * (y::rat)"; | |
| 198 | ||
| 199 | test "999*x = -396 * (y::rat)"; | |
| 200 | test "(999*x) / (-396 * (y::rat)) = z"; | |
| 201 | test "999*x < -396 * (y::rat)"; | |
| 202 | test "999*x <= -396 * (y::rat)"; | |
| 203 | ||
| 204 | test "(- ((2::rat) * x) <= 2 * y)"; | |
| 205 | test "-99*x = -81 * (y::rat)"; | |
| 206 | test "(-99*x) / (-81 * (y::rat)) = z"; | |
| 207 | test "-99*x <= -81 * (y::rat)"; | |
| 208 | test "-99*x < -81 * (y::rat)"; | |
| 209 | ||
| 210 | test "-2 * x = -1 * (y::rat)"; | |
| 211 | test "-2 * x = -(y::rat)"; | |
| 212 | test "(-2 * x) / (-1 * (y::rat)) = z"; | |
| 213 | test "-2 * x < -(y::rat)"; | |
| 214 | test "-2 * x <= -1 * (y::rat)"; | |
| 215 | test "-x < -23 * (y::rat)"; | |
| 216 | test "-x <= -23 * (y::rat)"; | |
| 217 | *) | |
| 218 | ||
| 10703 | 219 | |
| 220 | (** Declarations for ExtractCommonTerm **) | |
| 221 | ||
| 222 | local | |
| 223 | open Int_Numeral_Simprocs | |
| 224 | in | |
| 225 | ||
| 226 | (*Find first term that matches u*) | |
| 13462 | 227 | fun find_first past u []         = raise TERM("find_first", [])
 | 
| 10703 | 228 | | find_first past u (t::terms) = | 
| 13462 | 229 | if u aconv t then (rev past @ terms) | 
| 10703 | 230 | else find_first (t::past) u terms | 
| 13462 | 231 | handle TERM _ => find_first (t::past) u terms; | 
| 10703 | 232 | |
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 233 | (** Final simplification for the CancelFactor simprocs **) | 
| 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 234 | val simplify_one = | 
| 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 235 | Int_Numeral_Simprocs.simplify_meta_eq | 
| 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 236 | [mult_1_left, mult_1_right, zdiv_1, numeral_1_eq_1]; | 
| 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 237 | |
| 13462 | 238 | fun cancel_simplify_meta_eq cancel_th th = | 
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 239 | simplify_one (([th, cancel_th]) MRS trans); | 
| 10703 | 240 | |
| 14426 | 241 | (*At present, long_mk_prod creates Numeral1, so this requires the axclass | 
| 242 | number_ring*) | |
| 10703 | 243 | structure CancelFactorCommon = | 
| 244 | struct | |
| 13462 | 245 | val mk_sum = long_mk_prod | 
| 246 | val dest_sum = dest_prod | |
| 247 | val mk_coeff = mk_coeff | |
| 248 | val dest_coeff = dest_coeff | |
| 249 | val find_first = find_first [] | |
| 10703 | 250 | val trans_tac = trans_tac | 
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 251 | val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@mult_ac)) | 
| 10703 | 252 | end; | 
| 253 | ||
| 14738 | 254 | (*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 | 255 | rat_arith.ML works for all fields.*) | 
| 10703 | 256 | structure EqCancelFactor = ExtractCommonTermFun | 
| 257 | (open CancelFactorCommon | |
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 258 | val prove_conv = Bin_Simprocs.prove_conv | 
| 10703 | 259 | val mk_bal = HOLogic.mk_eq | 
| 260 | val dest_bal = HOLogic.dest_bin "op =" HOLogic.intT | |
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14113diff
changeset | 261 | val simplify_meta_eq = cancel_simplify_meta_eq mult_cancel_left | 
| 10703 | 262 | ); | 
| 263 | ||
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
14378diff
changeset | 264 | (*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 | 265 | rat_arith.ML works for all fields, using real division (/).*) | 
| 10703 | 266 | structure DivideCancelFactor = ExtractCommonTermFun | 
| 267 | (open CancelFactorCommon | |
| 13485 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 wenzelm parents: 
13462diff
changeset | 268 | val prove_conv = Bin_Simprocs.prove_conv | 
| 10703 | 269 | val mk_bal = HOLogic.mk_binop "Divides.op div" | 
| 270 | val dest_bal = HOLogic.dest_bin "Divides.op div" HOLogic.intT | |
| 271 | val simplify_meta_eq = cancel_simplify_meta_eq int_mult_div_cancel_disj | |
| 272 | ); | |
| 273 | ||
| 13462 | 274 | val int_cancel_factor = | 
| 11868 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 paulson parents: 
11704diff
changeset | 275 | map Bin_Simprocs.prep_simproc | 
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 276 |    [("ring_eq_cancel_factor", ["(l::int) * m = n", "(l::int) = m * n"],
 | 
| 14390 | 277 | EqCancelFactor.proc), | 
| 278 |     ("int_divide_cancel_factor", ["((l::int) * m) div n", "(l::int) div (m*n)"],
 | |
| 10703 | 279 | DivideCancelFactor.proc)]; | 
| 280 | ||
| 14390 | 281 | (** Versions for all fields, including unordered ones (type complex).*) | 
| 282 | ||
| 283 | structure FieldEqCancelFactor = ExtractCommonTermFun | |
| 284 | (open CancelFactorCommon | |
| 285 | val prove_conv = Bin_Simprocs.prove_conv | |
| 286 | val mk_bal = HOLogic.mk_eq | |
| 287 | val dest_bal = HOLogic.dest_bin "op =" Term.dummyT | |
| 288 | val simplify_meta_eq = cancel_simplify_meta_eq field_mult_cancel_left | |
| 289 | ); | |
| 290 | ||
| 291 | structure FieldDivideCancelFactor = ExtractCommonTermFun | |
| 292 | (open CancelFactorCommon | |
| 293 | val prove_conv = Bin_Simprocs.prove_conv | |
| 294 | val mk_bal = HOLogic.mk_binop "HOL.divide" | |
| 295 | val dest_bal = HOLogic.dest_bin "HOL.divide" Term.dummyT | |
| 296 | val simplify_meta_eq = cancel_simplify_meta_eq mult_divide_cancel_eq_if | |
| 297 | ); | |
| 298 | ||
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 299 | (*The number_ring class is necessary because the simprocs refer to the | 
| 14426 | 300 | binary number 1. FIXME: probably they could use 1 instead.*) | 
| 14390 | 301 | val field_cancel_factor = | 
| 302 | map Bin_Simprocs.prep_simproc | |
| 303 |    [("field_eq_cancel_factor",
 | |
| 14426 | 304 |      ["(l::'a::{field,number_ring}) * m = n",
 | 
| 15271 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 paulson parents: 
14738diff
changeset | 305 |       "(l::'a::{field,number_ring}) = m * n"],
 | 
| 14390 | 306 | FieldEqCancelFactor.proc), | 
| 307 |     ("field_divide_cancel_factor",
 | |
| 14426 | 308 |      ["((l::'a::{division_by_zero,field,number_ring}) * m) / n",
 | 
| 309 |       "(l::'a::{division_by_zero,field,number_ring}) / (m * n)"],
 | |
| 14390 | 310 | FieldDivideCancelFactor.proc)]; | 
| 311 | ||
| 10703 | 312 | end; | 
| 313 | ||
| 314 | Addsimprocs int_cancel_factor; | |
| 14390 | 315 | Addsimprocs field_cancel_factor; | 
| 10703 | 316 | |
| 317 | ||
| 318 | (*examples: | |
| 319 | print_depth 22; | |
| 320 | set timing; | |
| 321 | set trace_simp; | |
| 13462 | 322 | fun test s = (Goal s; by (Asm_simp_tac 1)); | 
| 10703 | 323 | |
| 324 | test "x*k = k*(y::int)"; | |
| 13462 | 325 | test "k = k*(y::int)"; | 
| 10703 | 326 | test "a*(b*c) = (b::int)"; | 
| 327 | test "a*(b*c) = d*(b::int)*(x*a)"; | |
| 328 | ||
| 329 | test "(x*k) div (k*(y::int)) = (uu::int)"; | |
| 13462 | 330 | test "(k) div (k*(y::int)) = (uu::int)"; | 
| 10703 | 331 | test "(a*(b*c)) div ((b::int)) = (uu::int)"; | 
| 332 | test "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)"; | |
| 333 | *) | |
| 334 | ||
| 14390 | 335 | (*And the same examples for fields such as rat or real: | 
| 336 | print_depth 22; | |
| 337 | set timing; | |
| 338 | set trace_simp; | |
| 339 | fun test s = (Goal s; by (Asm_simp_tac 1)); | |
| 340 | ||
| 341 | test "x*k = k*(y::rat)"; | |
| 342 | test "k = k*(y::rat)"; | |
| 343 | test "a*(b*c) = (b::rat)"; | |
| 344 | test "a*(b*c) = d*(b::rat)*(x*a)"; | |
| 345 | ||
| 346 | ||
| 347 | test "(x*k) / (k*(y::rat)) = (uu::rat)"; | |
| 348 | test "(k) / (k*(y::rat)) = (uu::rat)"; | |
| 349 | test "(a*(b*c)) / ((b::rat)) = (uu::rat)"; | |
| 350 | test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)"; | |
| 351 | ||
| 352 | (*FIXME: what do we do about this?*) | |
| 353 | test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z"; | |
| 354 | *) |