| author | nipkow | 
| Tue, 18 Jun 2013 17:37:51 +0200 | |
| changeset 52389 | 3971dd9ca831 | 
| parent 51717 | 9e7d1c139569 | 
| child 54230 | b1d955791529 | 
| permissions | -rw-r--r-- | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 1 | (* Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 2 | Copyright 2000 University of Cambridge | 
| 23164 | 3 | |
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 4 | Simprocs for the (integer) numerals. | 
| 23164 | 5 | *) | 
| 6 | ||
| 7 | (*To quote from Provers/Arith/cancel_numeral_factor.ML: | |
| 8 | ||
| 9 | Cancels common coefficients in balanced expressions: | |
| 10 | ||
| 11 | u*#m ~~ u'*#m' == #n*u ~~ #n'*u' | |
| 12 | ||
| 13 | where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /) | |
| 14 | and d = gcd(m,m') and n=m/d and n'=m'/d. | |
| 15 | *) | |
| 16 | ||
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 17 | signature NUMERAL_SIMPROCS = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 18 | sig | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 19 | val prep_simproc: theory -> string * string list * (Proof.context -> term -> thm option) | 
| 44945 | 20 | -> simproc | 
| 21 | val trans_tac: thm option -> tactic | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 22 | val assoc_fold: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 23 | val combine_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 24 | val eq_cancel_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 25 | val less_cancel_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 26 | val le_cancel_numerals: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 27 | val eq_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 28 | val le_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 29 | val less_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 30 | val div_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 31 | val mod_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 32 | val dvd_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 33 | val divide_cancel_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 34 | val eq_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 35 | val less_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 36 | val le_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 37 | val div_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 38 | val divide_cancel_numeral_factor: Proof.context -> cterm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 39 | val field_combine_numerals: Proof.context -> cterm -> thm option | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 40 | val field_cancel_numeral_factors: simproc list | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 41 | val num_ss: simpset | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 42 | val field_comp_conv: Proof.context -> conv | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 43 | end; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 44 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 45 | structure Numeral_Simprocs : NUMERAL_SIMPROCS = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 46 | struct | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 47 | |
| 44945 | 48 | fun prep_simproc thy (name, pats, proc) = | 
| 49 | Simplifier.simproc_global thy name pats proc; | |
| 50 | ||
| 51 | fun trans_tac NONE = all_tac | |
| 52 | | trans_tac (SOME th) = ALLGOALS (rtac (th RS trans)); | |
| 53 | ||
| 33359 | 54 | val mk_number = Arith_Data.mk_number; | 
| 55 | val mk_sum = Arith_Data.mk_sum; | |
| 56 | val long_mk_sum = Arith_Data.long_mk_sum; | |
| 57 | val dest_sum = Arith_Data.dest_sum; | |
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 58 | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 59 | val mk_diff = HOLogic.mk_binop @{const_name Groups.minus};
 | 
| 49323 | 60 | val dest_diff = HOLogic.dest_bin @{const_name Groups.minus} dummyT;
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 61 | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 62 | val mk_times = HOLogic.mk_binop @{const_name Groups.times};
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 63 | |
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 64 | fun one_of T = Const(@{const_name Groups.one}, T);
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 65 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 66 | (* build product with trailing 1 rather than Numeral 1 in order to avoid the | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 67 | unnecessary restriction to type class number_ring | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 68 | which is not required for cancellation of common factors in divisions. | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 69 | UPDATE: this reasoning no longer applies (number_ring is gone) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 70 | *) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 71 | fun mk_prod T = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 72 | let val one = one_of T | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 73 | fun mk [] = one | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 74 | | mk [t] = t | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 75 | | mk (t :: ts) = if t = one then mk ts else mk_times (t, mk ts) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 76 | in mk end; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 77 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 78 | (*This version ALWAYS includes a trailing one*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 79 | fun long_mk_prod T [] = one_of T | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 80 | | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 81 | |
| 49323 | 82 | val dest_times = HOLogic.dest_bin @{const_name Groups.times} dummyT;
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 83 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 84 | fun dest_prod t = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 85 | let val (t,u) = dest_times t | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 86 | in dest_prod t @ dest_prod u end | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 87 | handle TERM _ => [t]; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 88 | |
| 33359 | 89 | fun find_first_numeral past (t::terms) = | 
| 90 | ((snd (HOLogic.dest_number t), rev past @ terms) | |
| 91 | handle TERM _ => find_first_numeral (t::past) terms) | |
| 92 |   | find_first_numeral past [] = raise TERM("find_first_numeral", []);
 | |
| 93 | ||
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 94 | (*DON'T do the obvious simplifications; that would create special cases*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 95 | fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 96 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 97 | (*Express t as a product of (possibly) a numeral with other sorted terms*) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 98 | fun dest_coeff sign (Const (@{const_name Groups.uminus}, _) $ t) = dest_coeff (~sign) t
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 99 | | dest_coeff sign t = | 
| 35408 | 100 | let val ts = sort Term_Ord.term_ord (dest_prod t) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 101 | val (n, ts') = find_first_numeral [] ts | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 102 | handle TERM _ => (1, ts) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 103 | in (sign*n, mk_prod (Term.fastype_of t) ts') end; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 104 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 105 | (*Find first coefficient-term THAT MATCHES u*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 106 | fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 107 | | find_first_coeff past u (t::terms) = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 108 | let val (n,u') = dest_coeff 1 t | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 109 | in if u aconv u' then (n, rev past @ terms) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 110 | else find_first_coeff (t::past) u terms | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 111 | end | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 112 | handle TERM _ => find_first_coeff (t::past) u terms; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 113 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 114 | (*Fractions as pairs of ints. Can't use Rat.rat because the representation | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 115 | needs to preserve negative values in the denominator.*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 116 | fun mk_frac (p, q) = if q = 0 then raise Div else (p, q); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 117 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 118 | (*Don't reduce fractions; sums must be proved by rule add_frac_eq. | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 119 | Fractions are reduced later by the cancel_numeral_factor simproc.*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 120 | fun add_frac ((p1, q1), (p2, q2)) = (p1 * q2 + p2 * q1, q1 * q2); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 121 | |
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 122 | val mk_divide = HOLogic.mk_binop @{const_name Fields.divide};
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 123 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 124 | (*Build term (p / q) * t*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 125 | fun mk_fcoeff ((p, q), t) = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 126 | let val T = Term.fastype_of t | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 127 | in mk_times (mk_divide (mk_number T p, mk_number T q), t) end; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 128 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 129 | (*Express t as a product of a fraction with other sorted terms*) | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 130 | fun dest_fcoeff sign (Const (@{const_name Groups.uminus}, _) $ t) = dest_fcoeff (~sign) t
 | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 131 |   | dest_fcoeff sign (Const (@{const_name Fields.divide}, _) $ t $ u) =
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 132 | let val (p, t') = dest_coeff sign t | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 133 | val (q, u') = dest_coeff 1 u | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 134 | in (mk_frac (p, q), mk_divide (t', u')) end | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 135 | | dest_fcoeff sign t = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 136 | let val (p, t') = dest_coeff sign t | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 137 | val T = Term.fastype_of t | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 138 | in (mk_frac (p, 1), mk_divide (t', one_of T)) end; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 139 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 140 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 141 | (** New term ordering so that AC-rewriting brings numerals to the front **) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 142 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 143 | (*Order integers by absolute value and then by sign. The standard integer | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 144 | ordering is not well-founded.*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 145 | fun num_ord (i,j) = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 146 | (case int_ord (abs i, abs j) of | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 147 | EQUAL => int_ord (Int.sign i, Int.sign j) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 148 | | ord => ord); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 149 | |
| 35408 | 150 | (*This resembles Term_Ord.term_ord, but it puts binary numerals before other | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 151 | non-atomic terms.*) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 152 | local open Term | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 153 | in | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 154 | fun numterm_ord (t, u) = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 155 | case (try HOLogic.dest_number t, try HOLogic.dest_number u) of | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 156 | (SOME (_, i), SOME (_, j)) => num_ord (i, j) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 157 | | (SOME _, NONE) => LESS | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 158 | | (NONE, SOME _) => GREATER | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 159 | | _ => ( | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 160 | case (t, u) of | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 161 | (Abs (_, T, t), Abs(_, U, u)) => | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 162 | (prod_ord numterm_ord Term_Ord.typ_ord ((t, T), (u, U))) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 163 | | _ => ( | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 164 | case int_ord (size_of_term t, size_of_term u) of | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 165 | EQUAL => | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 166 | let val (f, ts) = strip_comb t and (g, us) = strip_comb u in | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 167 | (prod_ord Term_Ord.hd_ord numterms_ord ((f, ts), (g, us))) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 168 | end | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 169 | | ord => ord)) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 170 | and numterms_ord (ts, us) = list_ord numterm_ord (ts, us) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 171 | end; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 172 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 173 | fun numtermless tu = (numterm_ord tu = LESS); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 174 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 175 | val num_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 176 |   simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 177 | |> Simplifier.set_termless numtermless); | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 178 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 179 | (*Maps 1 to Numeral1 so that arithmetic isn't complicated by the abstract 1.*) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 180 | val numeral_syms = [@{thm numeral_1_eq_1} RS sym];
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 181 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 182 | (*Simplify 0+n, n+0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 183 | val add_0s =  @{thms add_0s};
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 184 | val mult_1s = @{thms mult_1s mult_1_left mult_1_right divide_1};
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 185 | |
| 45437 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 186 | (* For post-simplification of the rhs of simproc-generated rules *) | 
| 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 187 | val post_simps = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 188 |     [@{thm numeral_1_eq_1},
 | 
| 45437 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 189 |      @{thm add_0_left}, @{thm add_0_right},
 | 
| 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 190 |      @{thm mult_zero_left}, @{thm mult_zero_right},
 | 
| 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 191 |      @{thm mult_1_left}, @{thm mult_1_right},
 | 
| 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 192 |      @{thm mult_minus1}, @{thm mult_minus1_right}]
 | 
| 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 193 | |
| 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 194 | val field_post_simps = | 
| 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 195 |     post_simps @ [@{thm divide_zero_left}, @{thm divide_1}]
 | 
| 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 196 | |
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 197 | (*Simplify inverse Numeral1, a/Numeral1*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 198 | val inverse_1s = [@{thm inverse_numeral_1}];
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 199 | val divide_1s = [@{thm divide_numeral_1}];
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 200 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 201 | (*To perform binary arithmetic. The "left" rewriting handles patterns | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 202 | created by the Numeral_Simprocs, such as 3 * (5 * x). *) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 203 | val simps = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 204 |     [@{thm numeral_1_eq_1} RS sym] @
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 205 |     @{thms add_numeral_left} @
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 206 |     @{thms add_neg_numeral_left} @
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 207 |     @{thms mult_numeral_left} @
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 208 |     @{thms arith_simps} @ @{thms rel_simps};
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 209 | |
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 210 | (*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 211 | during re-arrangement*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 212 | val non_add_simps = | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 213 | subtract Thm.eq_thm | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 214 |     (@{thms add_numeral_left} @
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 215 |      @{thms add_neg_numeral_left} @
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 216 |      @{thms numeral_plus_numeral} @
 | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 217 |      @{thms add_neg_numeral_simps}) simps;
 | 
| 23164 | 218 | |
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 219 | (*To evaluate binary negations of coefficients*) | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 220 | val minus_simps = [@{thm minus_zero}, @{thm minus_one}, @{thm minus_numeral}, @{thm minus_neg_numeral}];
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 221 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 222 | (*To let us treat subtraction as addition*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 223 | val diff_simps = [@{thm diff_minus}, @{thm minus_add_distrib}, @{thm minus_minus}];
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 224 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 225 | (*To let us treat division as multiplication*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 226 | val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}];
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 227 | |
| 35020 
862a20ffa8e2
prefer explicit @{lemma} over adhoc forward reasoning;
 wenzelm parents: 
34974diff
changeset | 228 | (*push the unary minus down*) | 
| 
862a20ffa8e2
prefer explicit @{lemma} over adhoc forward reasoning;
 wenzelm parents: 
34974diff
changeset | 229 | val minus_mult_eq_1_to_2 = @{lemma "- (a::'a::ring) * b = a * - b" by simp};
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 230 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 231 | (*to extract again any uncancelled minuses*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 232 | val minus_from_mult_simps = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 233 |     [@{thm minus_minus}, @{thm mult_minus_left}, @{thm mult_minus_right}];
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 234 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 235 | (*combine unary minus with numeric literals, however nested within a product*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 236 | val mult_minus_simps = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 237 |     [@{thm mult_assoc}, @{thm minus_mult_left}, minus_mult_eq_1_to_2];
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 238 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 239 | val norm_ss1 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 240 |   simpset_of (put_simpset num_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 241 | addsimps numeral_syms @ add_0s @ mult_1s @ | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 242 |     diff_simps @ minus_simps @ @{thms add_ac})
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 243 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 244 | val norm_ss2 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 245 |   simpset_of (put_simpset num_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 246 | addsimps non_add_simps @ mult_minus_simps) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 247 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 248 | val norm_ss3 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 249 |   simpset_of (put_simpset num_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 250 |     addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac})
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 251 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 252 | structure CancelNumeralsCommon = | 
| 44945 | 253 | struct | 
| 254 | val mk_sum = mk_sum | |
| 255 | val dest_sum = dest_sum | |
| 256 | val mk_coeff = mk_coeff | |
| 257 | val dest_coeff = dest_coeff 1 | |
| 258 | val find_first_coeff = find_first_coeff [] | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 259 | val trans_tac = trans_tac | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 260 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 261 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 262 | ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 263 | THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 264 | THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 265 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 266 | val numeral_simp_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 267 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps add_0s @ simps)
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 268 | fun numeral_simp_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 269 | ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) | 
| 45437 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 270 | val simplify_meta_eq = Arith_Data.simplify_meta_eq post_simps | 
| 44945 | 271 | val prove_conv = Arith_Data.prove_conv | 
| 272 | end; | |
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 273 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 274 | structure EqCancelNumerals = CancelNumeralsFun | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 275 | (open CancelNumeralsCommon | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 276 | val mk_bal = HOLogic.mk_eq | 
| 49323 | 277 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} dummyT
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 278 |   val bal_add1 = @{thm eq_add_iff1} RS trans
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 279 |   val bal_add2 = @{thm eq_add_iff2} RS trans
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 280 | ); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 281 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 282 | structure LessCancelNumerals = CancelNumeralsFun | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 283 | (open CancelNumeralsCommon | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 284 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
 | 
| 49323 | 285 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} dummyT
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 286 |   val bal_add1 = @{thm less_add_iff1} RS trans
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 287 |   val bal_add2 = @{thm less_add_iff2} RS trans
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 288 | ); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 289 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 290 | structure LeCancelNumerals = CancelNumeralsFun | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 291 | (open CancelNumeralsCommon | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 292 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
 | 
| 49323 | 293 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} dummyT
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 294 |   val bal_add1 = @{thm le_add_iff1} RS trans
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 295 |   val bal_add2 = @{thm le_add_iff2} RS trans
 | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 296 | ); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 297 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 298 | fun eq_cancel_numerals ctxt ct = EqCancelNumerals.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 299 | fun less_cancel_numerals ctxt ct = LessCancelNumerals.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 300 | fun le_cancel_numerals ctxt ct = LeCancelNumerals.proc ctxt (term_of ct) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 301 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 302 | structure CombineNumeralsData = | 
| 44945 | 303 | struct | 
| 304 | type coeff = int | |
| 305 | val iszero = (fn x => x = 0) | |
| 306 | val add = op + | |
| 307 | val mk_sum = long_mk_sum (*to work for e.g. 2*x + 3*x *) | |
| 308 | val dest_sum = dest_sum | |
| 309 | val mk_coeff = mk_coeff | |
| 310 | val dest_coeff = dest_coeff 1 | |
| 311 |   val left_distrib = @{thm combine_common_factor} RS trans
 | |
| 312 | val prove_conv = Arith_Data.prove_conv_nohyps | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 313 | val trans_tac = trans_tac | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 314 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 315 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 316 | ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 317 | THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 318 | THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 319 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 320 | val numeral_simp_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 321 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps add_0s @ simps)
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 322 | fun numeral_simp_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 323 | ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) | 
| 45437 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 324 | val simplify_meta_eq = Arith_Data.simplify_meta_eq post_simps | 
| 44945 | 325 | end; | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 326 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 327 | structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 328 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 329 | (*Version for fields, where coefficients can be fractions*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 330 | structure FieldCombineNumeralsData = | 
| 44945 | 331 | struct | 
| 332 | type coeff = int * int | |
| 333 | val iszero = (fn (p, q) => p = 0) | |
| 334 | val add = add_frac | |
| 335 | val mk_sum = long_mk_sum | |
| 336 | val dest_sum = dest_sum | |
| 337 | val mk_coeff = mk_fcoeff | |
| 338 | val dest_coeff = dest_fcoeff 1 | |
| 339 |   val left_distrib = @{thm combine_common_factor} RS trans
 | |
| 340 | val prove_conv = Arith_Data.prove_conv_nohyps | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 341 | val trans_tac = trans_tac | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 342 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 343 | val norm_ss1a = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 344 |     simpset_of (put_simpset norm_ss1 @{context} addsimps inverse_1s @ divide_simps)
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 345 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 346 | ALLGOALS (simp_tac (put_simpset norm_ss1a ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 347 | THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 348 | THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 349 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 350 | val numeral_simp_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 351 |     simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 352 |       addsimps add_0s @ simps @ [@{thm add_frac_eq}, @{thm not_False_eq_True}])
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 353 | fun numeral_simp_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 354 | ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) | 
| 45437 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 355 | val simplify_meta_eq = Arith_Data.simplify_meta_eq field_post_simps | 
| 44945 | 356 | end; | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 357 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 358 | structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 359 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 360 | fun combine_numerals ctxt ct = CombineNumerals.proc ctxt (term_of ct) | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 361 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 362 | fun field_combine_numerals ctxt ct = FieldCombineNumerals.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 363 | |
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 364 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 365 | (** Constant folding for multiplication in semirings **) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 366 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 367 | (*We do not need folding for addition: combine_numerals does the same thing*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 368 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 369 | structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 370 | struct | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 371 |   val assoc_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms mult_ac})
 | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 372 | val eq_reflection = eq_reflection | 
| 35983 
27e2fa7d4ce7
slightly more general simproc (avoids errors of linarith)
 boehmes parents: 
35408diff
changeset | 373 | val is_numeral = can HOLogic.dest_number | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 374 | end; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 375 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 376 | structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 377 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 378 | fun assoc_fold ctxt ct = Semiring_Times_Assoc.proc ctxt (term_of ct) | 
| 23164 | 379 | |
| 380 | structure CancelNumeralFactorCommon = | |
| 44945 | 381 | struct | 
| 382 | val mk_coeff = mk_coeff | |
| 383 | val dest_coeff = dest_coeff 1 | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 384 | val trans_tac = trans_tac | 
| 23164 | 385 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 386 | val norm_ss1 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 387 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps minus_from_mult_simps @ mult_1s)
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 388 | val norm_ss2 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 389 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps simps @ mult_minus_simps)
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 390 | val norm_ss3 = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 391 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps @{thms mult_ac})
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 392 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 393 | ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 394 | THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 395 | THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) | 
| 23164 | 396 | |
| 45668 
0ea1c705eebb
use HOL_basic_ss instead of HOL_ss for internal simproc proofs (cf. b36eedcd1633)
 huffman parents: 
45625diff
changeset | 397 | (* simp_thms are necessary because some of the cancellation rules below | 
| 
0ea1c705eebb
use HOL_basic_ss instead of HOL_ss for internal simproc proofs (cf. b36eedcd1633)
 huffman parents: 
45625diff
changeset | 398 | (e.g. mult_less_cancel_left) introduce various logical connectives *) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 399 | val numeral_simp_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 400 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps simps @ @{thms simp_thms})
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 401 | fun numeral_simp_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 402 | ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) | 
| 30518 | 403 | val simplify_meta_eq = Arith_Data.simplify_meta_eq | 
| 45437 
958d19d3405b
tune post-processing of simproc-generated rules so they won't produce Numeral0 or Numeral1
 huffman parents: 
45306diff
changeset | 404 |     ([@{thm Nat.add_0}, @{thm Nat.add_0_right}] @ post_simps)
 | 
| 44945 | 405 | val prove_conv = Arith_Data.prove_conv | 
| 406 | end | |
| 23164 | 407 | |
| 30931 
86ca651da03e
generalized some simprocs from int to semiring_div
 haftmann parents: 
30685diff
changeset | 408 | (*Version for semiring_div*) | 
| 
86ca651da03e
generalized some simprocs from int to semiring_div
 haftmann parents: 
30685diff
changeset | 409 | structure DivCancelNumeralFactor = CancelNumeralFactorFun | 
| 23164 | 410 | (open CancelNumeralFactorCommon | 
| 411 |   val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
 | |
| 49323 | 412 |   val dest_bal = HOLogic.dest_bin @{const_name Divides.div} dummyT
 | 
| 30931 
86ca651da03e
generalized some simprocs from int to semiring_div
 haftmann parents: 
30685diff
changeset | 413 |   val cancel = @{thm div_mult_mult1} RS trans
 | 
| 23164 | 414 | val neg_exchanges = false | 
| 415 | ) | |
| 416 | ||
| 417 | (*Version for fields*) | |
| 418 | structure DivideCancelNumeralFactor = CancelNumeralFactorFun | |
| 419 | (open CancelNumeralFactorCommon | |
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 420 |   val mk_bal   = HOLogic.mk_binop @{const_name Fields.divide}
 | 
| 49323 | 421 |   val dest_bal = HOLogic.dest_bin @{const_name Fields.divide} dummyT
 | 
| 23413 
5caa2710dd5b
tuned laws for cancellation in divisions for fields.
 nipkow parents: 
23401diff
changeset | 422 |   val cancel = @{thm mult_divide_mult_cancel_left} RS trans
 | 
| 23164 | 423 | val neg_exchanges = false | 
| 424 | ) | |
| 425 | ||
| 426 | structure EqCancelNumeralFactor = CancelNumeralFactorFun | |
| 427 | (open CancelNumeralFactorCommon | |
| 428 | val mk_bal = HOLogic.mk_eq | |
| 49323 | 429 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} dummyT
 | 
| 23164 | 430 |   val cancel = @{thm mult_cancel_left} RS trans
 | 
| 431 | val neg_exchanges = false | |
| 432 | ) | |
| 433 | ||
| 434 | structure LessCancelNumeralFactor = CancelNumeralFactorFun | |
| 435 | (open CancelNumeralFactorCommon | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 436 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
 | 
| 49323 | 437 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} dummyT
 | 
| 23164 | 438 |   val cancel = @{thm mult_less_cancel_left} RS trans
 | 
| 439 | val neg_exchanges = true | |
| 440 | ) | |
| 441 | ||
| 442 | structure LeCancelNumeralFactor = CancelNumeralFactorFun | |
| 443 | (open CancelNumeralFactorCommon | |
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 444 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
 | 
| 49323 | 445 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} dummyT
 | 
| 23164 | 446 |   val cancel = @{thm mult_le_cancel_left} RS trans
 | 
| 447 | val neg_exchanges = true | |
| 448 | ) | |
| 449 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 450 | fun eq_cancel_numeral_factor ctxt ct = EqCancelNumeralFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 451 | fun less_cancel_numeral_factor ctxt ct = LessCancelNumeralFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 452 | fun le_cancel_numeral_factor ctxt ct = LeCancelNumeralFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 453 | fun div_cancel_numeral_factor ctxt ct = DivCancelNumeralFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 454 | fun divide_cancel_numeral_factor ctxt ct = DivideCancelNumeralFactor.proc ctxt (term_of ct) | 
| 23164 | 455 | |
| 456 | val field_cancel_numeral_factors = | |
| 44945 | 457 |   map (prep_simproc @{theory})
 | 
| 23164 | 458 |    [("field_eq_cancel_numeral_factor",
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 459 | ["(l::'a::field) * m = n", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 460 | "(l::'a::field) = m * n"], | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 461 | EqCancelNumeralFactor.proc), | 
| 23164 | 462 |     ("field_cancel_numeral_factor",
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 463 | ["((l::'a::field_inverse_zero) * m) / n", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 464 | "(l::'a::field_inverse_zero) / (m * n)", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 465 | "((numeral v)::'a::field_inverse_zero) / (numeral w)", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 466 | "((numeral v)::'a::field_inverse_zero) / (neg_numeral w)", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 467 | "((neg_numeral v)::'a::field_inverse_zero) / (numeral w)", | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 468 | "((neg_numeral v)::'a::field_inverse_zero) / (neg_numeral w)"], | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 469 | DivideCancelNumeralFactor.proc)] | 
| 23164 | 470 | |
| 471 | ||
| 472 | (** Declarations for ExtractCommonTerm **) | |
| 473 | ||
| 474 | (*Find first term that matches u*) | |
| 475 | fun find_first_t past u []         = raise TERM ("find_first_t", [])
 | |
| 476 | | find_first_t past u (t::terms) = | |
| 477 | if u aconv t then (rev past @ terms) | |
| 478 | else find_first_t (t::past) u terms | |
| 479 | handle TERM _ => find_first_t (t::past) u terms; | |
| 480 | ||
| 481 | (** Final simplification for the CancelFactor simprocs **) | |
| 30518 | 482 | val simplify_one = Arith_Data.simplify_meta_eq | 
| 30031 | 483 |   [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_by_1}, @{thm numeral_1_eq_1}];
 | 
| 23164 | 484 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 485 | fun cancel_simplify_meta_eq ctxt cancel_th th = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 486 | simplify_one ctxt (([th, cancel_th]) MRS trans); | 
| 23164 | 487 | |
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 488 | local | 
| 31067 | 489 |   val Tp_Eq = Thm.reflexive (Thm.cterm_of @{theory HOL} HOLogic.Trueprop)
 | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 490 | fun Eq_True_elim Eq = | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 491 |     Thm.equal_elim (Thm.combination Tp_Eq (Thm.symmetric Eq)) @{thm TrueI}
 | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 492 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 493 | fun sign_conv pos_th neg_th ctxt t = | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 494 | let val T = fastype_of t; | 
| 35267 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 haftmann parents: 
35092diff
changeset | 495 |       val zero = Const(@{const_name Groups.zero}, T);
 | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 496 |       val less = Const(@{const_name Orderings.less}, [T,T] ---> HOLogic.boolT);
 | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 497 | val pos = less $ zero $ t and neg = less $ t $ zero | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 498 | val thy = Proof_Context.theory_of ctxt | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 499 | fun prove p = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 500 | SOME (Eq_True_elim (Simplifier.asm_rewrite ctxt (Thm.cterm_of thy p))) | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 501 | handle THM _ => NONE | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 502 | in case prove pos of | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 503 | SOME th => SOME(th RS pos_th) | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 504 | | NONE => (case prove neg of | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 505 | SOME th => SOME(th RS neg_th) | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 506 | | NONE => NONE) | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 507 | end; | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 508 | end | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 509 | |
| 23164 | 510 | structure CancelFactorCommon = | 
| 44945 | 511 | struct | 
| 512 | val mk_sum = long_mk_prod | |
| 513 | val dest_sum = dest_prod | |
| 514 | val mk_coeff = mk_coeff | |
| 515 | val dest_coeff = dest_coeff | |
| 516 | val find_first = find_first_t [] | |
| 44947 
8ae418dfe561
dropped unused argument – avoids problem with SML/NJ
 haftmann parents: 
44945diff
changeset | 517 | val trans_tac = trans_tac | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 518 | val norm_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 519 |     simpset_of (put_simpset HOL_basic_ss @{context} addsimps mult_1s @ @{thms mult_ac})
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 520 | fun norm_tac ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 521 | ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 522 | 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: 
44984diff
changeset | 523 | fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) | 
| 44945 | 524 | end; | 
| 23164 | 525 | |
| 526 | (*mult_cancel_left requires a ring with no zero divisors.*) | |
| 527 | structure EqCancelFactor = ExtractCommonTermFun | |
| 528 | (open CancelFactorCommon | |
| 529 | val mk_bal = HOLogic.mk_eq | |
| 49323 | 530 |   val dest_bal = HOLogic.dest_bin @{const_name HOL.eq} dummyT
 | 
| 31368 | 531 |   fun simp_conv _ _ = SOME @{thm mult_cancel_left}
 | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 532 | ); | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 533 | |
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 534 | (*for ordered rings*) | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 535 | structure LeCancelFactor = ExtractCommonTermFun | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 536 | (open CancelFactorCommon | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 537 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less_eq}
 | 
| 49323 | 538 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} dummyT
 | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 539 | val simp_conv = sign_conv | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 540 |     @{thm mult_le_cancel_left_pos} @{thm mult_le_cancel_left_neg}
 | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 541 | ); | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 542 | |
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 543 | (*for ordered rings*) | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 544 | structure LessCancelFactor = ExtractCommonTermFun | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 545 | (open CancelFactorCommon | 
| 35092 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 haftmann parents: 
35084diff
changeset | 546 |   val mk_bal   = HOLogic.mk_binrel @{const_name Orderings.less}
 | 
| 49323 | 547 |   val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} dummyT
 | 
| 30649 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 548 | val simp_conv = sign_conv | 
| 
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
 nipkow parents: 
30518diff
changeset | 549 |     @{thm mult_less_cancel_left_pos} @{thm mult_less_cancel_left_neg}
 | 
| 23164 | 550 | ); | 
| 551 | ||
| 30931 
86ca651da03e
generalized some simprocs from int to semiring_div
 haftmann parents: 
30685diff
changeset | 552 | (*for semirings with division*) | 
| 
86ca651da03e
generalized some simprocs from int to semiring_div
 haftmann parents: 
30685diff
changeset | 553 | structure DivCancelFactor = ExtractCommonTermFun | 
| 23164 | 554 | (open CancelFactorCommon | 
| 555 |   val mk_bal   = HOLogic.mk_binop @{const_name Divides.div}
 | |
| 49323 | 556 |   val dest_bal = HOLogic.dest_bin @{const_name Divides.div} dummyT
 | 
| 31368 | 557 |   fun simp_conv _ _ = SOME @{thm div_mult_mult1_if}
 | 
| 23164 | 558 | ); | 
| 559 | ||
| 30931 
86ca651da03e
generalized some simprocs from int to semiring_div
 haftmann parents: 
30685diff
changeset | 560 | structure ModCancelFactor = ExtractCommonTermFun | 
| 24395 | 561 | (open CancelFactorCommon | 
| 562 |   val mk_bal   = HOLogic.mk_binop @{const_name Divides.mod}
 | |
| 49323 | 563 |   val dest_bal = HOLogic.dest_bin @{const_name Divides.mod} dummyT
 | 
| 31368 | 564 |   fun simp_conv _ _ = SOME @{thm mod_mult_mult1}
 | 
| 24395 | 565 | ); | 
| 566 | ||
| 30931 
86ca651da03e
generalized some simprocs from int to semiring_div
 haftmann parents: 
30685diff
changeset | 567 | (*for idoms*) | 
| 
86ca651da03e
generalized some simprocs from int to semiring_div
 haftmann parents: 
30685diff
changeset | 568 | structure DvdCancelFactor = ExtractCommonTermFun | 
| 23969 | 569 | (open CancelFactorCommon | 
| 35050 
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
 haftmann parents: 
35043diff
changeset | 570 |   val mk_bal   = HOLogic.mk_binrel @{const_name Rings.dvd}
 | 
| 49323 | 571 |   val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} dummyT
 | 
| 31368 | 572 |   fun simp_conv _ _ = SOME @{thm dvd_mult_cancel_left}
 | 
| 23969 | 573 | ); | 
| 574 | ||
| 23164 | 575 | (*Version for all fields, including unordered ones (type complex).*) | 
| 576 | structure DivideCancelFactor = ExtractCommonTermFun | |
| 577 | (open CancelFactorCommon | |
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 578 |   val mk_bal   = HOLogic.mk_binop @{const_name Fields.divide}
 | 
| 49323 | 579 |   val dest_bal = HOLogic.dest_bin @{const_name Fields.divide} dummyT
 | 
| 31368 | 580 |   fun simp_conv _ _ = SOME @{thm mult_divide_mult_cancel_left_if}
 | 
| 23164 | 581 | ); | 
| 582 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 583 | fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 584 | fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 585 | fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 586 | fun div_cancel_factor ctxt ct = DivCancelFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 587 | fun mod_cancel_factor ctxt ct = ModCancelFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 588 | fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (term_of ct) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 589 | fun divide_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (term_of ct) | 
| 23164 | 590 | |
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 591 | local | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 592 |  val zr = @{cpat "0"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 593 | val zT = ctyp_of_term zr | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38549diff
changeset | 594 |  val geq = @{cpat HOL.eq}
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 595 | val eqT = Thm.dest_ctyp (ctyp_of_term geq) |> hd | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 596 |  val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 597 |  val add_frac_num = mk_meta_eq @{thm "add_frac_num"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 598 |  val add_num_frac = mk_meta_eq @{thm "add_num_frac"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 599 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 600 | fun prove_nz ctxt T t = | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 601 | let | 
| 36945 | 602 | val z = Thm.instantiate_cterm ([(zT,T)],[]) zr | 
| 603 | val eq = Thm.instantiate_cterm ([(eqT,T)],[]) geq | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 604 |       val th = Simplifier.rewrite (ctxt addsimps @{thms simp_thms})
 | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
46240diff
changeset | 605 |            (Thm.apply @{cterm "Trueprop"} (Thm.apply @{cterm "Not"}
 | 
| 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
46240diff
changeset | 606 | (Thm.apply (Thm.apply eq t) z))) | 
| 36945 | 607 | in Thm.equal_elim (Thm.symmetric th) TrueI | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 608 | end | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 609 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 610 | fun proc phi ctxt ct = | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 611 | let | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 612 | val ((x,y),(w,z)) = | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 613 | (Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 614 | val _ = map (HOLogic.dest_number o term_of) [x,y,z,w] | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 615 | val T = ctyp_of_term x | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 616 | val [y_nz, z_nz] = map (prove_nz ctxt T) [y, z] | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 617 | val th = instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq | 
| 36945 | 618 | in SOME (Thm.implies_elim (Thm.implies_elim th y_nz) z_nz) | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 619 | end | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 620 | handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 621 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 622 | fun proc2 phi ctxt ct = | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 623 | let | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 624 | val (l,r) = Thm.dest_binop ct | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 625 | val T = ctyp_of_term l | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 626 | in (case (term_of l, term_of r) of | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 627 |       (Const(@{const_name Fields.divide},_)$_$_, _) =>
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 628 | let val (x,y) = Thm.dest_binop l val z = r | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 629 | val _ = map (HOLogic.dest_number o term_of) [x,y,z] | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 630 | val ynz = prove_nz ctxt T y | 
| 36945 | 631 | in SOME (Thm.implies_elim (instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz) | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 632 | end | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 633 |      | (_, Const (@{const_name Fields.divide},_)$_$_) =>
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 634 | let val (x,y) = Thm.dest_binop r val z = l | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 635 | val _ = map (HOLogic.dest_number o term_of) [x,y,z] | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 636 | val ynz = prove_nz ctxt T y | 
| 36945 | 637 | in SOME (Thm.implies_elim (instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz) | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 638 | end | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 639 | | _ => NONE) | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 640 | end | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 641 | handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 642 | |
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 643 |  fun is_number (Const(@{const_name Fields.divide},_)$a$b) = is_number a andalso is_number b
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 644 | | is_number t = can HOLogic.dest_number t | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 645 | |
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 646 | val is_number = is_number o term_of | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 647 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 648 | fun proc3 phi ctxt ct = | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 649 | (case term_of ct of | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 650 |     Const(@{const_name Orderings.less},_)$(Const(@{const_name Fields.divide},_)$_$_)$_ =>
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 651 | let | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 652 | val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 653 | val _ = map is_number [a,b,c] | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 654 | val T = ctyp_of_term c | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 655 |         val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_less_eq"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 656 | in SOME (mk_meta_eq th) end | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 657 |   | Const(@{const_name Orderings.less_eq},_)$(Const(@{const_name Fields.divide},_)$_$_)$_ =>
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 658 | let | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 659 | val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 660 | val _ = map is_number [a,b,c] | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 661 | val T = ctyp_of_term c | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 662 |         val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 663 | in SOME (mk_meta_eq th) end | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 664 |   | Const(@{const_name HOL.eq},_)$(Const(@{const_name Fields.divide},_)$_$_)$_ =>
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 665 | let | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 666 | val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 667 | val _ = map is_number [a,b,c] | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 668 | val T = ctyp_of_term c | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 669 |         val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_eq_eq"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 670 | in SOME (mk_meta_eq th) end | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 671 |   | Const(@{const_name Orderings.less},_)$_$(Const(@{const_name Fields.divide},_)$_$_) =>
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 672 | let | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 673 | val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 674 | val _ = map is_number [a,b,c] | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 675 | val T = ctyp_of_term c | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 676 |         val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "less_divide_eq"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 677 | in SOME (mk_meta_eq th) end | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 678 |   | Const(@{const_name Orderings.less_eq},_)$_$(Const(@{const_name Fields.divide},_)$_$_) =>
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 679 | let | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 680 | val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 681 | val _ = map is_number [a,b,c] | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 682 | val T = ctyp_of_term c | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 683 |         val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 684 | in SOME (mk_meta_eq th) end | 
| 44064 
5bce8ff0d9ae
moved division ring stuff from Rings.thy to Fields.thy
 huffman parents: 
40878diff
changeset | 685 |   | Const(@{const_name HOL.eq},_)$_$(Const(@{const_name Fields.divide},_)$_$_) =>
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 686 | let | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 687 | val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 688 | val _ = map is_number [a,b,c] | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 689 | val T = ctyp_of_term c | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 690 |         val th = instantiate' [SOME T] (map SOME [a,b,c]) @{thm "eq_divide_eq"}
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 691 | in SOME (mk_meta_eq th) end | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 692 | | _ => NONE) | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 693 | handle TERM _ => NONE | CTERM _ => NONE | THM _ => NONE | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 694 | |
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 695 | val add_frac_frac_simproc = | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 696 |        make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + (?w::?'a::field)/?z"}],
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 697 | name = "add_frac_frac_simproc", | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 698 | proc = proc, identifier = []} | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 699 | |
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 700 | val add_frac_num_simproc = | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 701 |        make_simproc {lhss = [@{cpat "(?x::?'a::field)/?y + ?z"}, @{cpat "?z + (?x::?'a::field)/?y"}],
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 702 | name = "add_frac_num_simproc", | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 703 | proc = proc2, identifier = []} | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 704 | |
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 705 | val ord_frac_simproc = | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 706 | make_simproc | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 707 |     {lhss = [@{cpat "(?a::(?'a::{field, ord}))/?b < ?c"},
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 708 |              @{cpat "(?a::(?'a::{field, ord}))/?b <= ?c"},
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 709 |              @{cpat "?c < (?a::(?'a::{field, ord}))/?b"},
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 710 |              @{cpat "?c <= (?a::(?'a::{field, ord}))/?b"},
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 711 |              @{cpat "?c = ((?a::(?'a::{field, ord}))/?b)"},
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 712 |              @{cpat "((?a::(?'a::{field, ord}))/ ?b) = ?c"}],
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 713 | name = "ord_frac_simproc", proc = proc3, identifier = []} | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 714 | |
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 715 | val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"},
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 716 |            @{thm "divide_Numeral1"},
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 717 |            @{thm "divide_zero"}, @{thm divide_zero_left},
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 718 |            @{thm "divide_divide_eq_left"}, 
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 719 |            @{thm "times_divide_eq_left"}, @{thm "times_divide_eq_right"},
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 720 |            @{thm "times_divide_times_eq"},
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 721 |            @{thm "divide_divide_eq_right"},
 | 
| 37887 | 722 |            @{thm "diff_minus"}, @{thm "minus_divide_left"},
 | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 723 |            @{thm "add_divide_distrib"} RS sym,
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 724 |            @{thm field_divide_inverse} RS sym, @{thm inverse_divide}, 
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 725 |            Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult_commute}))))   
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 726 |            (@{thm field_divide_inverse} RS sym)]
 | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 727 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 728 | val field_comp_ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 729 | simpset_of | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 730 |     (put_simpset HOL_basic_ss @{context}
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 731 |       addsimps @{thms "semiring_norm"}
 | 
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45437diff
changeset | 732 |       addsimps ths addsimps @{thms simp_thms}
 | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45437diff
changeset | 733 | addsimprocs field_cancel_numeral_factors | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45437diff
changeset | 734 | addsimprocs [add_frac_frac_simproc, add_frac_num_simproc, ord_frac_simproc] | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45437diff
changeset | 735 |       |> Simplifier.add_cong @{thm "if_weak_cong"})
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 736 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 737 | in | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 738 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 739 | fun field_comp_conv ctxt = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 740 | Simplifier.rewrite (put_simpset field_comp_ss ctxt) | 
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45437diff
changeset | 741 | then_conv | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
49323diff
changeset | 742 |   Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_1_eq_1}])
 | 
| 36751 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 743 | |
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 744 | end | 
| 
7f1da69cacb3
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
 haftmann parents: 
36409diff
changeset | 745 | |
| 23164 | 746 | end; | 
| 747 | ||
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 748 | (*examples: | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 749 | print_depth 22; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 750 | set timing; | 
| 40878 
7695e4de4d86
renamed trace_simp to simp_trace, and debug_simp to simp_debug;
 wenzelm parents: 
38864diff
changeset | 751 | set simp_trace; | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 752 | fun test s = (Goal s, by (Simp_tac 1)); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 753 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 754 | test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 755 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 756 | test "2*u = (u::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 757 | test "(i + j + 12 + (k::int)) - 15 = y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 758 | test "(i + j + 12 + (k::int)) - 5 = y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 759 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 760 | test "y - b < (b::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 761 | test "y - (3*b + c) < (b::int) - 2*c"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 762 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 763 | test "(2*x - (u*v) + y) - v*3*u = (w::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 764 | test "(2*x*u*v + (u*v)*4 + y) - v*u*4 = (w::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 765 | test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 766 | test "u*v - (x*u*v + (u*v)*4 + y) = (w::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 767 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 768 | test "(i + j + 12 + (k::int)) = u + 15 + y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 769 | test "(i + j*2 + 12 + (k::int)) = j + 5 + y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 770 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 771 | test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 772 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 773 | test "a + -(b+c) + b = (d::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 774 | test "a + -(b+c) - b = (d::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 775 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 776 | (*negative numerals*) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 777 | test "(i + j + -2 + (k::int)) - (u + 5 + y) = zz"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 778 | test "(i + j + -3 + (k::int)) < u + 5 + y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 779 | test "(i + j + 3 + (k::int)) < u + -6 + y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 780 | test "(i + j + -12 + (k::int)) - 15 = y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 781 | test "(i + j + 12 + (k::int)) - -15 = y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 782 | test "(i + j + -12 + (k::int)) - -15 = y"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 783 | *) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 784 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 785 | (*examples: | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 786 | print_depth 22; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 787 | set timing; | 
| 40878 
7695e4de4d86
renamed trace_simp to simp_trace, and debug_simp to simp_debug;
 wenzelm parents: 
38864diff
changeset | 788 | set simp_trace; | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 789 | fun test s = (Goal s; by (Simp_tac 1)); | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 790 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 791 | test "9*x = 12 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 792 | test "(9*x) div (12 * (y::int)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 793 | test "9*x < 12 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 794 | test "9*x <= 12 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 795 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 796 | test "-99*x = 132 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 797 | test "(-99*x) div (132 * (y::int)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 798 | test "-99*x < 132 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 799 | test "-99*x <= 132 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 800 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 801 | test "999*x = -396 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 802 | test "(999*x) div (-396 * (y::int)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 803 | test "999*x < -396 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 804 | test "999*x <= -396 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 805 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 806 | test "-99*x = -81 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 807 | test "(-99*x) div (-81 * (y::int)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 808 | test "-99*x <= -81 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 809 | test "-99*x < -81 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 810 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 811 | test "-2 * x = -1 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 812 | test "-2 * x = -(y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 813 | test "(-2 * x) div (-1 * (y::int)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 814 | test "-2 * x < -(y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 815 | test "-2 * x <= -1 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 816 | test "-x < -23 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 817 | test "-x <= -23 * (y::int)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 818 | *) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 819 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 820 | (*And the same examples for fields such as rat or real: | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 821 | test "0 <= (y::rat) * -2"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 822 | test "9*x = 12 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 823 | test "(9*x) / (12 * (y::rat)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 824 | test "9*x < 12 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 825 | test "9*x <= 12 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 826 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 827 | test "-99*x = 132 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 828 | test "(-99*x) / (132 * (y::rat)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 829 | test "-99*x < 132 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 830 | test "-99*x <= 132 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 831 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 832 | test "999*x = -396 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 833 | test "(999*x) / (-396 * (y::rat)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 834 | test "999*x < -396 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 835 | test "999*x <= -396 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 836 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 837 | test "(- ((2::rat) * x) <= 2 * y)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 838 | test "-99*x = -81 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 839 | test "(-99*x) / (-81 * (y::rat)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 840 | test "-99*x <= -81 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 841 | test "-99*x < -81 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 842 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 843 | test "-2 * x = -1 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 844 | test "-2 * x = -(y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 845 | test "(-2 * x) / (-1 * (y::rat)) = z"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 846 | test "-2 * x < -(y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 847 | test "-2 * x <= -1 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 848 | test "-x < -23 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 849 | test "-x <= -23 * (y::rat)"; | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 850 | *) | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31067diff
changeset | 851 | |
| 23164 | 852 | (*examples: | 
| 853 | print_depth 22; | |
| 854 | set timing; | |
| 40878 
7695e4de4d86
renamed trace_simp to simp_trace, and debug_simp to simp_debug;
 wenzelm parents: 
38864diff
changeset | 855 | set simp_trace; | 
| 23164 | 856 | fun test s = (Goal s; by (Asm_simp_tac 1)); | 
| 857 | ||
| 858 | test "x*k = k*(y::int)"; | |
| 859 | test "k = k*(y::int)"; | |
| 860 | test "a*(b*c) = (b::int)"; | |
| 861 | test "a*(b*c) = d*(b::int)*(x*a)"; | |
| 862 | ||
| 863 | test "(x*k) div (k*(y::int)) = (uu::int)"; | |
| 864 | test "(k) div (k*(y::int)) = (uu::int)"; | |
| 865 | test "(a*(b*c)) div ((b::int)) = (uu::int)"; | |
| 866 | test "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)"; | |
| 867 | *) | |
| 868 | ||
| 869 | (*And the same examples for fields such as rat or real: | |
| 870 | print_depth 22; | |
| 871 | set timing; | |
| 40878 
7695e4de4d86
renamed trace_simp to simp_trace, and debug_simp to simp_debug;
 wenzelm parents: 
38864diff
changeset | 872 | set simp_trace; | 
| 23164 | 873 | fun test s = (Goal s; by (Asm_simp_tac 1)); | 
| 874 | ||
| 875 | test "x*k = k*(y::rat)"; | |
| 876 | test "k = k*(y::rat)"; | |
| 877 | test "a*(b*c) = (b::rat)"; | |
| 878 | test "a*(b*c) = d*(b::rat)*(x*a)"; | |
| 879 | ||
| 880 | ||
| 881 | test "(x*k) / (k*(y::rat)) = (uu::rat)"; | |
| 882 | test "(k) / (k*(y::rat)) = (uu::rat)"; | |
| 883 | test "(a*(b*c)) / ((b::rat)) = (uu::rat)"; | |
| 884 | test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)"; | |
| 885 | ||
| 886 | (*FIXME: what do we do about this?*) | |
| 887 | test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z"; | |
| 888 | *) |