| author | blanchet |
| Wed, 28 Apr 2010 21:59:29 +0200 | |
| changeset 36551 | cc42df660808 |
| parent 36409 | d323e7773aa8 |
| child 36751 | 7f1da69cacb3 |
| permissions | -rw-r--r-- |
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
1 |
(* Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
2 |
Copyright 2000 University of Cambridge |
| 23164 | 3 |
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
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:
31067
diff
changeset
|
17 |
signature NUMERAL_SIMPROCS = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
18 |
sig |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
19 |
val assoc_fold_simproc: simproc |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
20 |
val combine_numerals: simproc |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
21 |
val cancel_numerals: simproc list |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
22 |
val cancel_factors: simproc list |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
23 |
val cancel_numeral_factors: simproc list |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
24 |
val field_combine_numerals: simproc |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
25 |
val field_cancel_numeral_factors: simproc list |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
26 |
val num_ss: simpset |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
27 |
end; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
28 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
29 |
structure Numeral_Simprocs : NUMERAL_SIMPROCS = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
30 |
struct |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
31 |
|
| 33359 | 32 |
val mk_number = Arith_Data.mk_number; |
33 |
val mk_sum = Arith_Data.mk_sum; |
|
34 |
val long_mk_sum = Arith_Data.long_mk_sum; |
|
35 |
val dest_sum = Arith_Data.dest_sum; |
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
36 |
|
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
37 |
val mk_diff = HOLogic.mk_binop @{const_name Groups.minus};
|
|
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
38 |
val dest_diff = HOLogic.dest_bin @{const_name Groups.minus} Term.dummyT;
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
39 |
|
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
40 |
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:
31067
diff
changeset
|
41 |
|
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
42 |
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:
31067
diff
changeset
|
43 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
44 |
(* 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:
31067
diff
changeset
|
45 |
unnecessary restriction to type class number_ring |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
46 |
which is not required for cancellation of common factors in divisions. |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
47 |
*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
48 |
fun mk_prod T = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
49 |
let val one = one_of T |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
50 |
fun mk [] = one |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
51 |
| mk [t] = t |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
52 |
| 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:
31067
diff
changeset
|
53 |
in mk end; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
54 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
55 |
(*This version ALWAYS includes a trailing one*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
56 |
fun long_mk_prod T [] = one_of T |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
57 |
| 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:
31067
diff
changeset
|
58 |
|
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
59 |
val dest_times = HOLogic.dest_bin @{const_name Groups.times} Term.dummyT;
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
60 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
61 |
fun dest_prod t = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
62 |
let val (t,u) = dest_times t |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
63 |
in dest_prod t @ dest_prod u end |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
64 |
handle TERM _ => [t]; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
65 |
|
| 33359 | 66 |
fun find_first_numeral past (t::terms) = |
67 |
((snd (HOLogic.dest_number t), rev past @ terms) |
|
68 |
handle TERM _ => find_first_numeral (t::past) terms) |
|
69 |
| find_first_numeral past [] = raise TERM("find_first_numeral", []);
|
|
70 |
||
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
71 |
(*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:
31067
diff
changeset
|
72 |
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:
31067
diff
changeset
|
73 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
74 |
(*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:
35092
diff
changeset
|
75 |
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:
31067
diff
changeset
|
76 |
| dest_coeff sign t = |
| 35408 | 77 |
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:
31067
diff
changeset
|
78 |
val (n, ts') = find_first_numeral [] ts |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
79 |
handle TERM _ => (1, ts) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
80 |
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:
31067
diff
changeset
|
81 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
82 |
(*Find first coefficient-term THAT MATCHES u*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
83 |
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:
31067
diff
changeset
|
84 |
| find_first_coeff past u (t::terms) = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
85 |
let val (n,u') = dest_coeff 1 t |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
86 |
in if u aconv u' then (n, rev past @ terms) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
87 |
else find_first_coeff (t::past) u terms |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
88 |
end |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
89 |
handle TERM _ => find_first_coeff (t::past) u terms; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
90 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
91 |
(*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:
31067
diff
changeset
|
92 |
needs to preserve negative values in the denominator.*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
93 |
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:
31067
diff
changeset
|
94 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
95 |
(*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:
31067
diff
changeset
|
96 |
Fractions are reduced later by the cancel_numeral_factor simproc.*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
97 |
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:
31067
diff
changeset
|
98 |
|
| 35084 | 99 |
val mk_divide = HOLogic.mk_binop @{const_name Rings.divide};
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
100 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
101 |
(*Build term (p / q) * t*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
102 |
fun mk_fcoeff ((p, q), t) = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
103 |
let val T = Term.fastype_of t |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
104 |
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:
31067
diff
changeset
|
105 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
106 |
(*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:
35092
diff
changeset
|
107 |
fun dest_fcoeff sign (Const (@{const_name Groups.uminus}, _) $ t) = dest_fcoeff (~sign) t
|
| 35084 | 108 |
| dest_fcoeff sign (Const (@{const_name Rings.divide}, _) $ t $ u) =
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
109 |
let val (p, t') = dest_coeff sign t |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
110 |
val (q, u') = dest_coeff 1 u |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
111 |
in (mk_frac (p, q), mk_divide (t', u')) end |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
112 |
| dest_fcoeff sign t = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
113 |
let val (p, t') = dest_coeff sign t |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
114 |
val T = Term.fastype_of t |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
115 |
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:
31067
diff
changeset
|
116 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
117 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
118 |
(** 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:
31067
diff
changeset
|
119 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
120 |
(*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:
31067
diff
changeset
|
121 |
ordering is not well-founded.*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
122 |
fun num_ord (i,j) = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
123 |
(case int_ord (abs i, abs j) of |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
124 |
EQUAL => int_ord (Int.sign i, Int.sign j) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
125 |
| ord => ord); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
126 |
|
| 35408 | 127 |
(*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:
31067
diff
changeset
|
128 |
non-atomic terms.*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
129 |
local open Term |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
130 |
in |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
131 |
fun numterm_ord (Abs (_, T, t), Abs(_, U, u)) = |
| 35408 | 132 |
(case numterm_ord (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord) |
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
133 |
| numterm_ord |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
134 |
(Const(@{const_name Int.number_of}, _) $ v, Const(@{const_name Int.number_of}, _) $ w) =
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
135 |
num_ord (HOLogic.dest_numeral v, HOLogic.dest_numeral w) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
136 |
| numterm_ord (Const(@{const_name Int.number_of}, _) $ _, _) = LESS
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
137 |
| numterm_ord (_, Const(@{const_name Int.number_of}, _) $ _) = GREATER
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
138 |
| numterm_ord (t, u) = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
139 |
(case int_ord (size_of_term t, size_of_term u) of |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
140 |
EQUAL => |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
141 |
let val (f, ts) = strip_comb t and (g, us) = strip_comb u in |
| 35408 | 142 |
(case Term_Ord.hd_ord (f, g) of EQUAL => numterms_ord (ts, us) | ord => ord) |
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
143 |
end |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
144 |
| ord => ord) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
145 |
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:
31067
diff
changeset
|
146 |
end; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
147 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
148 |
fun numtermless tu = (numterm_ord tu = LESS); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
149 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
150 |
val num_ss = HOL_ss settermless numtermless; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
151 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
152 |
(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic isn't complicated by the abstract 0 and 1.*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
153 |
val numeral_syms = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym];
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
154 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
155 |
(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
156 |
val add_0s = @{thms add_0s};
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
157 |
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:
31067
diff
changeset
|
158 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
159 |
(*Simplify inverse Numeral1, a/Numeral1*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
160 |
val inverse_1s = [@{thm inverse_numeral_1}];
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
161 |
val divide_1s = [@{thm divide_numeral_1}];
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
162 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
163 |
(*To perform binary arithmetic. The "left" rewriting handles patterns |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
164 |
created by the Numeral_Simprocs, such as 3 * (5 * x). *) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
165 |
val simps = [@{thm numeral_0_eq_0} RS sym, @{thm numeral_1_eq_1} RS sym,
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
166 |
@{thm add_number_of_left}, @{thm mult_number_of_left}] @
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
167 |
@{thms arith_simps} @ @{thms rel_simps};
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
168 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
169 |
(*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:
31067
diff
changeset
|
170 |
during re-arrangement*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
171 |
val non_add_simps = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
172 |
subtract Thm.eq_thm [@{thm add_number_of_left}, @{thm number_of_add} RS sym] simps;
|
| 23164 | 173 |
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
174 |
(*To evaluate binary negations of coefficients*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
175 |
val minus_simps = [@{thm numeral_m1_eq_minus_1} RS sym, @{thm number_of_minus} RS sym] @
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
176 |
@{thms minus_bin_simps} @ @{thms pred_bin_simps};
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
177 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
178 |
(*To let us treat subtraction as addition*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
179 |
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:
31067
diff
changeset
|
180 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
181 |
(*To let us treat division as multiplication*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
182 |
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:
31067
diff
changeset
|
183 |
|
|
35020
862a20ffa8e2
prefer explicit @{lemma} over adhoc forward reasoning;
wenzelm
parents:
34974
diff
changeset
|
184 |
(*push the unary minus down*) |
|
862a20ffa8e2
prefer explicit @{lemma} over adhoc forward reasoning;
wenzelm
parents:
34974
diff
changeset
|
185 |
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:
31067
diff
changeset
|
186 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
187 |
(*to extract again any uncancelled minuses*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
188 |
val minus_from_mult_simps = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
189 |
[@{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:
31067
diff
changeset
|
190 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
191 |
(*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:
31067
diff
changeset
|
192 |
val mult_minus_simps = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
193 |
[@{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:
31067
diff
changeset
|
194 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
195 |
val norm_ss1 = num_ss addsimps numeral_syms @ add_0s @ mult_1s @ |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
196 |
diff_simps @ minus_simps @ @{thms add_ac}
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
197 |
val norm_ss2 = num_ss addsimps non_add_simps @ mult_minus_simps |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
198 |
val norm_ss3 = num_ss addsimps minus_from_mult_simps @ @{thms add_ac} @ @{thms mult_ac}
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
199 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
200 |
structure CancelNumeralsCommon = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
201 |
struct |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
202 |
val mk_sum = mk_sum |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
203 |
val dest_sum = dest_sum |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
204 |
val mk_coeff = mk_coeff |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
205 |
val dest_coeff = dest_coeff 1 |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
206 |
val find_first_coeff = find_first_coeff [] |
| 31368 | 207 |
fun trans_tac _ = Arith_Data.trans_tac |
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
208 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
209 |
fun norm_tac ss = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
210 |
ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
211 |
THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
212 |
THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
213 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
214 |
val numeral_simp_ss = HOL_ss addsimps add_0s @ simps |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
215 |
fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
216 |
val simplify_meta_eq = Arith_Data.simplify_meta_eq (add_0s @ mult_1s) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
217 |
end; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
218 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
219 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
220 |
structure EqCancelNumerals = CancelNumeralsFun |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
221 |
(open CancelNumeralsCommon |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
222 |
val prove_conv = Arith_Data.prove_conv |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
223 |
val mk_bal = HOLogic.mk_eq |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
224 |
val dest_bal = HOLogic.dest_bin "op =" Term.dummyT |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
225 |
val bal_add1 = @{thm eq_add_iff1} RS trans
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
226 |
val bal_add2 = @{thm eq_add_iff2} RS trans
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
227 |
); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
228 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
229 |
structure LessCancelNumerals = CancelNumeralsFun |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
230 |
(open CancelNumeralsCommon |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
231 |
val prove_conv = Arith_Data.prove_conv |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
232 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less}
|
|
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
233 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} Term.dummyT
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
234 |
val bal_add1 = @{thm less_add_iff1} RS trans
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
235 |
val bal_add2 = @{thm less_add_iff2} RS trans
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
236 |
); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
237 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
238 |
structure LeCancelNumerals = CancelNumeralsFun |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
239 |
(open CancelNumeralsCommon |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
240 |
val prove_conv = Arith_Data.prove_conv |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
241 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less_eq}
|
|
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
242 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} Term.dummyT
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
243 |
val bal_add1 = @{thm le_add_iff1} RS trans
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
244 |
val bal_add2 = @{thm le_add_iff2} RS trans
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
245 |
); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
246 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
247 |
val cancel_numerals = |
| 32155 | 248 |
map (Arith_Data.prep_simproc @{theory})
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
249 |
[("inteq_cancel_numerals",
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
250 |
["(l::'a::number_ring) + m = n", |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
251 |
"(l::'a::number_ring) = m + n", |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
252 |
"(l::'a::number_ring) - m = n", |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
253 |
"(l::'a::number_ring) = m - n", |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
254 |
"(l::'a::number_ring) * m = n", |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
255 |
"(l::'a::number_ring) = m * n"], |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
256 |
K EqCancelNumerals.proc), |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
257 |
("intless_cancel_numerals",
|
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
258 |
["(l::'a::{linordered_idom,number_ring}) + m < n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
259 |
"(l::'a::{linordered_idom,number_ring}) < m + n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
260 |
"(l::'a::{linordered_idom,number_ring}) - m < n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
261 |
"(l::'a::{linordered_idom,number_ring}) < m - n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
262 |
"(l::'a::{linordered_idom,number_ring}) * m < n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
263 |
"(l::'a::{linordered_idom,number_ring}) < m * n"],
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
264 |
K LessCancelNumerals.proc), |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
265 |
("intle_cancel_numerals",
|
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
266 |
["(l::'a::{linordered_idom,number_ring}) + m <= n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
267 |
"(l::'a::{linordered_idom,number_ring}) <= m + n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
268 |
"(l::'a::{linordered_idom,number_ring}) - m <= n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
269 |
"(l::'a::{linordered_idom,number_ring}) <= m - n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
270 |
"(l::'a::{linordered_idom,number_ring}) * m <= n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
271 |
"(l::'a::{linordered_idom,number_ring}) <= m * n"],
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
272 |
K LeCancelNumerals.proc)]; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
273 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
274 |
structure CombineNumeralsData = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
275 |
struct |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
276 |
type coeff = int |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
277 |
val iszero = (fn x => x = 0) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
278 |
val add = op + |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
279 |
val mk_sum = long_mk_sum (*to work for e.g. 2*x + 3*x *) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
280 |
val dest_sum = dest_sum |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
281 |
val mk_coeff = mk_coeff |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
282 |
val dest_coeff = dest_coeff 1 |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
283 |
val left_distrib = @{thm combine_common_factor} RS trans
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
284 |
val prove_conv = Arith_Data.prove_conv_nohyps |
| 31368 | 285 |
fun trans_tac _ = Arith_Data.trans_tac |
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
286 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
287 |
fun norm_tac ss = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
288 |
ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
289 |
THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
290 |
THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
291 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
292 |
val numeral_simp_ss = HOL_ss addsimps add_0s @ simps |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
293 |
fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
294 |
val simplify_meta_eq = Arith_Data.simplify_meta_eq (add_0s @ mult_1s) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
295 |
end; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
296 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
297 |
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
298 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
299 |
(*Version for fields, where coefficients can be fractions*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
300 |
structure FieldCombineNumeralsData = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
301 |
struct |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
302 |
type coeff = int * int |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
303 |
val iszero = (fn (p, q) => p = 0) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
304 |
val add = add_frac |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
305 |
val mk_sum = long_mk_sum |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
306 |
val dest_sum = dest_sum |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
307 |
val mk_coeff = mk_fcoeff |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
308 |
val dest_coeff = dest_fcoeff 1 |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
309 |
val left_distrib = @{thm combine_common_factor} RS trans
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
310 |
val prove_conv = Arith_Data.prove_conv_nohyps |
| 31368 | 311 |
fun trans_tac _ = Arith_Data.trans_tac |
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
312 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
313 |
val norm_ss1a = norm_ss1 addsimps inverse_1s @ divide_simps |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
314 |
fun norm_tac ss = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
315 |
ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1a)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
316 |
THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
317 |
THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
318 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
319 |
val numeral_simp_ss = HOL_ss addsimps add_0s @ simps @ [@{thm add_frac_eq}]
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
320 |
fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
321 |
val simplify_meta_eq = Arith_Data.simplify_meta_eq (add_0s @ mult_1s @ divide_1s) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
322 |
end; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
323 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
324 |
structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
325 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
326 |
val combine_numerals = |
| 32155 | 327 |
Arith_Data.prep_simproc @{theory}
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
328 |
("int_combine_numerals",
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
329 |
["(i::'a::number_ring) + j", "(i::'a::number_ring) - j"], |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
330 |
K CombineNumerals.proc); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
331 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
332 |
val field_combine_numerals = |
| 32155 | 333 |
Arith_Data.prep_simproc @{theory}
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
334 |
("field_combine_numerals",
|
| 36409 | 335 |
["(i::'a::{field_inverse_zero, number_ring}) + j",
|
336 |
"(i::'a::{field_inverse_zero, number_ring}) - j"],
|
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
337 |
K FieldCombineNumerals.proc); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
338 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
339 |
(** Constant folding for multiplication in semirings **) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
340 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
341 |
(*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:
31067
diff
changeset
|
342 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
343 |
structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
344 |
struct |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
345 |
val assoc_ss = HOL_ss addsimps @{thms mult_ac}
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
346 |
val eq_reflection = eq_reflection |
|
35983
27e2fa7d4ce7
slightly more general simproc (avoids errors of linarith)
boehmes
parents:
35408
diff
changeset
|
347 |
val is_numeral = can HOLogic.dest_number |
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
348 |
end; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
349 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
350 |
structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
351 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
352 |
val assoc_fold_simproc = |
| 32155 | 353 |
Arith_Data.prep_simproc @{theory}
|
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
354 |
("semiring_assoc_fold", ["(a::'a::comm_semiring_1_cancel) * b"],
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
355 |
K Semiring_Times_Assoc.proc); |
| 23164 | 356 |
|
357 |
structure CancelNumeralFactorCommon = |
|
358 |
struct |
|
359 |
val mk_coeff = mk_coeff |
|
360 |
val dest_coeff = dest_coeff 1 |
|
| 31368 | 361 |
fun trans_tac _ = Arith_Data.trans_tac |
| 23164 | 362 |
|
363 |
val norm_ss1 = HOL_ss addsimps minus_from_mult_simps @ mult_1s |
|
364 |
val norm_ss2 = HOL_ss addsimps simps @ mult_minus_simps |
|
| 23881 | 365 |
val norm_ss3 = HOL_ss addsimps @{thms mult_ac}
|
| 23164 | 366 |
fun norm_tac ss = |
367 |
ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss1)) |
|
368 |
THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss2)) |
|
369 |
THEN ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss3)) |
|
370 |
||
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
371 |
val numeral_simp_ss = HOL_ss addsimps |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
372 |
[@{thm eq_number_of_eq}, @{thm less_number_of}, @{thm le_number_of}] @ simps
|
| 23164 | 373 |
fun numeral_simp_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss numeral_simp_ss)) |
| 30518 | 374 |
val simplify_meta_eq = Arith_Data.simplify_meta_eq |
|
35064
1bdef0c013d3
hide fact names clashing with fact names from Group.thy
haftmann
parents:
35050
diff
changeset
|
375 |
[@{thm Nat.add_0}, @{thm Nat.add_0_right}, @{thm mult_zero_left},
|
|
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
25481
diff
changeset
|
376 |
@{thm mult_zero_right}, @{thm mult_Bit1}, @{thm mult_1_right}];
|
| 23164 | 377 |
end |
378 |
||
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
379 |
(*Version for semiring_div*) |
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
380 |
structure DivCancelNumeralFactor = CancelNumeralFactorFun |
| 23164 | 381 |
(open CancelNumeralFactorCommon |
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
382 |
val prove_conv = Arith_Data.prove_conv |
| 23164 | 383 |
val mk_bal = HOLogic.mk_binop @{const_name Divides.div}
|
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
384 |
val dest_bal = HOLogic.dest_bin @{const_name Divides.div} Term.dummyT
|
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
385 |
val cancel = @{thm div_mult_mult1} RS trans
|
| 23164 | 386 |
val neg_exchanges = false |
387 |
) |
|
388 |
||
389 |
(*Version for fields*) |
|
390 |
structure DivideCancelNumeralFactor = CancelNumeralFactorFun |
|
391 |
(open CancelNumeralFactorCommon |
|
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
392 |
val prove_conv = Arith_Data.prove_conv |
| 35084 | 393 |
val mk_bal = HOLogic.mk_binop @{const_name Rings.divide}
|
394 |
val dest_bal = HOLogic.dest_bin @{const_name Rings.divide} Term.dummyT
|
|
|
23413
5caa2710dd5b
tuned laws for cancellation in divisions for fields.
nipkow
parents:
23401
diff
changeset
|
395 |
val cancel = @{thm mult_divide_mult_cancel_left} RS trans
|
| 23164 | 396 |
val neg_exchanges = false |
397 |
) |
|
398 |
||
399 |
structure EqCancelNumeralFactor = CancelNumeralFactorFun |
|
400 |
(open CancelNumeralFactorCommon |
|
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
401 |
val prove_conv = Arith_Data.prove_conv |
| 23164 | 402 |
val mk_bal = HOLogic.mk_eq |
403 |
val dest_bal = HOLogic.dest_bin "op =" Term.dummyT |
|
404 |
val cancel = @{thm mult_cancel_left} RS trans
|
|
405 |
val neg_exchanges = false |
|
406 |
) |
|
407 |
||
408 |
structure LessCancelNumeralFactor = CancelNumeralFactorFun |
|
409 |
(open CancelNumeralFactorCommon |
|
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
410 |
val prove_conv = Arith_Data.prove_conv |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
411 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less}
|
|
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
412 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} Term.dummyT
|
| 23164 | 413 |
val cancel = @{thm mult_less_cancel_left} RS trans
|
414 |
val neg_exchanges = true |
|
415 |
) |
|
416 |
||
417 |
structure LeCancelNumeralFactor = CancelNumeralFactorFun |
|
418 |
(open CancelNumeralFactorCommon |
|
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
419 |
val prove_conv = Arith_Data.prove_conv |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
420 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less_eq}
|
|
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
421 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} Term.dummyT
|
| 23164 | 422 |
val cancel = @{thm mult_le_cancel_left} RS trans
|
423 |
val neg_exchanges = true |
|
424 |
) |
|
425 |
||
426 |
val cancel_numeral_factors = |
|
| 32155 | 427 |
map (Arith_Data.prep_simproc @{theory})
|
| 23164 | 428 |
[("ring_eq_cancel_numeral_factor",
|
429 |
["(l::'a::{idom,number_ring}) * m = n",
|
|
430 |
"(l::'a::{idom,number_ring}) = m * n"],
|
|
431 |
K EqCancelNumeralFactor.proc), |
|
432 |
("ring_less_cancel_numeral_factor",
|
|
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
433 |
["(l::'a::{linordered_idom,number_ring}) * m < n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
434 |
"(l::'a::{linordered_idom,number_ring}) < m * n"],
|
| 23164 | 435 |
K LessCancelNumeralFactor.proc), |
436 |
("ring_le_cancel_numeral_factor",
|
|
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
437 |
["(l::'a::{linordered_idom,number_ring}) * m <= n",
|
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
438 |
"(l::'a::{linordered_idom,number_ring}) <= m * n"],
|
| 23164 | 439 |
K LeCancelNumeralFactor.proc), |
440 |
("int_div_cancel_numeral_factors",
|
|
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
441 |
["((l::'a::{semiring_div,number_ring}) * m) div n",
|
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
442 |
"(l::'a::{semiring_div,number_ring}) div (m * n)"],
|
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
443 |
K DivCancelNumeralFactor.proc), |
| 23164 | 444 |
("divide_cancel_numeral_factor",
|
| 36409 | 445 |
["((l::'a::{field_inverse_zero,number_ring}) * m) / n",
|
446 |
"(l::'a::{field_inverse_zero,number_ring}) / (m * n)",
|
|
447 |
"((number_of v)::'a::{field_inverse_zero,number_ring}) / (number_of w)"],
|
|
| 23164 | 448 |
K DivideCancelNumeralFactor.proc)]; |
449 |
||
450 |
val field_cancel_numeral_factors = |
|
| 32155 | 451 |
map (Arith_Data.prep_simproc @{theory})
|
| 23164 | 452 |
[("field_eq_cancel_numeral_factor",
|
453 |
["(l::'a::{field,number_ring}) * m = n",
|
|
454 |
"(l::'a::{field,number_ring}) = m * n"],
|
|
455 |
K EqCancelNumeralFactor.proc), |
|
456 |
("field_cancel_numeral_factor",
|
|
| 36409 | 457 |
["((l::'a::{field_inverse_zero,number_ring}) * m) / n",
|
458 |
"(l::'a::{field_inverse_zero,number_ring}) / (m * n)",
|
|
459 |
"((number_of v)::'a::{field_inverse_zero,number_ring}) / (number_of w)"],
|
|
| 23164 | 460 |
K DivideCancelNumeralFactor.proc)] |
461 |
||
462 |
||
463 |
(** Declarations for ExtractCommonTerm **) |
|
464 |
||
465 |
(*Find first term that matches u*) |
|
466 |
fun find_first_t past u [] = raise TERM ("find_first_t", [])
|
|
467 |
| find_first_t past u (t::terms) = |
|
468 |
if u aconv t then (rev past @ terms) |
|
469 |
else find_first_t (t::past) u terms |
|
470 |
handle TERM _ => find_first_t (t::past) u terms; |
|
471 |
||
472 |
(** Final simplification for the CancelFactor simprocs **) |
|
| 30518 | 473 |
val simplify_one = Arith_Data.simplify_meta_eq |
| 30031 | 474 |
[@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_by_1}, @{thm numeral_1_eq_1}];
|
| 23164 | 475 |
|
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
476 |
fun cancel_simplify_meta_eq ss cancel_th th = |
| 23164 | 477 |
simplify_one ss (([th, cancel_th]) MRS trans); |
478 |
||
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
479 |
local |
| 31067 | 480 |
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:
30518
diff
changeset
|
481 |
fun Eq_True_elim Eq = |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
482 |
Thm.equal_elim (Thm.combination Tp_Eq (Thm.symmetric Eq)) @{thm TrueI}
|
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
483 |
in |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
484 |
fun sign_conv pos_th neg_th ss t = |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
485 |
let val T = fastype_of t; |
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
35092
diff
changeset
|
486 |
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:
35084
diff
changeset
|
487 |
val less = Const(@{const_name Orderings.less}, [T,T] ---> HOLogic.boolT);
|
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
488 |
val pos = less $ zero $ t and neg = less $ t $ zero |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
489 |
fun prove p = |
|
31101
26c7bb764a38
qualified names for Lin_Arith tactics and simprocs
haftmann
parents:
31068
diff
changeset
|
490 |
Option.map Eq_True_elim (Lin_Arith.simproc ss p) |
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
491 |
handle THM _ => NONE |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
492 |
in case prove pos of |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
493 |
SOME th => SOME(th RS pos_th) |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
494 |
| NONE => (case prove neg of |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
495 |
SOME th => SOME(th RS neg_th) |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
496 |
| NONE => NONE) |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
497 |
end; |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
498 |
end |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
499 |
|
| 23164 | 500 |
structure CancelFactorCommon = |
501 |
struct |
|
502 |
val mk_sum = long_mk_prod |
|
503 |
val dest_sum = dest_prod |
|
504 |
val mk_coeff = mk_coeff |
|
505 |
val dest_coeff = dest_coeff |
|
506 |
val find_first = find_first_t [] |
|
| 31368 | 507 |
fun trans_tac _ = Arith_Data.trans_tac |
| 23881 | 508 |
val norm_ss = HOL_ss addsimps mult_1s @ @{thms mult_ac}
|
| 23164 | 509 |
fun norm_tac ss = ALLGOALS (simp_tac (Simplifier.inherit_context ss norm_ss)) |
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
510 |
val simplify_meta_eq = cancel_simplify_meta_eq |
| 23164 | 511 |
end; |
512 |
||
513 |
(*mult_cancel_left requires a ring with no zero divisors.*) |
|
514 |
structure EqCancelFactor = ExtractCommonTermFun |
|
515 |
(open CancelFactorCommon |
|
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
516 |
val prove_conv = Arith_Data.prove_conv |
| 23164 | 517 |
val mk_bal = HOLogic.mk_eq |
518 |
val dest_bal = HOLogic.dest_bin "op =" Term.dummyT |
|
| 31368 | 519 |
fun simp_conv _ _ = SOME @{thm mult_cancel_left}
|
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
520 |
); |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
521 |
|
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
522 |
(*for ordered rings*) |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
523 |
structure LeCancelFactor = ExtractCommonTermFun |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
524 |
(open CancelFactorCommon |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
525 |
val prove_conv = Arith_Data.prove_conv |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
526 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less_eq}
|
|
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
527 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less_eq} Term.dummyT
|
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
528 |
val simp_conv = sign_conv |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
529 |
@{thm mult_le_cancel_left_pos} @{thm mult_le_cancel_left_neg}
|
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
530 |
); |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
531 |
|
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
532 |
(*for ordered rings*) |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
533 |
structure LessCancelFactor = ExtractCommonTermFun |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
534 |
(open CancelFactorCommon |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
535 |
val prove_conv = Arith_Data.prove_conv |
|
35092
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
536 |
val mk_bal = HOLogic.mk_binrel @{const_name Orderings.less}
|
|
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents:
35084
diff
changeset
|
537 |
val dest_bal = HOLogic.dest_bin @{const_name Orderings.less} Term.dummyT
|
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
538 |
val simp_conv = sign_conv |
|
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
539 |
@{thm mult_less_cancel_left_pos} @{thm mult_less_cancel_left_neg}
|
| 23164 | 540 |
); |
541 |
||
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
542 |
(*for semirings with division*) |
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
543 |
structure DivCancelFactor = ExtractCommonTermFun |
| 23164 | 544 |
(open CancelFactorCommon |
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
545 |
val prove_conv = Arith_Data.prove_conv |
| 23164 | 546 |
val mk_bal = HOLogic.mk_binop @{const_name Divides.div}
|
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
547 |
val dest_bal = HOLogic.dest_bin @{const_name Divides.div} Term.dummyT
|
| 31368 | 548 |
fun simp_conv _ _ = SOME @{thm div_mult_mult1_if}
|
| 23164 | 549 |
); |
550 |
||
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
551 |
structure ModCancelFactor = ExtractCommonTermFun |
| 24395 | 552 |
(open CancelFactorCommon |
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
553 |
val prove_conv = Arith_Data.prove_conv |
| 24395 | 554 |
val mk_bal = HOLogic.mk_binop @{const_name Divides.mod}
|
| 31067 | 555 |
val dest_bal = HOLogic.dest_bin @{const_name Divides.mod} Term.dummyT
|
| 31368 | 556 |
fun simp_conv _ _ = SOME @{thm mod_mult_mult1}
|
| 24395 | 557 |
); |
558 |
||
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
559 |
(*for idoms*) |
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
560 |
structure DvdCancelFactor = ExtractCommonTermFun |
| 23969 | 561 |
(open CancelFactorCommon |
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
562 |
val prove_conv = Arith_Data.prove_conv |
|
35050
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents:
35043
diff
changeset
|
563 |
val mk_bal = HOLogic.mk_binrel @{const_name Rings.dvd}
|
|
9f841f20dca6
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents:
35043
diff
changeset
|
564 |
val dest_bal = HOLogic.dest_bin @{const_name Rings.dvd} Term.dummyT
|
| 31368 | 565 |
fun simp_conv _ _ = SOME @{thm dvd_mult_cancel_left}
|
| 23969 | 566 |
); |
567 |
||
| 23164 | 568 |
(*Version for all fields, including unordered ones (type complex).*) |
569 |
structure DivideCancelFactor = ExtractCommonTermFun |
|
570 |
(open CancelFactorCommon |
|
|
30496
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents:
30242
diff
changeset
|
571 |
val prove_conv = Arith_Data.prove_conv |
| 35084 | 572 |
val mk_bal = HOLogic.mk_binop @{const_name Rings.divide}
|
573 |
val dest_bal = HOLogic.dest_bin @{const_name Rings.divide} Term.dummyT
|
|
| 31368 | 574 |
fun simp_conv _ _ = SOME @{thm mult_divide_mult_cancel_left_if}
|
| 23164 | 575 |
); |
576 |
||
577 |
val cancel_factors = |
|
| 32155 | 578 |
map (Arith_Data.prep_simproc @{theory})
|
| 23164 | 579 |
[("ring_eq_cancel_factor",
|
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
580 |
["(l::'a::idom) * m = n", |
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
581 |
"(l::'a::idom) = m * n"], |
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
582 |
K EqCancelFactor.proc), |
|
35043
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents:
35030
diff
changeset
|
583 |
("linordered_ring_le_cancel_factor",
|
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
584 |
["(l::'a::linordered_ring) * m <= n", |
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
585 |
"(l::'a::linordered_ring) <= m * n"], |
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
586 |
K LeCancelFactor.proc), |
|
35043
07dbdf60d5ad
dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents:
35030
diff
changeset
|
587 |
("linordered_ring_less_cancel_factor",
|
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
588 |
["(l::'a::linordered_ring) * m < n", |
|
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
34974
diff
changeset
|
589 |
"(l::'a::linordered_ring) < m * n"], |
|
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
30518
diff
changeset
|
590 |
K LessCancelFactor.proc), |
| 23164 | 591 |
("int_div_cancel_factor",
|
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
592 |
["((l::'a::semiring_div) * m) div n", "(l::'a::semiring_div) div (m * n)"], |
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
593 |
K DivCancelFactor.proc), |
| 24395 | 594 |
("int_mod_cancel_factor",
|
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
595 |
["((l::'a::semiring_div) * m) mod n", "(l::'a::semiring_div) mod (m * n)"], |
|
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
596 |
K ModCancelFactor.proc), |
|
29981
7d0ed261b712
generalize int_dvd_cancel_factor simproc to idom class
huffman
parents:
29038
diff
changeset
|
597 |
("dvd_cancel_factor",
|
|
7d0ed261b712
generalize int_dvd_cancel_factor simproc to idom class
huffman
parents:
29038
diff
changeset
|
598 |
["((l::'a::idom) * m) dvd n", "(l::'a::idom) dvd (m * n)"], |
|
30931
86ca651da03e
generalized some simprocs from int to semiring_div
haftmann
parents:
30685
diff
changeset
|
599 |
K DvdCancelFactor.proc), |
| 23164 | 600 |
("divide_cancel_factor",
|
| 36409 | 601 |
["((l::'a::field_inverse_zero) * m) / n", |
602 |
"(l::'a::field_inverse_zero) / (m * n)"], |
|
| 23164 | 603 |
K DivideCancelFactor.proc)]; |
604 |
||
605 |
end; |
|
606 |
||
|
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
607 |
Addsimprocs Numeral_Simprocs.cancel_numerals; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
608 |
Addsimprocs [Numeral_Simprocs.combine_numerals]; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
609 |
Addsimprocs [Numeral_Simprocs.field_combine_numerals]; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
610 |
Addsimprocs [Numeral_Simprocs.assoc_fold_simproc]; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
611 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
612 |
(*examples: |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
613 |
print_depth 22; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
614 |
set timing; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
615 |
set trace_simp; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
616 |
fun test s = (Goal s, by (Simp_tac 1)); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
617 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
618 |
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:
31067
diff
changeset
|
619 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
620 |
test "2*u = (u::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
621 |
test "(i + j + 12 + (k::int)) - 15 = y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
622 |
test "(i + j + 12 + (k::int)) - 5 = y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
623 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
624 |
test "y - b < (b::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
625 |
test "y - (3*b + c) < (b::int) - 2*c"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
626 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
627 |
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:
31067
diff
changeset
|
628 |
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:
31067
diff
changeset
|
629 |
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:
31067
diff
changeset
|
630 |
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:
31067
diff
changeset
|
631 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
632 |
test "(i + j + 12 + (k::int)) = u + 15 + y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
633 |
test "(i + j*2 + 12 + (k::int)) = j + 5 + y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
634 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
635 |
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:
31067
diff
changeset
|
636 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
637 |
test "a + -(b+c) + b = (d::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
638 |
test "a + -(b+c) - b = (d::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
639 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
640 |
(*negative numerals*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
641 |
test "(i + j + -2 + (k::int)) - (u + 5 + y) = zz"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
642 |
test "(i + j + -3 + (k::int)) < u + 5 + y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
643 |
test "(i + j + 3 + (k::int)) < u + -6 + y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
644 |
test "(i + j + -12 + (k::int)) - 15 = y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
645 |
test "(i + j + 12 + (k::int)) - -15 = y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
646 |
test "(i + j + -12 + (k::int)) - -15 = y"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
647 |
*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
648 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
649 |
Addsimprocs Numeral_Simprocs.cancel_numeral_factors; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
650 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
651 |
(*examples: |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
652 |
print_depth 22; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
653 |
set timing; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
654 |
set trace_simp; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
655 |
fun test s = (Goal s; by (Simp_tac 1)); |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
656 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
657 |
test "9*x = 12 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
658 |
test "(9*x) div (12 * (y::int)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
659 |
test "9*x < 12 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
660 |
test "9*x <= 12 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
661 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
662 |
test "-99*x = 132 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
663 |
test "(-99*x) div (132 * (y::int)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
664 |
test "-99*x < 132 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
665 |
test "-99*x <= 132 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
666 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
667 |
test "999*x = -396 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
668 |
test "(999*x) div (-396 * (y::int)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
669 |
test "999*x < -396 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
670 |
test "999*x <= -396 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
671 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
672 |
test "-99*x = -81 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
673 |
test "(-99*x) div (-81 * (y::int)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
674 |
test "-99*x <= -81 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
675 |
test "-99*x < -81 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
676 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
677 |
test "-2 * x = -1 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
678 |
test "-2 * x = -(y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
679 |
test "(-2 * x) div (-1 * (y::int)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
680 |
test "-2 * x < -(y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
681 |
test "-2 * x <= -1 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
682 |
test "-x < -23 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
683 |
test "-x <= -23 * (y::int)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
684 |
*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
685 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
686 |
(*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:
31067
diff
changeset
|
687 |
test "0 <= (y::rat) * -2"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
688 |
test "9*x = 12 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
689 |
test "(9*x) / (12 * (y::rat)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
690 |
test "9*x < 12 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
691 |
test "9*x <= 12 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
692 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
693 |
test "-99*x = 132 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
694 |
test "(-99*x) / (132 * (y::rat)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
695 |
test "-99*x < 132 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
696 |
test "-99*x <= 132 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
697 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
698 |
test "999*x = -396 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
699 |
test "(999*x) / (-396 * (y::rat)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
700 |
test "999*x < -396 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
701 |
test "999*x <= -396 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
702 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
703 |
test "(- ((2::rat) * x) <= 2 * y)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
704 |
test "-99*x = -81 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
705 |
test "(-99*x) / (-81 * (y::rat)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
706 |
test "-99*x <= -81 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
707 |
test "-99*x < -81 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
708 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
709 |
test "-2 * x = -1 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
710 |
test "-2 * x = -(y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
711 |
test "(-2 * x) / (-1 * (y::rat)) = z"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
712 |
test "-2 * x < -(y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
713 |
test "-2 * x <= -1 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
714 |
test "-x < -23 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
715 |
test "-x <= -23 * (y::rat)"; |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
716 |
*) |
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
717 |
|
|
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31067
diff
changeset
|
718 |
Addsimprocs Numeral_Simprocs.cancel_factors; |
| 23164 | 719 |
|
720 |
||
721 |
(*examples: |
|
722 |
print_depth 22; |
|
723 |
set timing; |
|
724 |
set trace_simp; |
|
725 |
fun test s = (Goal s; by (Asm_simp_tac 1)); |
|
726 |
||
727 |
test "x*k = k*(y::int)"; |
|
728 |
test "k = k*(y::int)"; |
|
729 |
test "a*(b*c) = (b::int)"; |
|
730 |
test "a*(b*c) = d*(b::int)*(x*a)"; |
|
731 |
||
732 |
test "(x*k) div (k*(y::int)) = (uu::int)"; |
|
733 |
test "(k) div (k*(y::int)) = (uu::int)"; |
|
734 |
test "(a*(b*c)) div ((b::int)) = (uu::int)"; |
|
735 |
test "(a*(b*c)) div (d*(b::int)*(x*a)) = (uu::int)"; |
|
736 |
*) |
|
737 |
||
738 |
(*And the same examples for fields such as rat or real: |
|
739 |
print_depth 22; |
|
740 |
set timing; |
|
741 |
set trace_simp; |
|
742 |
fun test s = (Goal s; by (Asm_simp_tac 1)); |
|
743 |
||
744 |
test "x*k = k*(y::rat)"; |
|
745 |
test "k = k*(y::rat)"; |
|
746 |
test "a*(b*c) = (b::rat)"; |
|
747 |
test "a*(b*c) = d*(b::rat)*(x*a)"; |
|
748 |
||
749 |
||
750 |
test "(x*k) / (k*(y::rat)) = (uu::rat)"; |
|
751 |
test "(k) / (k*(y::rat)) = (uu::rat)"; |
|
752 |
test "(a*(b*c)) / ((b::rat)) = (uu::rat)"; |
|
753 |
test "(a*(b*c)) / (d*(b::rat)*(x*a)) = (uu::rat)"; |
|
754 |
||
755 |
(*FIXME: what do we do about this?*) |
|
756 |
test "a*(b*c)/(y*z) = d*(b::rat)*(x*a)/z"; |
|
757 |
*) |