| author | Timothy Bourke | 
| Wed, 06 May 2009 10:55:47 +1000 | |
| changeset 31042 | d452117ba564 | 
| parent 30685 | dd5fe091ff04 | 
| permissions | -rw-r--r-- | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29039diff
changeset | 1 | (* Title: HOL/Tools/nat_simprocs.ML | 
| 23164 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | ||
| 4 | Simprocs for nat numerals. | |
| 5 | *) | |
| 6 | ||
| 7 | structure Nat_Numeral_Simprocs = | |
| 8 | struct | |
| 9 | ||
| 10 | (*Maps n to #n for n = 0, 1, 2*) | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 11 | val numeral_syms = [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym, @{thm numeral_2_eq_2} RS sym];
 | 
| 23164 | 12 | val numeral_sym_ss = HOL_ss addsimps numeral_syms; | 
| 13 | ||
| 14 | fun rename_numerals th = | |
| 15 | simplify numeral_sym_ss (Thm.transfer (the_context ()) th); | |
| 16 | ||
| 17 | (*Utilities*) | |
| 18 | ||
| 19 | fun mk_number n = HOLogic.number_of_const HOLogic.natT $ HOLogic.mk_numeral n; | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24431diff
changeset | 20 | fun dest_number t = Int.max (0, snd (HOLogic.dest_number t)); | 
| 23164 | 21 | |
| 22 | fun find_first_numeral past (t::terms) = | |
| 23 | ((dest_number t, t, rev past @ terms) | |
| 24 | handle TERM _ => find_first_numeral (t::past) terms) | |
| 25 |   | find_first_numeral past [] = raise TERM("find_first_numeral", []);
 | |
| 26 | ||
| 27 | val zero = mk_number 0; | |
| 28 | val mk_plus = HOLogic.mk_binop @{const_name HOL.plus};
 | |
| 29 | ||
| 30 | (*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) | |
| 31 | fun mk_sum [] = zero | |
| 32 | | mk_sum [t,u] = mk_plus (t, u) | |
| 33 | | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); | |
| 34 | ||
| 35 | (*this version ALWAYS includes a trailing zero*) | |
| 36 | fun long_mk_sum [] = HOLogic.zero | |
| 37 | | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); | |
| 38 | ||
| 39 | val dest_plus = HOLogic.dest_bin @{const_name HOL.plus} HOLogic.natT;
 | |
| 40 | ||
| 41 | ||
| 42 | (** Other simproc items **) | |
| 43 | ||
| 44 | val bin_simps = | |
| 23471 | 45 |      [@{thm nat_numeral_0_eq_0} RS sym, @{thm nat_numeral_1_eq_1} RS sym,
 | 
| 46 |       @{thm add_nat_number_of}, @{thm nat_number_of_add_left}, 
 | |
| 47 |       @{thm diff_nat_number_of}, @{thm le_number_of_eq_not_less},
 | |
| 48 |       @{thm mult_nat_number_of}, @{thm nat_number_of_mult_left}, 
 | |
| 49 |       @{thm less_nat_number_of}, 
 | |
| 50 |       @{thm Let_number_of}, @{thm nat_number_of}] @
 | |
| 29039 | 51 |      @{thms arith_simps} @ @{thms rel_simps} @ @{thms neg_simps};
 | 
| 23164 | 52 | |
| 53 | ||
| 54 | (*** CancelNumerals simprocs ***) | |
| 55 | ||
| 56 | val one = mk_number 1; | |
| 57 | val mk_times = HOLogic.mk_binop @{const_name HOL.times};
 | |
| 58 | ||
| 59 | fun mk_prod [] = one | |
| 60 | | mk_prod [t] = t | |
| 61 | | mk_prod (t :: ts) = if t = one then mk_prod ts | |
| 62 | else mk_times (t, mk_prod ts); | |
| 63 | ||
| 64 | val dest_times = HOLogic.dest_bin @{const_name HOL.times} HOLogic.natT;
 | |
| 65 | ||
| 66 | fun dest_prod t = | |
| 67 | let val (t,u) = dest_times t | |
| 68 | in dest_prod t @ dest_prod u end | |
| 69 | handle TERM _ => [t]; | |
| 70 | ||
| 71 | (*DON'T do the obvious simplifications; that would create special cases*) | |
| 72 | fun mk_coeff (k,t) = mk_times (mk_number k, t); | |
| 73 | ||
| 74 | (*Express t as a product of (possibly) a numeral with other factors, sorted*) | |
| 75 | fun dest_coeff t = | |
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29039diff
changeset | 76 | let val ts = sort TermOrd.term_ord (dest_prod t) | 
| 23164 | 77 | val (n, _, ts') = find_first_numeral [] ts | 
| 78 | handle TERM _ => (1, one, ts) | |
| 79 | in (n, mk_prod ts') end; | |
| 80 | ||
| 81 | (*Find first coefficient-term THAT MATCHES u*) | |
| 82 | fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
 | |
| 83 | | find_first_coeff past u (t::terms) = | |
| 84 | let val (n,u') = dest_coeff t | |
| 85 | in if u aconv u' then (n, rev past @ terms) | |
| 86 | else find_first_coeff (t::past) u terms | |
| 87 | end | |
| 88 | handle TERM _ => find_first_coeff (t::past) u terms; | |
| 89 | ||
| 90 | ||
| 91 | (*Split up a sum into the list of its constituent terms, on the way removing any | |
| 92 | Sucs and counting them.*) | |
| 93 | fun dest_Suc_sum (Const ("Suc", _) $ t, (k,ts)) = dest_Suc_sum (t, (k+1,ts))
 | |
| 94 | | dest_Suc_sum (t, (k,ts)) = | |
| 95 | let val (t1,t2) = dest_plus t | |
| 96 | in dest_Suc_sum (t1, dest_Suc_sum (t2, (k,ts))) end | |
| 97 | handle TERM _ => (k, t::ts); | |
| 98 | ||
| 99 | (*Code for testing whether numerals are already used in the goal*) | |
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25481diff
changeset | 100 | fun is_numeral (Const(@{const_name Int.number_of}, _) $ w) = true
 | 
| 23164 | 101 | | is_numeral _ = false; | 
| 102 | ||
| 103 | fun prod_has_numeral t = exists is_numeral (dest_prod t); | |
| 104 | ||
| 105 | (*The Sucs found in the term are converted to a binary numeral. If relaxed is false, | |
| 106 | an exception is raised unless the original expression contains at least one | |
| 107 | numeral in a coefficient position. This prevents nat_combine_numerals from | |
| 108 | introducing numerals to goals.*) | |
| 109 | fun dest_Sucs_sum relaxed t = | |
| 110 | let val (k,ts) = dest_Suc_sum (t,(0,[])) | |
| 111 | in | |
| 112 | if relaxed orelse exists prod_has_numeral ts then | |
| 113 | if k=0 then ts | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24431diff
changeset | 114 | else mk_number k :: ts | 
| 23164 | 115 |      else raise TERM("Nat_Numeral_Simprocs.dest_Sucs_sum", [t])
 | 
| 116 | end; | |
| 117 | ||
| 118 | ||
| 119 | (*Simplify 1*n and n*1 to n*) | |
| 23881 | 120 | val add_0s  = map rename_numerals [@{thm add_0}, @{thm add_0_right}];
 | 
| 23164 | 121 | val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}];
 | 
| 122 | ||
| 123 | (*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*) | |
| 124 | ||
| 125 | (*And these help the simproc return False when appropriate, which helps | |
| 126 | the arith prover.*) | |
| 23881 | 127 | val contra_rules = [@{thm add_Suc}, @{thm add_Suc_right}, @{thm Zero_not_Suc},
 | 
| 128 |   @{thm Suc_not_Zero}, @{thm le_0_eq}];
 | |
| 23164 | 129 | |
| 130 | val simplify_meta_eq = | |
| 30518 | 131 | Arith_Data.simplify_meta_eq | 
| 23471 | 132 |         ([@{thm nat_numeral_0_eq_0}, @{thm numeral_1_eq_Suc_0}, @{thm add_0}, @{thm add_0_right},
 | 
| 133 |           @{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules);
 | |
| 23164 | 134 | |
| 135 | ||
| 136 | (*** Applying CancelNumeralsFun ***) | |
| 137 | ||
| 138 | structure CancelNumeralsCommon = | |
| 139 | struct | |
| 140 | val mk_sum = (fn T:typ => mk_sum) | |
| 141 | val dest_sum = dest_Sucs_sum true | |
| 142 | val mk_coeff = mk_coeff | |
| 143 | val dest_coeff = dest_coeff | |
| 144 | val find_first_coeff = find_first_coeff [] | |
| 30518 | 145 | val trans_tac = K Arith_Data.trans_tac | 
| 23164 | 146 | |
| 30518 | 147 | val norm_ss1 = Int_Numeral_Simprocs.num_ss addsimps numeral_syms @ add_0s @ mult_1s @ | 
| 23881 | 148 |     [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac}
 | 
| 30518 | 149 |   val norm_ss2 = Int_Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac}
 | 
| 23164 | 150 | fun norm_tac ss = | 
| 151 | ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) | |
| 152 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) | |
| 153 | ||
| 154 | val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps; | |
| 155 | fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)); | |
| 156 | val simplify_meta_eq = simplify_meta_eq | |
| 157 | end; | |
| 158 | ||
| 159 | ||
| 160 | structure EqCancelNumerals = CancelNumeralsFun | |
| 161 | (open CancelNumeralsCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 162 | val prove_conv = Arith_Data.prove_conv | 
| 23164 | 163 | val mk_bal = HOLogic.mk_eq | 
| 164 | val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT | |
| 23471 | 165 |   val bal_add1 = @{thm nat_eq_add_iff1} RS trans
 | 
| 166 |   val bal_add2 = @{thm nat_eq_add_iff2} RS trans
 | |
| 23164 | 167 | ); | 
| 168 | ||
| 169 | structure LessCancelNumerals = CancelNumeralsFun | |
| 170 | (open CancelNumeralsCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 171 | val prove_conv = Arith_Data.prove_conv | 
| 23881 | 172 |   val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
 | 
| 173 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
 | |
| 23471 | 174 |   val bal_add1 = @{thm nat_less_add_iff1} RS trans
 | 
| 175 |   val bal_add2 = @{thm nat_less_add_iff2} RS trans
 | |
| 23164 | 176 | ); | 
| 177 | ||
| 178 | structure LeCancelNumerals = CancelNumeralsFun | |
| 179 | (open CancelNumeralsCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 180 | val prove_conv = Arith_Data.prove_conv | 
| 23881 | 181 |   val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
 | 
| 182 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
 | |
| 23471 | 183 |   val bal_add1 = @{thm nat_le_add_iff1} RS trans
 | 
| 184 |   val bal_add2 = @{thm nat_le_add_iff2} RS trans
 | |
| 23164 | 185 | ); | 
| 186 | ||
| 187 | structure DiffCancelNumerals = CancelNumeralsFun | |
| 188 | (open CancelNumeralsCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 189 | val prove_conv = Arith_Data.prove_conv | 
| 23164 | 190 |   val mk_bal   = HOLogic.mk_binop @{const_name HOL.minus}
 | 
| 191 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.minus} HOLogic.natT
 | |
| 23471 | 192 |   val bal_add1 = @{thm nat_diff_add_eq1} RS trans
 | 
| 193 |   val bal_add2 = @{thm nat_diff_add_eq2} RS trans
 | |
| 23164 | 194 | ); | 
| 195 | ||
| 196 | ||
| 197 | val cancel_numerals = | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 198 | map Arith_Data.prep_simproc | 
| 23164 | 199 |    [("nateq_cancel_numerals",
 | 
| 200 | ["(l::nat) + m = n", "(l::nat) = m + n", | |
| 201 | "(l::nat) * m = n", "(l::nat) = m * n", | |
| 202 | "Suc m = n", "m = Suc n"], | |
| 203 | K EqCancelNumerals.proc), | |
| 204 |     ("natless_cancel_numerals",
 | |
| 205 | ["(l::nat) + m < n", "(l::nat) < m + n", | |
| 206 | "(l::nat) * m < n", "(l::nat) < m * n", | |
| 207 | "Suc m < n", "m < Suc n"], | |
| 208 | K LessCancelNumerals.proc), | |
| 209 |     ("natle_cancel_numerals",
 | |
| 210 | ["(l::nat) + m <= n", "(l::nat) <= m + n", | |
| 211 | "(l::nat) * m <= n", "(l::nat) <= m * n", | |
| 212 | "Suc m <= n", "m <= Suc n"], | |
| 213 | K LeCancelNumerals.proc), | |
| 214 |     ("natdiff_cancel_numerals",
 | |
| 215 | ["((l::nat) + m) - n", "(l::nat) - (m + n)", | |
| 216 | "(l::nat) * m - n", "(l::nat) - m * n", | |
| 217 | "Suc m - n", "m - Suc n"], | |
| 218 | K DiffCancelNumerals.proc)]; | |
| 219 | ||
| 220 | ||
| 221 | (*** Applying CombineNumeralsFun ***) | |
| 222 | ||
| 223 | structure CombineNumeralsData = | |
| 224 | struct | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24431diff
changeset | 225 | type coeff = int | 
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24431diff
changeset | 226 | val iszero = (fn x => x = 0) | 
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24431diff
changeset | 227 | val add = op + | 
| 23164 | 228 | val mk_sum = (fn T:typ => long_mk_sum) (*to work for 2*x + 3*x *) | 
| 229 | val dest_sum = dest_Sucs_sum false | |
| 230 | val mk_coeff = mk_coeff | |
| 231 | val dest_coeff = dest_coeff | |
| 23471 | 232 |   val left_distrib      = @{thm left_add_mult_distrib} RS trans
 | 
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 233 | val prove_conv = Arith_Data.prove_conv_nohyps | 
| 30518 | 234 | val trans_tac = K Arith_Data.trans_tac | 
| 23164 | 235 | |
| 30518 | 236 |   val norm_ss1 = Int_Numeral_Simprocs.num_ss addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1}] @ @{thms add_ac}
 | 
| 237 |   val norm_ss2 = Int_Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac}
 | |
| 23164 | 238 | fun norm_tac ss = | 
| 239 | ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) | |
| 240 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) | |
| 241 | ||
| 242 | val numeral_simp_ss = HOL_ss addsimps add_0s @ bin_simps; | |
| 243 | fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) | |
| 244 | val simplify_meta_eq = simplify_meta_eq | |
| 245 | end; | |
| 246 | ||
| 247 | structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); | |
| 248 | ||
| 249 | val combine_numerals = | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 250 |   Arith_Data.prep_simproc ("nat_combine_numerals", ["(i::nat) + j", "Suc (i + j)"], K CombineNumerals.proc);
 | 
| 23164 | 251 | |
| 252 | ||
| 253 | (*** Applying CancelNumeralFactorFun ***) | |
| 254 | ||
| 255 | structure CancelNumeralFactorCommon = | |
| 256 | struct | |
| 257 | val mk_coeff = mk_coeff | |
| 258 | val dest_coeff = dest_coeff | |
| 30518 | 259 | val trans_tac = K Arith_Data.trans_tac | 
| 23164 | 260 | |
| 30518 | 261 | val norm_ss1 = Int_Numeral_Simprocs.num_ss addsimps | 
| 23881 | 262 |     numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_add_numeral_1_left}] @ @{thms add_ac}
 | 
| 30518 | 263 |   val norm_ss2 = Int_Numeral_Simprocs.num_ss addsimps bin_simps @ @{thms add_ac} @ @{thms mult_ac}
 | 
| 23164 | 264 | fun norm_tac ss = | 
| 265 | ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) | |
| 266 | THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) | |
| 267 | ||
| 268 | val numeral_simp_ss = HOL_ss addsimps bin_simps | |
| 269 | fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) | |
| 270 | val simplify_meta_eq = simplify_meta_eq | |
| 271 | end | |
| 272 | ||
| 273 | structure DivCancelNumeralFactor = CancelNumeralFactorFun | |
| 274 | (open CancelNumeralFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 275 | val prove_conv = Arith_Data.prove_conv | 
| 23164 | 276 |   val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
 | 
| 277 |   val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT
 | |
| 23471 | 278 |   val cancel = @{thm nat_mult_div_cancel1} RS trans
 | 
| 23164 | 279 | val neg_exchanges = false | 
| 280 | ) | |
| 281 | ||
| 23969 | 282 | structure DvdCancelNumeralFactor = CancelNumeralFactorFun | 
| 283 | (open CancelNumeralFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 284 | val prove_conv = Arith_Data.prove_conv | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
25919diff
changeset | 285 |   val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
 | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
25919diff
changeset | 286 |   val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
 | 
| 23969 | 287 |   val cancel = @{thm nat_mult_dvd_cancel1} RS trans
 | 
| 288 | val neg_exchanges = false | |
| 289 | ) | |
| 290 | ||
| 23164 | 291 | structure EqCancelNumeralFactor = CancelNumeralFactorFun | 
| 292 | (open CancelNumeralFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 293 | val prove_conv = Arith_Data.prove_conv | 
| 23164 | 294 | val mk_bal = HOLogic.mk_eq | 
| 295 | val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT | |
| 23471 | 296 |   val cancel = @{thm nat_mult_eq_cancel1} RS trans
 | 
| 23164 | 297 | val neg_exchanges = false | 
| 298 | ) | |
| 299 | ||
| 300 | structure LessCancelNumeralFactor = CancelNumeralFactorFun | |
| 301 | (open CancelNumeralFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 302 | val prove_conv = Arith_Data.prove_conv | 
| 23881 | 303 |   val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
 | 
| 304 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
 | |
| 23471 | 305 |   val cancel = @{thm nat_mult_less_cancel1} RS trans
 | 
| 23164 | 306 | val neg_exchanges = true | 
| 307 | ) | |
| 308 | ||
| 309 | structure LeCancelNumeralFactor = CancelNumeralFactorFun | |
| 310 | (open CancelNumeralFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 311 | val prove_conv = Arith_Data.prove_conv | 
| 23881 | 312 |   val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
 | 
| 313 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
 | |
| 23471 | 314 |   val cancel = @{thm nat_mult_le_cancel1} RS trans
 | 
| 23164 | 315 | val neg_exchanges = true | 
| 316 | ) | |
| 317 | ||
| 318 | val cancel_numeral_factors = | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 319 | map Arith_Data.prep_simproc | 
| 23164 | 320 |    [("nateq_cancel_numeral_factors",
 | 
| 321 | ["(l::nat) * m = n", "(l::nat) = m * n"], | |
| 322 | K EqCancelNumeralFactor.proc), | |
| 323 |     ("natless_cancel_numeral_factors",
 | |
| 324 | ["(l::nat) * m < n", "(l::nat) < m * n"], | |
| 325 | K LessCancelNumeralFactor.proc), | |
| 326 |     ("natle_cancel_numeral_factors",
 | |
| 327 | ["(l::nat) * m <= n", "(l::nat) <= m * n"], | |
| 328 | K LeCancelNumeralFactor.proc), | |
| 329 |     ("natdiv_cancel_numeral_factors",
 | |
| 330 | ["((l::nat) * m) div n", "(l::nat) div (m * n)"], | |
| 23969 | 331 | K DivCancelNumeralFactor.proc), | 
| 332 |     ("natdvd_cancel_numeral_factors",
 | |
| 333 | ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"], | |
| 334 | K DvdCancelNumeralFactor.proc)]; | |
| 23164 | 335 | |
| 336 | ||
| 337 | ||
| 338 | (*** Applying ExtractCommonTermFun ***) | |
| 339 | ||
| 340 | (*this version ALWAYS includes a trailing one*) | |
| 341 | fun long_mk_prod [] = one | |
| 342 | | long_mk_prod (t :: ts) = mk_times (t, mk_prod ts); | |
| 343 | ||
| 344 | (*Find first term that matches u*) | |
| 345 | fun find_first_t past u []         = raise TERM("find_first_t", [])
 | |
| 346 | | find_first_t past u (t::terms) = | |
| 347 | if u aconv t then (rev past @ terms) | |
| 348 | else find_first_t (t::past) u terms | |
| 349 | handle TERM _ => find_first_t (t::past) u terms; | |
| 350 | ||
| 351 | (** Final simplification for the CancelFactor simprocs **) | |
| 30518 | 352 | val simplify_one = Arith_Data.simplify_meta_eq | 
| 23164 | 353 |   [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_1}, @{thm numeral_1_eq_Suc_0}];
 | 
| 354 | ||
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 355 | fun cancel_simplify_meta_eq ss cancel_th th = | 
| 23164 | 356 | simplify_one ss (([th, cancel_th]) MRS trans); | 
| 357 | ||
| 358 | structure CancelFactorCommon = | |
| 359 | struct | |
| 360 | val mk_sum = (fn T:typ => long_mk_prod) | |
| 361 | val dest_sum = dest_prod | |
| 362 | val mk_coeff = mk_coeff | |
| 363 | val dest_coeff = dest_coeff | |
| 364 | val find_first = find_first_t [] | |
| 30518 | 365 | val trans_tac = K Arith_Data.trans_tac | 
| 23881 | 366 |   val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac}
 | 
| 23164 | 367 | fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 368 | val simplify_meta_eq = cancel_simplify_meta_eq | 
| 23164 | 369 | end; | 
| 370 | ||
| 371 | structure EqCancelFactor = ExtractCommonTermFun | |
| 372 | (open CancelFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 373 | val prove_conv = Arith_Data.prove_conv | 
| 23164 | 374 | val mk_bal = HOLogic.mk_eq | 
| 375 | val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT | |
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 376 |   val simp_conv = K(K (SOME @{thm nat_mult_eq_cancel_disj}))
 | 
| 23164 | 377 | ); | 
| 378 | ||
| 379 | structure LessCancelFactor = ExtractCommonTermFun | |
| 380 | (open CancelFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 381 | val prove_conv = Arith_Data.prove_conv | 
| 23881 | 382 |   val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less}
 | 
| 383 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.less} HOLogic.natT
 | |
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 384 |   val simp_conv = K(K (SOME @{thm nat_mult_less_cancel_disj}))
 | 
| 23164 | 385 | ); | 
| 386 | ||
| 387 | structure LeCancelFactor = ExtractCommonTermFun | |
| 388 | (open CancelFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 389 | val prove_conv = Arith_Data.prove_conv | 
| 23881 | 390 |   val mk_bal   = HOLogic.mk_binrel @{const_name HOL.less_eq}
 | 
| 391 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.less_eq} HOLogic.natT
 | |
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 392 |   val simp_conv = K(K (SOME @{thm nat_mult_le_cancel_disj}))
 | 
| 23164 | 393 | ); | 
| 394 | ||
| 395 | structure DivideCancelFactor = ExtractCommonTermFun | |
| 396 | (open CancelFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 397 | val prove_conv = Arith_Data.prove_conv | 
| 23164 | 398 |   val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
 | 
| 399 |   val dest_bal = HOLogic.dest_bin @{const_name Divides.div} HOLogic.natT
 | |
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 400 |   val simp_conv = K(K (SOME @{thm nat_mult_div_cancel_disj}))
 | 
| 23164 | 401 | ); | 
| 402 | ||
| 23969 | 403 | structure DvdCancelFactor = ExtractCommonTermFun | 
| 404 | (open CancelFactorCommon | |
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 405 | val prove_conv = Arith_Data.prove_conv | 
| 27651 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
25919diff
changeset | 406 |   val mk_bal   = HOLogic.mk_binrel @{const_name Ring_and_Field.dvd}
 | 
| 
16a26996c30e
moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
 haftmann parents: 
25919diff
changeset | 407 |   val dest_bal = HOLogic.dest_bin @{const_name Ring_and_Field.dvd} HOLogic.natT
 | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 408 |   val simp_conv = K(K (SOME @{thm nat_mult_dvd_cancel_disj}))
 | 
| 23969 | 409 | ); | 
| 410 | ||
| 23164 | 411 | val cancel_factor = | 
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 412 | map Arith_Data.prep_simproc | 
| 23164 | 413 |    [("nat_eq_cancel_factor",
 | 
| 414 | ["(l::nat) * m = n", "(l::nat) = m * n"], | |
| 415 | K EqCancelFactor.proc), | |
| 416 |     ("nat_less_cancel_factor",
 | |
| 417 | ["(l::nat) * m < n", "(l::nat) < m * n"], | |
| 418 | K LessCancelFactor.proc), | |
| 419 |     ("nat_le_cancel_factor",
 | |
| 420 | ["(l::nat) * m <= n", "(l::nat) <= m * n"], | |
| 421 | K LeCancelFactor.proc), | |
| 422 |     ("nat_divide_cancel_factor",
 | |
| 423 | ["((l::nat) * m) div n", "(l::nat) div (m * n)"], | |
| 23969 | 424 | K DivideCancelFactor.proc), | 
| 425 |     ("nat_dvd_cancel_factor",
 | |
| 426 | ["((l::nat) * m) dvd n", "(l::nat) dvd (m * n)"], | |
| 427 | K DvdCancelFactor.proc)]; | |
| 23164 | 428 | |
| 429 | end; | |
| 430 | ||
| 431 | ||
| 432 | Addsimprocs Nat_Numeral_Simprocs.cancel_numerals; | |
| 433 | Addsimprocs [Nat_Numeral_Simprocs.combine_numerals]; | |
| 434 | Addsimprocs Nat_Numeral_Simprocs.cancel_numeral_factors; | |
| 435 | Addsimprocs Nat_Numeral_Simprocs.cancel_factor; | |
| 436 | ||
| 437 | ||
| 438 | (*examples: | |
| 439 | print_depth 22; | |
| 440 | set timing; | |
| 441 | set trace_simp; | |
| 442 | fun test s = (Goal s; by (Simp_tac 1)); | |
| 443 | ||
| 444 | (*cancel_numerals*) | |
| 445 | test "l +( 2) + (2) + 2 + (l + 2) + (oo + 2) = (uu::nat)"; | |
| 446 | test "(2*length xs < 2*length xs + j)"; | |
| 447 | test "(2*length xs < length xs * 2 + j)"; | |
| 448 | test "2*u = (u::nat)"; | |
| 449 | test "2*u = Suc (u)"; | |
| 450 | test "(i + j + 12 + (k::nat)) - 15 = y"; | |
| 451 | test "(i + j + 12 + (k::nat)) - 5 = y"; | |
| 452 | test "Suc u - 2 = y"; | |
| 453 | test "Suc (Suc (Suc u)) - 2 = y"; | |
| 454 | test "(i + j + 2 + (k::nat)) - 1 = y"; | |
| 455 | test "(i + j + 1 + (k::nat)) - 2 = y"; | |
| 456 | ||
| 457 | test "(2*x + (u*v) + y) - v*3*u = (w::nat)"; | |
| 458 | test "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = (w::nat)"; | |
| 459 | test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::nat)"; | |
| 460 | test "Suc (Suc (2*x*u*v + u*4 + y)) - u = w"; | |
| 461 | test "Suc ((u*v)*4) - v*3*u = w"; | |
| 462 | test "Suc (Suc ((u*v)*3)) - v*3*u = w"; | |
| 463 | ||
| 464 | test "(i + j + 12 + (k::nat)) = u + 15 + y"; | |
| 465 | test "(i + j + 32 + (k::nat)) - (u + 15 + y) = zz"; | |
| 466 | test "(i + j + 12 + (k::nat)) = u + 5 + y"; | |
| 467 | (*Suc*) | |
| 468 | test "(i + j + 12 + k) = Suc (u + y)"; | |
| 469 | test "Suc (Suc (Suc (Suc (Suc (u + y))))) <= ((i + j) + 41 + k)"; | |
| 470 | test "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))"; | |
| 471 | test "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v"; | |
| 472 | test "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))"; | |
| 473 | test "2*y + 3*z + 2*u = Suc (u)"; | |
| 474 | test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)"; | |
| 475 | 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::nat)"; | |
| 476 | test "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)"; | |
| 477 | test "(2*n*m) < (3*(m*n)) + (u::nat)"; | |
| 478 | ||
| 479 | test "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) <= Suc 0)"; | |
| 480 | ||
| 481 | test "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) <= length l1"; | |
| 482 | ||
| 483 | test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) <= length (compT P E A ST mxr e))"; | |
| 484 | ||
| 485 | test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un \<A> e) ST mxr c))))))) <= length (compT P E A ST mxr e))"; | |
| 486 | ||
| 487 | ||
| 488 | (*negative numerals: FAIL*) | |
| 489 | test "(i + j + -23 + (k::nat)) < u + 15 + y"; | |
| 490 | test "(i + j + 3 + (k::nat)) < u + -15 + y"; | |
| 491 | test "(i + j + -12 + (k::nat)) - 15 = y"; | |
| 492 | test "(i + j + 12 + (k::nat)) - -15 = y"; | |
| 493 | test "(i + j + -12 + (k::nat)) - -15 = y"; | |
| 494 | ||
| 495 | (*combine_numerals*) | |
| 496 | test "k + 3*k = (u::nat)"; | |
| 497 | test "Suc (i + 3) = u"; | |
| 498 | test "Suc (i + j + 3 + k) = u"; | |
| 499 | test "k + j + 3*k + j = (u::nat)"; | |
| 500 | test "Suc (j*i + i + k + 5 + 3*k + i*j*4) = (u::nat)"; | |
| 501 | test "(2*n*m) + (3*(m*n)) = (u::nat)"; | |
| 502 | (*negative numerals: FAIL*) | |
| 503 | test "Suc (i + j + -3 + k) = u"; | |
| 504 | ||
| 505 | (*cancel_numeral_factors*) | |
| 506 | test "9*x = 12 * (y::nat)"; | |
| 507 | test "(9*x) div (12 * (y::nat)) = z"; | |
| 508 | test "9*x < 12 * (y::nat)"; | |
| 509 | test "9*x <= 12 * (y::nat)"; | |
| 510 | ||
| 511 | (*cancel_factor*) | |
| 512 | test "x*k = k*(y::nat)"; | |
| 513 | test "k = k*(y::nat)"; | |
| 514 | test "a*(b*c) = (b::nat)"; | |
| 515 | test "a*(b*c) = d*(b::nat)*(x*a)"; | |
| 516 | ||
| 517 | test "x*k < k*(y::nat)"; | |
| 518 | test "k < k*(y::nat)"; | |
| 519 | test "a*(b*c) < (b::nat)"; | |
| 520 | test "a*(b*c) < d*(b::nat)*(x*a)"; | |
| 521 | ||
| 522 | test "x*k <= k*(y::nat)"; | |
| 523 | test "k <= k*(y::nat)"; | |
| 524 | test "a*(b*c) <= (b::nat)"; | |
| 525 | test "a*(b*c) <= d*(b::nat)*(x*a)"; | |
| 526 | ||
| 527 | test "(x*k) div (k*(y::nat)) = (uu::nat)"; | |
| 528 | test "(k) div (k*(y::nat)) = (uu::nat)"; | |
| 529 | test "(a*(b*c)) div ((b::nat)) = (uu::nat)"; | |
| 530 | test "(a*(b*c)) div (d*(b::nat)*(x*a)) = (uu::nat)"; | |
| 531 | *) | |
| 532 | ||
| 533 | ||
| 534 | (*** Prepare linear arithmetic for nat numerals ***) | |
| 535 | ||
| 536 | local | |
| 537 | ||
| 538 | (* reduce contradictory <= to False *) | |
| 24431 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 539 | val add_rules = @{thms ring_distribs} @
 | 
| 23471 | 540 |   [@{thm Let_number_of}, @{thm Let_0}, @{thm Let_1}, @{thm nat_0}, @{thm nat_1},
 | 
| 541 |    @{thm add_nat_number_of}, @{thm diff_nat_number_of}, @{thm mult_nat_number_of},
 | |
| 542 |    @{thm eq_nat_number_of}, @{thm less_nat_number_of}, @{thm le_number_of_eq_not_less},
 | |
| 543 |    @{thm le_Suc_number_of}, @{thm le_number_of_Suc},
 | |
| 544 |    @{thm less_Suc_number_of}, @{thm less_number_of_Suc},
 | |
| 545 |    @{thm Suc_eq_number_of}, @{thm eq_number_of_Suc},
 | |
| 546 |    @{thm mult_Suc}, @{thm mult_Suc_right},
 | |
| 24431 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 547 |    @{thm add_Suc}, @{thm add_Suc_right},
 | 
| 23471 | 548 |    @{thm eq_number_of_0}, @{thm eq_0_number_of}, @{thm less_0_number_of},
 | 
| 549 |    @{thm of_int_number_of_eq}, @{thm of_nat_number_of_eq}, @{thm nat_number_of}, @{thm if_True}, @{thm if_False}];
 | |
| 23164 | 550 | |
| 24431 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 551 | (* Products are multiplied out during proof (re)construction via | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 552 | ring_distribs. Ideally they should remain atomic. But that is | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 553 | currently not possible because 1 is replaced by Suc 0, and then some | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 554 | simprocs start to mess around with products like (n+1)*m. The rule | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 555 | 1 == Suc 0 is necessary for early parts of HOL where numerals and | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 556 | simprocs are not yet available. But then it is difficult to remove | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 557 | that rule later on, because it may find its way back in when theories | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 558 | (and thus lin-arith simpsets) are merged. Otherwise one could turn the | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 559 | rule around (Suc n = n+1) and see if that helps products being left | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 560 | alone. *) | 
| 
02d29baa42ff
tuned linear arith (once again) with ring_distribs
 nipkow parents: 
24093diff
changeset | 561 | |
| 23164 | 562 | val simprocs = Nat_Numeral_Simprocs.combine_numerals | 
| 563 | :: Nat_Numeral_Simprocs.cancel_numerals; | |
| 564 | ||
| 565 | in | |
| 566 | ||
| 567 | val nat_simprocs_setup = | |
| 30685 | 568 |   Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
 | 
| 23164 | 569 |    {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
 | 
| 570 | inj_thms = inj_thms, lessD = lessD, neqE = neqE, | |
| 571 | simpset = simpset addsimps add_rules | |
| 572 | addsimprocs simprocs}); | |
| 573 | ||
| 574 | end; |