| author | wenzelm | 
| Sun, 06 Mar 2016 13:19:19 +0100 | |
| changeset 62528 | c8c532b22947 | 
| parent 61694 | 6571c78c9667 | 
| child 64244 | e7102c40783c | 
| permissions | -rw-r--r-- | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
30685diff
changeset | 1 | (* Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 23164 | 2 | |
| 3 | Simprocs for nat numerals. | |
| 4 | *) | |
| 5 | ||
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
30685diff
changeset | 6 | signature NAT_NUMERAL_SIMPROCS = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
30685diff
changeset | 7 | sig | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 8 | val combine_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 9 | val eq_cancel_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 10 | val less_cancel_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 11 | val le_cancel_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 12 | val diff_cancel_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 13 | val eq_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 14 | val less_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 15 | val le_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 16 | val div_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 17 | val dvd_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 18 | val eq_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 19 | val less_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 20 | val le_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 21 | val div_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 22 | val dvd_cancel_factor: Proof.context -> cterm -> thm option | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
30685diff
changeset | 23 | end; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
30685diff
changeset | 24 | |
| 45462 
aba629d6cee5
use simproc_setup for more nat_numeral simprocs; add simproc tests
 huffman parents: 
45436diff
changeset | 25 | structure Nat_Numeral_Simprocs : NAT_NUMERAL_SIMPROCS = | 
| 23164 | 26 | struct | 
| 27 | ||
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 28 | (*Maps n to #n for n = 1, 2*) | 
| 61694 
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
 paulson <lp15@cam.ac.uk> parents: 
61144diff
changeset | 29 | val numeral_syms = [@{thm numeral_One} RS sym, @{thm numeral_2_eq_2} RS sym];
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 30 | val numeral_sym_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 31 |   simpset_of (put_simpset HOL_basic_ss @{context} addsimps numeral_syms);
 | 
| 23164 | 32 | |
| 33 | (*Utilities*) | |
| 34 | ||
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 35 | fun mk_number 1 = HOLogic.numeral_const HOLogic.natT $ HOLogic.one_const | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 36 | | mk_number n = HOLogic.mk_number HOLogic.natT n; | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24431diff
changeset | 37 | fun dest_number t = Int.max (0, snd (HOLogic.dest_number t)); | 
| 23164 | 38 | |
| 39 | fun find_first_numeral past (t::terms) = | |
| 40 | ((dest_number t, t, rev past @ terms) | |
| 41 | handle TERM _ => find_first_numeral (t::past) terms) | |
| 42 |   | find_first_numeral past [] = raise TERM("find_first_numeral", []);
 | |
| 43 | ||
| 44 | val zero = mk_number 0; | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 45 | val mk_plus = HOLogic.mk_binop @{const_name Groups.plus};
 | 
| 23164 | 46 | |
| 47 | (*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) | |
| 48 | fun mk_sum [] = zero | |
| 49 | | mk_sum [t,u] = mk_plus (t, u) | |
| 50 | | mk_sum (t :: ts) = mk_plus (t, mk_sum ts); | |
| 51 | ||
| 52 | (*this version ALWAYS includes a trailing zero*) | |
| 53 | fun long_mk_sum [] = HOLogic.zero | |
| 54 | | long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); | |
| 55 | ||
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 56 | val dest_plus = HOLogic.dest_bin @{const_name Groups.plus} HOLogic.natT;
 | 
| 23164 | 57 | |
| 58 | ||
| 59 | (** Other simproc items **) | |
| 60 | ||
| 61 | val bin_simps = | |
| 61694 
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
 paulson <lp15@cam.ac.uk> parents: 
61144diff
changeset | 62 |      [@{thm numeral_One} RS sym,
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 63 |       @{thm numeral_plus_numeral}, @{thm add_numeral_left},
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 64 |       @{thm diff_nat_numeral}, @{thm diff_0_eq_0}, @{thm diff_0},
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 65 |       @{thm numeral_times_numeral}, @{thm mult_numeral_left(1)},
 | 
| 45668 
0ea1c705eebb
use HOL_basic_ss instead of HOL_ss for internal simproc proofs (cf. b36eedcd1633)
 huffman parents: 
45463diff
changeset | 66 |       @{thm if_True}, @{thm if_False}, @{thm not_False_eq_True},
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 67 |       @{thm nat_0}, @{thm nat_numeral}, @{thm nat_neg_numeral}] @
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 68 |      @{thms arith_simps} @ @{thms rel_simps};
 | 
| 23164 | 69 | |
| 70 | ||
| 71 | (*** CancelNumerals simprocs ***) | |
| 72 | ||
| 73 | val one = mk_number 1; | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 74 | val mk_times = HOLogic.mk_binop @{const_name Groups.times};
 | 
| 23164 | 75 | |
| 76 | fun mk_prod [] = one | |
| 77 | | mk_prod [t] = t | |
| 78 | | mk_prod (t :: ts) = if t = one then mk_prod ts | |
| 79 | else mk_times (t, mk_prod ts); | |
| 80 | ||
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 81 | val dest_times = HOLogic.dest_bin @{const_name Groups.times} HOLogic.natT;
 | 
| 23164 | 82 | |
| 83 | fun dest_prod t = | |
| 84 | let val (t,u) = dest_times t | |
| 85 | in dest_prod t @ dest_prod u end | |
| 86 | handle TERM _ => [t]; | |
| 87 | ||
| 88 | (*DON'T do the obvious simplifications; that would create special cases*) | |
| 89 | fun mk_coeff (k,t) = mk_times (mk_number k, t); | |
| 90 | ||
| 91 | (*Express t as a product of (possibly) a numeral with other factors, sorted*) | |
| 92 | fun dest_coeff t = | |
| 35408 | 93 | let val ts = sort Term_Ord.term_ord (dest_prod t) | 
| 23164 | 94 | val (n, _, ts') = find_first_numeral [] ts | 
| 95 | handle TERM _ => (1, one, ts) | |
| 96 | in (n, mk_prod ts') end; | |
| 97 | ||
| 98 | (*Find first coefficient-term THAT MATCHES u*) | |
| 99 | fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
 | |
| 100 | | find_first_coeff past u (t::terms) = | |
| 101 | let val (n,u') = dest_coeff t | |
| 102 | in if u aconv u' then (n, rev past @ terms) | |
| 103 | else find_first_coeff (t::past) u terms | |
| 104 | end | |
| 105 | handle TERM _ => find_first_coeff (t::past) u terms; | |
| 106 | ||
| 107 | ||
| 108 | (*Split up a sum into the list of its constituent terms, on the way removing any | |
| 109 | Sucs and counting them.*) | |
| 37388 | 110 | fun dest_Suc_sum (Const (@{const_name Suc}, _) $ t, (k,ts)) = dest_Suc_sum (t, (k+1,ts))
 | 
| 23164 | 111 | | dest_Suc_sum (t, (k,ts)) = | 
| 112 | let val (t1,t2) = dest_plus t | |
| 113 | in dest_Suc_sum (t1, dest_Suc_sum (t2, (k,ts))) end | |
| 114 | handle TERM _ => (k, t::ts); | |
| 115 | ||
| 116 | (*Code for testing whether numerals are already used in the goal*) | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 117 | fun is_numeral (Const(@{const_name Num.numeral}, _) $ w) = true
 | 
| 23164 | 118 | | is_numeral _ = false; | 
| 119 | ||
| 120 | fun prod_has_numeral t = exists is_numeral (dest_prod t); | |
| 121 | ||
| 122 | (*The Sucs found in the term are converted to a binary numeral. If relaxed is false, | |
| 123 | an exception is raised unless the original expression contains at least one | |
| 124 | numeral in a coefficient position. This prevents nat_combine_numerals from | |
| 125 | introducing numerals to goals.*) | |
| 126 | fun dest_Sucs_sum relaxed t = | |
| 127 | let val (k,ts) = dest_Suc_sum (t,(0,[])) | |
| 128 | in | |
| 129 | if relaxed orelse exists prod_has_numeral ts then | |
| 130 | if k=0 then ts | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
24431diff
changeset | 131 | else mk_number k :: ts | 
| 23164 | 132 |      else raise TERM("Nat_Numeral_Simprocs.dest_Sucs_sum", [t])
 | 
| 133 | end; | |
| 134 | ||
| 135 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 136 | (* FIXME !? *) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 137 | val rename_numerals = simplify (put_simpset numeral_sym_ss @{context}) o Thm.transfer @{theory};
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 138 | |
| 23164 | 139 | (*Simplify 1*n and n*1 to n*) | 
| 35064 
1bdef0c013d3
hide fact names clashing with fact names from Group.thy
 haftmann parents: 
35050diff
changeset | 140 | val add_0s  = map rename_numerals [@{thm Nat.add_0}, @{thm Nat.add_0_right}];
 | 
| 23164 | 141 | val mult_1s = map rename_numerals [@{thm nat_mult_1}, @{thm nat_mult_1_right}];
 | 
| 142 | ||
| 143 | (*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*) | |
| 144 | ||
| 145 | (*And these help the simproc return False when appropriate, which helps | |
| 146 | the arith prover.*) | |
| 23881 | 147 | val contra_rules = [@{thm add_Suc}, @{thm add_Suc_right}, @{thm Zero_not_Suc},
 | 
| 148 |   @{thm Suc_not_Zero}, @{thm le_0_eq}];
 | |
| 23164 | 149 | |
| 150 | val simplify_meta_eq = | |
| 30518 | 151 | Arith_Data.simplify_meta_eq | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
45668diff
changeset | 152 |         ([@{thm numeral_1_eq_Suc_0}, @{thm Nat.add_0}, @{thm Nat.add_0_right},
 | 
| 23471 | 153 |           @{thm mult_0}, @{thm mult_0_right}, @{thm mult_1}, @{thm mult_1_right}] @ contra_rules);
 | 
| 23164 | 154 | |
| 155 | ||
| 156 | (*** Applying CancelNumeralsFun ***) | |
| 157 | ||
| 158 | structure CancelNumeralsCommon = | |
| 44945 | 159 | struct | 
| 160 | val mk_sum = (fn T : typ => mk_sum) | |
| 161 | val dest_sum = dest_Sucs_sum true | |
| 162 | val mk_coeff = mk_coeff | |
| 163 | val dest_coeff = dest_coeff | |
| 164 | val find_first_coeff = find_first_coeff [] | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 165 | val trans_tac = Numeral_Simprocs.trans_tac | 
| 23164 | 166 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 167 | val norm_ss1 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 168 |     simpset_of (put_simpset Numeral_Simprocs.num_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 169 | addsimps numeral_syms @ add_0s @ mult_1s @ | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
51717diff
changeset | 170 |         [@{thm Suc_eq_plus1_left}] @ @{thms ac_simps})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 171 | val norm_ss2 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 172 |     simpset_of (put_simpset Numeral_Simprocs.num_ss @{context}
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
51717diff
changeset | 173 |       addsimps bin_simps @ @{thms ac_simps} @ @{thms ac_simps})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 174 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 175 | ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 176 | THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) | 
| 23164 | 177 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 178 | val numeral_simp_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 179 |     simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 180 | addsimps add_0s @ bin_simps); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 181 | fun numeral_simp_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 182 | ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)); | 
| 23164 | 183 | val simplify_meta_eq = simplify_meta_eq | 
| 44945 | 184 | val prove_conv = Arith_Data.prove_conv | 
| 185 | end; | |
| 23164 | 186 | |
| 187 | structure EqCancelNumerals = CancelNumeralsFun | |
| 188 | (open CancelNumeralsCommon | |
| 189 | val mk_bal = HOLogic.mk_eq | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
37388diff
changeset | 190 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} HOLogic.natT
 | 
| 23471 | 191 |   val bal_add1 = @{thm nat_eq_add_iff1} RS trans
 | 
| 192 |   val bal_add2 = @{thm nat_eq_add_iff2} RS trans
 | |
| 23164 | 193 | ); | 
| 194 | ||
| 195 | structure LessCancelNumerals = CancelNumeralsFun | |
| 196 | (open CancelNumeralsCommon | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 197 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 198 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
 | 
| 23471 | 199 |   val bal_add1 = @{thm nat_less_add_iff1} RS trans
 | 
| 200 |   val bal_add2 = @{thm nat_less_add_iff2} RS trans
 | |
| 23164 | 201 | ); | 
| 202 | ||
| 203 | structure LeCancelNumerals = CancelNumeralsFun | |
| 204 | (open CancelNumeralsCommon | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 205 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 206 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
 | 
| 23471 | 207 |   val bal_add1 = @{thm nat_le_add_iff1} RS trans
 | 
| 208 |   val bal_add2 = @{thm nat_le_add_iff2} RS trans
 | |
| 23164 | 209 | ); | 
| 210 | ||
| 211 | structure DiffCancelNumerals = CancelNumeralsFun | |
| 212 | (open CancelNumeralsCommon | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 213 |   val mk_bal   = HOLogic.mk_binop @{const_name Groups.minus}
 | 
| 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 214 |   val dest_bal = HOLogic.dest_bin @{const_name Groups.minus} HOLogic.natT
 | 
| 23471 | 215 |   val bal_add1 = @{thm nat_diff_add_eq1} RS trans
 | 
| 216 |   val bal_add2 = @{thm nat_diff_add_eq2} RS trans
 | |
| 23164 | 217 | ); | 
| 218 | ||
| 61144 | 219 | val eq_cancel_numerals = EqCancelNumerals.proc | 
| 220 | val less_cancel_numerals = LessCancelNumerals.proc | |
| 221 | val le_cancel_numerals = LeCancelNumerals.proc | |
| 222 | val diff_cancel_numerals = DiffCancelNumerals.proc | |
| 23164 | 223 | |
| 224 | ||
| 225 | (*** Applying CombineNumeralsFun ***) | |
| 226 | ||
| 227 | structure CombineNumeralsData = | |
| 44945 | 228 | struct | 
| 229 | type coeff = int | |
| 230 | val iszero = (fn x => x = 0) | |
| 231 | val add = op + | |
| 232 | val mk_sum = (fn T : typ => long_mk_sum) (*to work for 2*x + 3*x *) | |
| 233 | val dest_sum = dest_Sucs_sum false | |
| 234 | val mk_coeff = mk_coeff | |
| 235 | val dest_coeff = dest_coeff | |
| 236 |   val left_distrib = @{thm left_add_mult_distrib} RS trans
 | |
| 237 | val prove_conv = Arith_Data.prove_conv_nohyps | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 238 | val trans_tac = Numeral_Simprocs.trans_tac | 
| 23164 | 239 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 240 | val norm_ss1 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 241 |     simpset_of (put_simpset Numeral_Simprocs.num_ss @{context}
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
51717diff
changeset | 242 |       addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_plus1}] @ @{thms ac_simps})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 243 | val norm_ss2 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 244 |     simpset_of (put_simpset Numeral_Simprocs.num_ss @{context}
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
51717diff
changeset | 245 |       addsimps bin_simps @ @{thms ac_simps} @ @{thms ac_simps})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 246 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 247 | ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 248 | THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) | 
| 23164 | 249 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 250 | val numeral_simp_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 251 |     simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 252 | addsimps add_0s @ bin_simps); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 253 | fun numeral_simp_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 254 | ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) | 
| 44945 | 255 | val simplify_meta_eq = simplify_meta_eq | 
| 256 | end; | |
| 23164 | 257 | |
| 258 | structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); | |
| 259 | ||
| 61144 | 260 | val combine_numerals = CombineNumerals.proc | 
| 23164 | 261 | |
| 262 | ||
| 263 | (*** Applying CancelNumeralFactorFun ***) | |
| 264 | ||
| 265 | structure CancelNumeralFactorCommon = | |
| 44945 | 266 | struct | 
| 267 | val mk_coeff = mk_coeff | |
| 268 | val dest_coeff = dest_coeff | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 269 | val trans_tac = Numeral_Simprocs.trans_tac | 
| 23164 | 270 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 271 | val norm_ss1 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 272 |     simpset_of (put_simpset Numeral_Simprocs.num_ss @{context}
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
51717diff
changeset | 273 |       addsimps numeral_syms @ add_0s @ mult_1s @ [@{thm Suc_eq_plus1_left}] @ @{thms ac_simps})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 274 | val norm_ss2 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 275 |     simpset_of (put_simpset Numeral_Simprocs.num_ss @{context}
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
51717diff
changeset | 276 |       addsimps bin_simps @ @{thms ac_simps} @ @{thms ac_simps})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 277 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 278 | ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 279 | THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) | 
| 23164 | 280 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 281 | val numeral_simp_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 282 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps bin_simps)
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 283 | fun numeral_simp_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 284 | ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) | 
| 44945 | 285 | val simplify_meta_eq = simplify_meta_eq | 
| 286 | val prove_conv = Arith_Data.prove_conv | |
| 287 | end; | |
| 23164 | 288 | |
| 289 | structure DivCancelNumeralFactor = CancelNumeralFactorFun | |
| 290 | (open CancelNumeralFactorCommon | |
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59582diff
changeset | 291 |   val mk_bal   = HOLogic.mk_binop @{const_name Rings.divide}
 | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59582diff
changeset | 292 |   val dest_bal = HOLogic.dest_bin @{const_name Rings.divide} HOLogic.natT
 | 
| 23471 | 293 |   val cancel = @{thm nat_mult_div_cancel1} RS trans
 | 
| 23164 | 294 | val neg_exchanges = false | 
| 44945 | 295 | ); | 
| 23164 | 296 | |
| 23969 | 297 | structure DvdCancelNumeralFactor = CancelNumeralFactorFun | 
| 298 | (open CancelNumeralFactorCommon | |
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34974diff
changeset | 299 |   val mk_bal   = HOLogic.mk_binrel @{const_name Rings.dvd}
 | 
| 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34974diff
changeset | 300 |   val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} HOLogic.natT
 | 
| 23969 | 301 |   val cancel = @{thm nat_mult_dvd_cancel1} RS trans
 | 
| 302 | val neg_exchanges = false | |
| 44945 | 303 | ); | 
| 23969 | 304 | |
| 23164 | 305 | structure EqCancelNumeralFactor = CancelNumeralFactorFun | 
| 306 | (open CancelNumeralFactorCommon | |
| 307 | val mk_bal = HOLogic.mk_eq | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
37388diff
changeset | 308 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} HOLogic.natT
 | 
| 23471 | 309 |   val cancel = @{thm nat_mult_eq_cancel1} RS trans
 | 
| 23164 | 310 | val neg_exchanges = false | 
| 44945 | 311 | ); | 
| 23164 | 312 | |
| 313 | structure LessCancelNumeralFactor = CancelNumeralFactorFun | |
| 61144 | 314 | ( | 
| 315 | open CancelNumeralFactorCommon | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 316 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 317 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
 | 
| 23471 | 318 |   val cancel = @{thm nat_mult_less_cancel1} RS trans
 | 
| 23164 | 319 | val neg_exchanges = true | 
| 44945 | 320 | ); | 
| 23164 | 321 | |
| 322 | structure LeCancelNumeralFactor = CancelNumeralFactorFun | |
| 61144 | 323 | ( | 
| 324 | open CancelNumeralFactorCommon | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 325 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 326 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
 | 
| 23471 | 327 |   val cancel = @{thm nat_mult_le_cancel1} RS trans
 | 
| 23164 | 328 | val neg_exchanges = true | 
| 329 | ) | |
| 330 | ||
| 61144 | 331 | val eq_cancel_numeral_factor = EqCancelNumeralFactor.proc | 
| 332 | val less_cancel_numeral_factor = LessCancelNumeralFactor.proc | |
| 333 | val le_cancel_numeral_factor = LeCancelNumeralFactor.proc | |
| 334 | val div_cancel_numeral_factor = DivCancelNumeralFactor.proc | |
| 335 | val dvd_cancel_numeral_factor = DvdCancelNumeralFactor.proc | |
| 23164 | 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 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 355 | fun cancel_simplify_meta_eq ctxt cancel_th th = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 356 | simplify_one ctxt (([th, cancel_th]) MRS trans); | 
| 23164 | 357 | |
| 358 | structure CancelFactorCommon = | |
| 44945 | 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 [] | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 365 | val trans_tac = Numeral_Simprocs.trans_tac | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 366 | val norm_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 367 |     simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
51717diff
changeset | 368 |       addsimps mult_1s @ @{thms ac_simps})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 369 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47108diff
changeset | 370 | ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 371 | val simplify_meta_eq = cancel_simplify_meta_eq | 
| 45270 
d5b5c9259afd
fix bug in cancel_factor simprocs so they will work on goals like 'x * y < x * z' where the common term is already on the left
 huffman parents: 
44947diff
changeset | 372 | fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) | 
| 44945 | 373 | end; | 
| 23164 | 374 | |
| 375 | structure EqCancelFactor = ExtractCommonTermFun | |
| 376 | (open CancelFactorCommon | |
| 377 | val mk_bal = HOLogic.mk_eq | |
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
37388diff
changeset | 378 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} HOLogic.natT
 | 
| 31368 | 379 |   fun simp_conv _ _ = SOME @{thm nat_mult_eq_cancel_disj}
 | 
| 23164 | 380 | ); | 
| 381 | ||
| 44945 | 382 | structure LeCancelFactor = ExtractCommonTermFun | 
| 383 | (open CancelFactorCommon | |
| 384 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
 | |
| 385 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} HOLogic.natT
 | |
| 386 |   fun simp_conv _ _ = SOME @{thm nat_mult_le_cancel_disj}
 | |
| 387 | ); | |
| 388 | ||
| 23164 | 389 | structure LessCancelFactor = ExtractCommonTermFun | 
| 390 | (open CancelFactorCommon | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 391 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
 | 
| 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35064diff
changeset | 392 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} HOLogic.natT
 | 
| 31368 | 393 |   fun simp_conv _ _ = SOME @{thm nat_mult_less_cancel_disj}
 | 
| 23164 | 394 | ); | 
| 395 | ||
| 396 | structure DivideCancelFactor = ExtractCommonTermFun | |
| 397 | (open CancelFactorCommon | |
| 60352 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59582diff
changeset | 398 |   val mk_bal   = HOLogic.mk_binop @{const_name Rings.divide}
 | 
| 
d46de31a50c4
separate class for division operator, with particular syntax added in more specific classes
 haftmann parents: 
59582diff
changeset | 399 |   val dest_bal = HOLogic.dest_bin @{const_name Rings.divide} HOLogic.natT
 | 
| 31368 | 400 |   fun simp_conv _ _ = SOME @{thm nat_mult_div_cancel_disj}
 | 
| 23164 | 401 | ); | 
| 402 | ||
| 23969 | 403 | structure DvdCancelFactor = ExtractCommonTermFun | 
| 404 | (open CancelFactorCommon | |
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34974diff
changeset | 405 |   val mk_bal   = HOLogic.mk_binrel @{const_name Rings.dvd}
 | 
| 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
34974diff
changeset | 406 |   val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} HOLogic.natT
 | 
| 31368 | 407 |   fun simp_conv _ _ = SOME @{thm nat_mult_dvd_cancel_disj}
 | 
| 23969 | 408 | ); | 
| 409 | ||
| 59582 | 410 | fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (Thm.term_of ct) | 
| 411 | fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (Thm.term_of ct) | |
| 412 | fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (Thm.term_of ct) | |
| 413 | fun div_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (Thm.term_of ct) | |
| 414 | fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (Thm.term_of ct) | |
| 23164 | 415 | |
| 416 | end; |