| author | haftmann | 
| Tue, 22 Nov 2005 12:59:25 +0100 | |
| changeset 18220 | 43cf5767f992 | 
| parent 17877 | 67d5ab1cb0d8 | 
| child 18328 | 841261f303a1 | 
| permissions | -rw-r--r-- | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/nat_simprocs.ML  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
Copyright 2000 University of Cambridge  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
Simprocs for nat numerals.  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 
14430
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
9  | 
val Let_number_of = thm"Let_number_of";  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
10  | 
val Let_0 = thm"Let_0";  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
11  | 
val Let_1 = thm"Let_1";  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
12  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
structure Nat_Numeral_Simprocs =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
|
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
16  | 
(*Maps n to #n for n = 0, 1, 2*)  | 
| 
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
17  | 
val numeral_syms =  | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14369 
diff
changeset
 | 
18  | 
[nat_numeral_0_eq_0 RS sym, nat_numeral_1_eq_1 RS sym, numeral_2_eq_2 RS sym];  | 
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
19  | 
val numeral_sym_ss = HOL_ss addsimps numeral_syms;  | 
| 
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
20  | 
|
| 13462 | 21  | 
fun rename_numerals th =  | 
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
22  | 
simplify numeral_sym_ss (Thm.transfer (the_context ()) th);  | 
| 
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
23  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
(*Utilities*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
|
| 10693 | 26  | 
fun mk_numeral n = HOLogic.number_of_const HOLogic.natT $ HOLogic.mk_bin n;  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
(*Decodes a unary or binary numeral to a NATURAL NUMBER*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
fun dest_numeral (Const ("0", _)) = 0
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
  | dest_numeral (Const ("Suc", _) $ t) = 1 + dest_numeral t
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
  | dest_numeral (Const("Numeral.number_of", _) $ w) =
 | 
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15921 
diff
changeset
 | 
32  | 
(IntInf.max (0, HOLogic.dest_binum w)  | 
| 10890 | 33  | 
       handle TERM _ => raise TERM("Nat_Numeral_Simprocs.dest_numeral:1", [w]))
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
  | dest_numeral t = raise TERM("Nat_Numeral_Simprocs.dest_numeral:2", [t]);
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
fun find_first_numeral past (t::terms) =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
((dest_numeral t, t, rev past @ terms)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
handle TERM _ => find_first_numeral (t::past) terms)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
  | find_first_numeral past [] = raise TERM("find_first_numeral", []);
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
val zero = mk_numeral 0;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
val mk_plus = HOLogic.mk_binop "op +";  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
|
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
44  | 
(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
fun mk_sum [] = zero  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
| mk_sum [t,u] = mk_plus (t, u)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
| mk_sum (t :: ts) = mk_plus (t, mk_sum ts);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
(*this version ALWAYS includes a trailing zero*)  | 
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
50  | 
fun long_mk_sum [] = HOLogic.zero  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
| long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
val dest_plus = HOLogic.dest_bin "op +" HOLogic.natT;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
(*extract the outer Sucs from a term and convert them to a binary numeral*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
fun dest_Sucs (k, Const ("Suc", _) $ t) = dest_Sucs (k+1, t)
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
| dest_Sucs (0, t) = t  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
| dest_Sucs (k, t) = mk_plus (mk_numeral k, t);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
fun dest_sum t =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
let val (t,u) = dest_plus t  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
in dest_sum t @ dest_sum u end  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
handle TERM _ => [t];  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
fun dest_Sucs_sum t = dest_sum (dest_Sucs (0,t));  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
|
| 
11377
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
67  | 
|
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
68  | 
(** Other simproc items **)  | 
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
69  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
val trans_tac = Int_Numeral_Simprocs.trans_tac;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
|
| 
14430
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
72  | 
val bin_simps =  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
73  | 
[nat_numeral_0_eq_0 RS sym, nat_numeral_1_eq_1 RS sym,  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
74  | 
add_nat_number_of, nat_number_of_add_left,  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
75  | 
diff_nat_number_of, le_number_of_eq_not_less,  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
76  | 
mult_nat_number_of, nat_number_of_mult_left,  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
77  | 
less_nat_number_of,  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
78  | 
Let_number_of, nat_number_of] @  | 
| 
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
79  | 
bin_arith_simps @ bin_rel_simps;  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
80  | 
|
| 13462 | 81  | 
fun prep_simproc (name, pats, proc) =  | 
82  | 
Simplifier.simproc (Theory.sign_of (the_context ())) name pats proc;  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
85  | 
(*** CancelNumerals simprocs ***)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
86  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
87  | 
val one = mk_numeral 1;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
val mk_times = HOLogic.mk_binop "op *";  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
90  | 
fun mk_prod [] = one  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
| mk_prod [t] = t  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
92  | 
| mk_prod (t :: ts) = if t = one then mk_prod ts  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
else mk_times (t, mk_prod ts);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
val dest_times = HOLogic.dest_bin "op *" HOLogic.natT;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
fun dest_prod t =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
let val (t,u) = dest_times t  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
in dest_prod t @ dest_prod u end  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
handle TERM _ => [t];  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
(*DON'T do the obvious simplifications; that would create special cases*)  | 
| 
9544
 
f9202e219a29
added a dummy "thm list" argument to prove_conv for the new interface to
 
paulson 
parents: 
9436 
diff
changeset
 | 
103  | 
fun mk_coeff (k,t) = mk_times (mk_numeral k, t);  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
|
| 
9544
 
f9202e219a29
added a dummy "thm list" argument to prove_conv for the new interface to
 
paulson 
parents: 
9436 
diff
changeset
 | 
105  | 
(*Express t as a product of (possibly) a numeral with other factors, sorted*)  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
fun dest_coeff t =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
let val ts = sort Term.term_ord (dest_prod t)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
val (n, _, ts') = find_first_numeral [] ts  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
109  | 
handle TERM _ => (1, one, ts)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
in (n, mk_prod ts') end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
111  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
112  | 
(*Find first coefficient-term THAT MATCHES u*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
fun find_first_coeff past u [] = raise TERM("find_first_coeff", [])
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
| find_first_coeff past u (t::terms) =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
let val (n,u') = dest_coeff t  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
in if u aconv u' then (n, rev past @ terms)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
117  | 
else find_first_coeff (t::past) u terms  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
118  | 
end  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
handle TERM _ => find_first_coeff (t::past) u terms;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
120  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
121  | 
|
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
122  | 
(*Simplify 1*n and n*1 to n*)  | 
| 
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
123  | 
val add_0s = map rename_numerals [add_0, add_0_right];  | 
| 
14267
 
b963e9cee2a0
More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
 
paulson 
parents: 
13485 
diff
changeset
 | 
124  | 
val mult_1s = map rename_numerals [thm"nat_mult_1", thm"nat_mult_1_right"];  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
|
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11377 
diff
changeset
 | 
126  | 
(*Final simplification: cancel + and *; replace Numeral0 by 0 and Numeral1 by 1*)  | 
| 
12949
 
b94843ffc0d1
Fixed the natxx_cancel_numerals simprocs so that they pull out Sucs and
 
paulson 
parents: 
12338 
diff
changeset
 | 
127  | 
|
| 
 
b94843ffc0d1
Fixed the natxx_cancel_numerals simprocs so that they pull out Sucs and
 
paulson 
parents: 
12338 
diff
changeset
 | 
128  | 
(*And these help the simproc return False when appropriate, which helps  | 
| 
 
b94843ffc0d1
Fixed the natxx_cancel_numerals simprocs so that they pull out Sucs and
 
paulson 
parents: 
12338 
diff
changeset
 | 
129  | 
the arith prover.*)  | 
| 13462 | 130  | 
val contra_rules = [add_Suc, add_Suc_right, Zero_not_Suc, Suc_not_Zero,  | 
| 
12949
 
b94843ffc0d1
Fixed the natxx_cancel_numerals simprocs so that they pull out Sucs and
 
paulson 
parents: 
12338 
diff
changeset
 | 
131  | 
le_0_eq];  | 
| 
 
b94843ffc0d1
Fixed the natxx_cancel_numerals simprocs so that they pull out Sucs and
 
paulson 
parents: 
12338 
diff
changeset
 | 
132  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
val simplify_meta_eq =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
Int_Numeral_Simprocs.simplify_meta_eq  | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14369 
diff
changeset
 | 
135  | 
([nat_numeral_0_eq_0, numeral_1_eq_Suc_0, add_0, add_0_right,  | 
| 
12949
 
b94843ffc0d1
Fixed the natxx_cancel_numerals simprocs so that they pull out Sucs and
 
paulson 
parents: 
12338 
diff
changeset
 | 
136  | 
mult_0, mult_0_right, mult_1, mult_1_right] @ contra_rules);  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
|
| 10536 | 138  | 
|
| 
11377
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
139  | 
(** Restricted version of dest_Sucs_sum for nat_combine_numerals:  | 
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
140  | 
Simprocs never apply unless the original expression contains at least one  | 
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
141  | 
numeral in a coefficient position.  | 
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
142  | 
**)  | 
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
143  | 
|
| 12196 | 144  | 
fun ignore_Sucs (Const ("Suc", _) $ t) = ignore_Sucs t
 | 
145  | 
| ignore_Sucs t = t;  | 
|
146  | 
||
| 
11377
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
147  | 
fun is_numeral (Const("Numeral.number_of", _) $ w) = true
 | 
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
148  | 
| is_numeral _ = false;  | 
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
149  | 
|
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
150  | 
fun prod_has_numeral t = exists is_numeral (dest_prod t);  | 
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
151  | 
|
| 13462 | 152  | 
fun restricted_dest_Sucs_sum t =  | 
| 12196 | 153  | 
if exists prod_has_numeral (dest_sum (ignore_Sucs t))  | 
154  | 
then dest_Sucs_sum t  | 
|
155  | 
    else raise TERM("Nat_Numeral_Simprocs.restricted_dest_Sucs_sum", [t]);
 | 
|
| 
11377
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
156  | 
|
| 
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
157  | 
|
| 
14474
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
158  | 
(*Like HOL_ss but with an ordering that brings numerals to the front  | 
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
159  | 
under AC-rewriting.*)  | 
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
160  | 
val num_ss = Int_Numeral_Simprocs.num_ss;  | 
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
161  | 
|
| 10704 | 162  | 
(*** Applying CancelNumeralsFun ***)  | 
| 10536 | 163  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
164  | 
structure CancelNumeralsCommon =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
165  | 
struct  | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14369 
diff
changeset
 | 
166  | 
val mk_sum = (fn T:typ => mk_sum)  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
167  | 
val dest_sum = dest_Sucs_sum  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
val mk_coeff = mk_coeff  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
val dest_coeff = dest_coeff  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
170  | 
val find_first_coeff = find_first_coeff []  | 
| 16973 | 171  | 
val trans_tac = fn _ => trans_tac  | 
172  | 
fun norm_tac ss =  | 
|
| 
17877
 
67d5ab1cb0d8
Simplifier.inherit_context instead of Simplifier.inherit_bounds;
 
wenzelm 
parents: 
16973 
diff
changeset
 | 
173  | 
let val num_ss' = Simplifier.inherit_context ss num_ss in  | 
| 16973 | 174  | 
ALLGOALS (simp_tac (num_ss' addsimps numeral_syms @ add_0s @ mult_1s @  | 
| 
14474
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
175  | 
[Suc_eq_add_numeral_1_left] @ add_ac))  | 
| 16973 | 176  | 
THEN ALLGOALS (simp_tac (num_ss' addsimps bin_simps @ add_ac @ mult_ac))  | 
177  | 
end  | 
|
178  | 
fun numeral_simp_tac ss =  | 
|
| 
17877
 
67d5ab1cb0d8
Simplifier.inherit_context instead of Simplifier.inherit_bounds;
 
wenzelm 
parents: 
16973 
diff
changeset
 | 
179  | 
ALLGOALS (simp_tac (Simplifier.inherit_context ss HOL_ss addsimps add_0s @ bin_simps))  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
180  | 
val simplify_meta_eq = simplify_meta_eq  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
181  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
182  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
184  | 
structure EqCancelNumerals = CancelNumeralsFun  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
185  | 
(open CancelNumeralsCommon  | 
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
186  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
187  | 
val mk_bal = HOLogic.mk_eq  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
188  | 
val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
val bal_add1 = nat_eq_add_iff1 RS trans  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
190  | 
val bal_add2 = nat_eq_add_iff2 RS trans  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
192  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
193  | 
structure LessCancelNumerals = CancelNumeralsFun  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
194  | 
(open CancelNumeralsCommon  | 
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
195  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
196  | 
val mk_bal = HOLogic.mk_binrel "op <"  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
197  | 
val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
198  | 
val bal_add1 = nat_less_add_iff1 RS trans  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
199  | 
val bal_add2 = nat_less_add_iff2 RS trans  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
201  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
structure LeCancelNumerals = CancelNumeralsFun  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
203  | 
(open CancelNumeralsCommon  | 
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
204  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
205  | 
val mk_bal = HOLogic.mk_binrel "op <="  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
206  | 
val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
val bal_add1 = nat_le_add_iff1 RS trans  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
val bal_add2 = nat_le_add_iff2 RS trans  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
211  | 
structure DiffCancelNumerals = CancelNumeralsFun  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
212  | 
(open CancelNumeralsCommon  | 
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
213  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
214  | 
val mk_bal = HOLogic.mk_binop "op -"  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
215  | 
val dest_bal = HOLogic.dest_bin "op -" HOLogic.natT  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
216  | 
val bal_add1 = nat_diff_add_eq1 RS trans  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
217  | 
val bal_add2 = nat_diff_add_eq2 RS trans  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
218  | 
);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
221  | 
val cancel_numerals =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
222  | 
map prep_simproc  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
223  | 
   [("nateq_cancel_numerals",
 | 
| 13462 | 224  | 
["(l::nat) + m = n", "(l::nat) = m + n",  | 
225  | 
"(l::nat) * m = n", "(l::nat) = m * n",  | 
|
226  | 
"Suc m = n", "m = Suc n"],  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
227  | 
EqCancelNumerals.proc),  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
228  | 
    ("natless_cancel_numerals",
 | 
| 13462 | 229  | 
["(l::nat) + m < n", "(l::nat) < m + n",  | 
230  | 
"(l::nat) * m < n", "(l::nat) < m * n",  | 
|
231  | 
"Suc m < n", "m < Suc n"],  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
232  | 
LessCancelNumerals.proc),  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
233  | 
    ("natle_cancel_numerals",
 | 
| 13462 | 234  | 
["(l::nat) + m <= n", "(l::nat) <= m + n",  | 
235  | 
"(l::nat) * m <= n", "(l::nat) <= m * n",  | 
|
236  | 
"Suc m <= n", "m <= Suc n"],  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
237  | 
LeCancelNumerals.proc),  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
238  | 
    ("natdiff_cancel_numerals",
 | 
| 13462 | 239  | 
["((l::nat) + m) - n", "(l::nat) - (m + n)",  | 
240  | 
"(l::nat) * m - n", "(l::nat) - m * n",  | 
|
241  | 
"Suc m - n", "m - Suc n"],  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
242  | 
DiffCancelNumerals.proc)];  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
243  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
244  | 
|
| 10704 | 245  | 
(*** Applying CombineNumeralsFun ***)  | 
| 10536 | 246  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
247  | 
structure CombineNumeralsData =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
248  | 
struct  | 
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15921 
diff
changeset
 | 
249  | 
val add = IntInf.+  | 
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14369 
diff
changeset
 | 
250  | 
val mk_sum = (fn T:typ => long_mk_sum) (*to work for 2*x + 3*x *)  | 
| 
11377
 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 
paulson 
parents: 
11334 
diff
changeset
 | 
251  | 
val dest_sum = restricted_dest_Sucs_sum  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
252  | 
val mk_coeff = mk_coeff  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
253  | 
val dest_coeff = dest_coeff  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
254  | 
val left_distrib = left_add_mult_distrib RS trans  | 
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
255  | 
val prove_conv = Bin_Simprocs.prove_conv_nohyps  | 
| 16973 | 256  | 
val trans_tac = fn _ => trans_tac  | 
257  | 
fun norm_tac ss =  | 
|
| 
17877
 
67d5ab1cb0d8
Simplifier.inherit_context instead of Simplifier.inherit_bounds;
 
wenzelm 
parents: 
16973 
diff
changeset
 | 
258  | 
let val num_ss' = Simplifier.inherit_context ss num_ss in  | 
| 16973 | 259  | 
ALLGOALS (simp_tac (num_ss' addsimps numeral_syms @ add_0s @ mult_1s @  | 
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
260  | 
[Suc_eq_add_numeral_1] @ add_ac))  | 
| 16973 | 261  | 
THEN ALLGOALS (simp_tac (num_ss' addsimps bin_simps @ add_ac @ mult_ac))  | 
262  | 
end  | 
|
263  | 
fun numeral_simp_tac ss =  | 
|
| 
17877
 
67d5ab1cb0d8
Simplifier.inherit_context instead of Simplifier.inherit_bounds;
 
wenzelm 
parents: 
16973 
diff
changeset
 | 
264  | 
ALLGOALS (simp_tac (Simplifier.inherit_context ss HOL_ss addsimps add_0s @ bin_simps))  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
265  | 
val simplify_meta_eq = simplify_meta_eq  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
266  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
267  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
268  | 
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
269  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
270  | 
val combine_numerals =  | 
| 13462 | 271  | 
  prep_simproc ("nat_combine_numerals", ["(i::nat) + j", "Suc (i + j)"], CombineNumerals.proc);
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
272  | 
|
| 10536 | 273  | 
|
| 10704 | 274  | 
(*** Applying CancelNumeralFactorFun ***)  | 
| 10536 | 275  | 
|
276  | 
structure CancelNumeralFactorCommon =  | 
|
277  | 
struct  | 
|
| 13462 | 278  | 
val mk_coeff = mk_coeff  | 
279  | 
val dest_coeff = dest_coeff  | 
|
| 16973 | 280  | 
val trans_tac = fn _ => trans_tac  | 
281  | 
fun norm_tac ss =  | 
|
| 
17877
 
67d5ab1cb0d8
Simplifier.inherit_context instead of Simplifier.inherit_bounds;
 
wenzelm 
parents: 
16973 
diff
changeset
 | 
282  | 
let val num_ss' = Simplifier.inherit_context ss num_ss in  | 
| 16973 | 283  | 
ALLGOALS (simp_tac (num_ss' addsimps numeral_syms @ add_0s @ mult_1s @  | 
| 
14474
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
284  | 
[Suc_eq_add_numeral_1_left] @ add_ac))  | 
| 16973 | 285  | 
THEN ALLGOALS (simp_tac (num_ss' addsimps bin_simps @ add_ac @ mult_ac))  | 
286  | 
end  | 
|
287  | 
fun numeral_simp_tac ss =  | 
|
| 
17877
 
67d5ab1cb0d8
Simplifier.inherit_context instead of Simplifier.inherit_bounds;
 
wenzelm 
parents: 
16973 
diff
changeset
 | 
288  | 
ALLGOALS (simp_tac (Simplifier.inherit_context ss HOL_ss addsimps bin_simps))  | 
| 10536 | 289  | 
val simplify_meta_eq = simplify_meta_eq  | 
290  | 
end  | 
|
291  | 
||
292  | 
structure DivCancelNumeralFactor = CancelNumeralFactorFun  | 
|
293  | 
(open CancelNumeralFactorCommon  | 
|
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
294  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 10536 | 295  | 
val mk_bal = HOLogic.mk_binop "Divides.op div"  | 
296  | 
val dest_bal = HOLogic.dest_bin "Divides.op div" HOLogic.natT  | 
|
297  | 
val cancel = nat_mult_div_cancel1 RS trans  | 
|
298  | 
val neg_exchanges = false  | 
|
299  | 
)  | 
|
300  | 
||
301  | 
structure EqCancelNumeralFactor = CancelNumeralFactorFun  | 
|
302  | 
(open CancelNumeralFactorCommon  | 
|
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
303  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 10536 | 304  | 
val mk_bal = HOLogic.mk_eq  | 
305  | 
val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT  | 
|
306  | 
val cancel = nat_mult_eq_cancel1 RS trans  | 
|
307  | 
val neg_exchanges = false  | 
|
308  | 
)  | 
|
309  | 
||
310  | 
structure LessCancelNumeralFactor = CancelNumeralFactorFun  | 
|
311  | 
(open CancelNumeralFactorCommon  | 
|
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
312  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 10536 | 313  | 
val mk_bal = HOLogic.mk_binrel "op <"  | 
314  | 
val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT  | 
|
315  | 
val cancel = nat_mult_less_cancel1 RS trans  | 
|
316  | 
val neg_exchanges = true  | 
|
317  | 
)  | 
|
318  | 
||
319  | 
structure LeCancelNumeralFactor = CancelNumeralFactorFun  | 
|
320  | 
(open CancelNumeralFactorCommon  | 
|
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
321  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 10536 | 322  | 
val mk_bal = HOLogic.mk_binrel "op <="  | 
323  | 
val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT  | 
|
324  | 
val cancel = nat_mult_le_cancel1 RS trans  | 
|
325  | 
val neg_exchanges = true  | 
|
326  | 
)  | 
|
327  | 
||
| 13462 | 328  | 
val cancel_numeral_factors =  | 
| 10536 | 329  | 
map prep_simproc  | 
330  | 
   [("nateq_cancel_numeral_factors",
 | 
|
| 13462 | 331  | 
["(l::nat) * m = n", "(l::nat) = m * n"],  | 
| 10536 | 332  | 
EqCancelNumeralFactor.proc),  | 
| 13462 | 333  | 
    ("natless_cancel_numeral_factors",
 | 
334  | 
["(l::nat) * m < n", "(l::nat) < m * n"],  | 
|
| 10536 | 335  | 
LessCancelNumeralFactor.proc),  | 
| 13462 | 336  | 
    ("natle_cancel_numeral_factors",
 | 
337  | 
["(l::nat) * m <= n", "(l::nat) <= m * n"],  | 
|
| 10536 | 338  | 
LeCancelNumeralFactor.proc),  | 
| 13462 | 339  | 
    ("natdiv_cancel_numeral_factors",
 | 
340  | 
["((l::nat) * m) div n", "(l::nat) div (m * n)"],  | 
|
| 10536 | 341  | 
DivCancelNumeralFactor.proc)];  | 
342  | 
||
| 10704 | 343  | 
|
344  | 
||
345  | 
(*** Applying ExtractCommonTermFun ***)  | 
|
346  | 
||
347  | 
(*this version ALWAYS includes a trailing one*)  | 
|
348  | 
fun long_mk_prod [] = one  | 
|
349  | 
| long_mk_prod (t :: ts) = mk_times (t, mk_prod ts);  | 
|
350  | 
||
351  | 
(*Find first term that matches u*)  | 
|
| 13462 | 352  | 
fun find_first past u []         = raise TERM("find_first", [])
 | 
| 10704 | 353  | 
| find_first past u (t::terms) =  | 
| 13462 | 354  | 
if u aconv t then (rev past @ terms)  | 
| 10704 | 355  | 
else find_first (t::past) u terms  | 
| 13462 | 356  | 
handle TERM _ => find_first (t::past) u terms;  | 
| 10704 | 357  | 
|
| 
15271
 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 
paulson 
parents: 
14474 
diff
changeset
 | 
358  | 
(** Final simplification for the CancelFactor simprocs **)  | 
| 
 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 
paulson 
parents: 
14474 
diff
changeset
 | 
359  | 
val simplify_one =  | 
| 
 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 
paulson 
parents: 
14474 
diff
changeset
 | 
360  | 
Int_Numeral_Simprocs.simplify_meta_eq  | 
| 
 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 
paulson 
parents: 
14474 
diff
changeset
 | 
361  | 
[mult_1_left, mult_1_right, div_1, numeral_1_eq_Suc_0];  | 
| 
 
3c32a26510c4
fixed some awkward problems with nat/int simprocs
 
paulson 
parents: 
14474 
diff
changeset
 | 
362  | 
|
| 16973 | 363  | 
fun cancel_simplify_meta_eq cancel_th ss th =  | 
364  | 
simplify_one ss (([th, cancel_th]) MRS trans);  | 
|
| 10704 | 365  | 
|
366  | 
structure CancelFactorCommon =  | 
|
367  | 
struct  | 
|
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14369 
diff
changeset
 | 
368  | 
val mk_sum = (fn T:typ => long_mk_prod)  | 
| 13462 | 369  | 
val dest_sum = dest_prod  | 
370  | 
val mk_coeff = mk_coeff  | 
|
371  | 
val dest_coeff = dest_coeff  | 
|
372  | 
val find_first = find_first []  | 
|
| 16973 | 373  | 
val trans_tac = fn _ => trans_tac  | 
374  | 
fun norm_tac ss =  | 
|
| 
17877
 
67d5ab1cb0d8
Simplifier.inherit_context instead of Simplifier.inherit_bounds;
 
wenzelm 
parents: 
16973 
diff
changeset
 | 
375  | 
ALLGOALS (simp_tac (Simplifier.inherit_context ss HOL_ss addsimps mult_1s @ mult_ac))  | 
| 10704 | 376  | 
end;  | 
377  | 
||
378  | 
structure EqCancelFactor = ExtractCommonTermFun  | 
|
379  | 
(open CancelFactorCommon  | 
|
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
380  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 10704 | 381  | 
val mk_bal = HOLogic.mk_eq  | 
382  | 
val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT  | 
|
383  | 
val simplify_meta_eq = cancel_simplify_meta_eq nat_mult_eq_cancel_disj  | 
|
384  | 
);  | 
|
385  | 
||
386  | 
structure LessCancelFactor = ExtractCommonTermFun  | 
|
387  | 
(open CancelFactorCommon  | 
|
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
388  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 10704 | 389  | 
val mk_bal = HOLogic.mk_binrel "op <"  | 
390  | 
val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT  | 
|
391  | 
val simplify_meta_eq = cancel_simplify_meta_eq nat_mult_less_cancel_disj  | 
|
392  | 
);  | 
|
393  | 
||
394  | 
structure LeCancelFactor = ExtractCommonTermFun  | 
|
395  | 
(open CancelFactorCommon  | 
|
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
396  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 10704 | 397  | 
val mk_bal = HOLogic.mk_binrel "op <="  | 
398  | 
val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT  | 
|
399  | 
val simplify_meta_eq = cancel_simplify_meta_eq nat_mult_le_cancel_disj  | 
|
400  | 
);  | 
|
401  | 
||
402  | 
structure DivideCancelFactor = ExtractCommonTermFun  | 
|
403  | 
(open CancelFactorCommon  | 
|
| 
13485
 
acf39e924091
tuned prove_conv (error reporting done within meta_simplifier.ML);
 
wenzelm 
parents: 
13462 
diff
changeset
 | 
404  | 
val prove_conv = Bin_Simprocs.prove_conv  | 
| 10704 | 405  | 
val mk_bal = HOLogic.mk_binop "Divides.op div"  | 
406  | 
val dest_bal = HOLogic.dest_bin "Divides.op div" HOLogic.natT  | 
|
407  | 
val simplify_meta_eq = cancel_simplify_meta_eq nat_mult_div_cancel_disj  | 
|
408  | 
);  | 
|
409  | 
||
| 13462 | 410  | 
val cancel_factor =  | 
| 10704 | 411  | 
map prep_simproc  | 
412  | 
   [("nat_eq_cancel_factor",
 | 
|
| 13462 | 413  | 
["(l::nat) * m = n", "(l::nat) = m * n"],  | 
| 10704 | 414  | 
EqCancelFactor.proc),  | 
415  | 
    ("nat_less_cancel_factor",
 | 
|
| 13462 | 416  | 
["(l::nat) * m < n", "(l::nat) < m * n"],  | 
| 10704 | 417  | 
LessCancelFactor.proc),  | 
418  | 
    ("nat_le_cancel_factor",
 | 
|
| 13462 | 419  | 
["(l::nat) * m <= n", "(l::nat) <= m * n"],  | 
| 10704 | 420  | 
LeCancelFactor.proc),  | 
| 13462 | 421  | 
    ("nat_divide_cancel_factor",
 | 
422  | 
["((l::nat) * m) div n", "(l::nat) div (m * n)"],  | 
|
| 10704 | 423  | 
DivideCancelFactor.proc)];  | 
424  | 
||
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
425  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
426  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
427  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
428  | 
Addsimprocs Nat_Numeral_Simprocs.cancel_numerals;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
429  | 
Addsimprocs [Nat_Numeral_Simprocs.combine_numerals];  | 
| 10536 | 430  | 
Addsimprocs Nat_Numeral_Simprocs.cancel_numeral_factors;  | 
| 10704 | 431  | 
Addsimprocs Nat_Numeral_Simprocs.cancel_factor;  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
432  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
433  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
434  | 
(*examples:  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
435  | 
print_depth 22;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
436  | 
set timing;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
437  | 
set trace_simp;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
438  | 
fun test s = (Goal s; by (Simp_tac 1));  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
439  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
440  | 
(*cancel_numerals*)  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
441  | 
test "l +( 2) + (2) + 2 + (l + 2) + (oo + 2) = (uu::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
442  | 
test "(2*length xs < 2*length xs + j)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
443  | 
test "(2*length xs < length xs * 2 + j)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
444  | 
test "2*u = (u::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
445  | 
test "2*u = Suc (u)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
446  | 
test "(i + j + 12 + (k::nat)) - 15 = y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
447  | 
test "(i + j + 12 + (k::nat)) - 5 = y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
448  | 
test "Suc u - 2 = y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
449  | 
test "Suc (Suc (Suc u)) - 2 = y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
450  | 
test "(i + j + 2 + (k::nat)) - 1 = y";  | 
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
451  | 
test "(i + j + 1 + (k::nat)) - 2 = y";  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
452  | 
|
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
453  | 
test "(2*x + (u*v) + y) - v*3*u = (w::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
454  | 
test "(2*x*u*v + 5 + (u*v)*4 + y) - v*u*4 = (w::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
455  | 
test "(2*x*u*v + (u*v)*4 + y) - v*u = (w::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
456  | 
test "Suc (Suc (2*x*u*v + u*4 + y)) - u = w";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
457  | 
test "Suc ((u*v)*4) - v*3*u = w";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
458  | 
test "Suc (Suc ((u*v)*3)) - v*3*u = w";  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
459  | 
|
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
460  | 
test "(i + j + 12 + (k::nat)) = u + 15 + y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
461  | 
test "(i + j + 32 + (k::nat)) - (u + 15 + y) = zz";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
462  | 
test "(i + j + 12 + (k::nat)) = u + 5 + y";  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
463  | 
(*Suc*)  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
464  | 
test "(i + j + 12 + k) = Suc (u + y)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
465  | 
test "Suc (Suc (Suc (Suc (Suc (u + y))))) <= ((i + j) + 41 + k)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
466  | 
test "(i + j + 5 + k) < Suc (Suc (Suc (Suc (Suc (u + y)))))";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
467  | 
test "Suc (Suc (Suc (Suc (Suc (u + y))))) - 5 = v";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
468  | 
test "(i + j + 5 + k) = Suc (Suc (Suc (Suc (Suc (Suc (Suc (u + y)))))))";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
469  | 
test "2*y + 3*z + 2*u = Suc (u)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
470  | 
test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = Suc (u)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
471  | 
test "2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
472  | 
test "6 + 2*y + 3*z + 4*u = Suc (vv + 2*u + z)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
473  | 
test "(2*n*m) < (3*(m*n)) + (u::nat)";  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
474  | 
|
| 
14474
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
475  | 
test "(Suc (Suc (Suc (Suc (Suc (Suc (case length (f c) of 0 => 0 | Suc k => k)))))) <= Suc 0)";  | 
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
476  | 
|
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
477  | 
test "Suc (Suc (Suc (Suc (Suc (Suc (length l1 + length l2)))))) <= length l1";  | 
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
478  | 
|
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
479  | 
test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length l3)))))) <= length (compT P E A ST mxr e))";  | 
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
480  | 
|
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
481  | 
test "( (Suc (Suc (Suc (Suc (Suc (length (compT P E A ST mxr e) + length (compT P E (A Un \<A> e) ST mxr c))))))) <= length (compT P E A ST mxr e))";  | 
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
482  | 
|
| 
 
00292f6f8d13
New simplification ordering to move numerals together. Fixes a bug in the
 
paulson 
parents: 
14430 
diff
changeset
 | 
483  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
484  | 
(*negative numerals: FAIL*)  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
485  | 
test "(i + j + -23 + (k::nat)) < u + 15 + y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
486  | 
test "(i + j + 3 + (k::nat)) < u + -15 + y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
487  | 
test "(i + j + -12 + (k::nat)) - 15 = y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
488  | 
test "(i + j + 12 + (k::nat)) - -15 = y";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
489  | 
test "(i + j + -12 + (k::nat)) - -15 = y";  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
490  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
491  | 
(*combine_numerals*)  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
492  | 
test "k + 3*k = (u::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
493  | 
test "Suc (i + 3) = u";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
494  | 
test "Suc (i + j + 3 + k) = u";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
495  | 
test "k + j + 3*k + j = (u::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
496  | 
test "Suc (j*i + i + k + 5 + 3*k + i*j*4) = (u::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
497  | 
test "(2*n*m) + (3*(m*n)) = (u::nat)";  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
498  | 
(*negative numerals: FAIL*)  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
499  | 
test "Suc (i + j + -3 + k) = u";  | 
| 10536 | 500  | 
|
| 10704 | 501  | 
(*cancel_numeral_factors*)  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
502  | 
test "9*x = 12 * (y::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
503  | 
test "(9*x) div (12 * (y::nat)) = z";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
504  | 
test "9*x < 12 * (y::nat)";  | 
| 
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
505  | 
test "9*x <= 12 * (y::nat)";  | 
| 10704 | 506  | 
|
507  | 
(*cancel_factor*)  | 
|
508  | 
test "x*k = k*(y::nat)";  | 
|
| 13462 | 509  | 
test "k = k*(y::nat)";  | 
| 10704 | 510  | 
test "a*(b*c) = (b::nat)";  | 
511  | 
test "a*(b*c) = d*(b::nat)*(x*a)";  | 
|
512  | 
||
513  | 
test "x*k < k*(y::nat)";  | 
|
| 13462 | 514  | 
test "k < k*(y::nat)";  | 
| 10704 | 515  | 
test "a*(b*c) < (b::nat)";  | 
516  | 
test "a*(b*c) < d*(b::nat)*(x*a)";  | 
|
517  | 
||
518  | 
test "x*k <= k*(y::nat)";  | 
|
| 13462 | 519  | 
test "k <= k*(y::nat)";  | 
| 10704 | 520  | 
test "a*(b*c) <= (b::nat)";  | 
521  | 
test "a*(b*c) <= d*(b::nat)*(x*a)";  | 
|
522  | 
||
523  | 
test "(x*k) div (k*(y::nat)) = (uu::nat)";  | 
|
| 13462 | 524  | 
test "(k) div (k*(y::nat)) = (uu::nat)";  | 
| 10704 | 525  | 
test "(a*(b*c)) div ((b::nat)) = (uu::nat)";  | 
526  | 
test "(a*(b*c)) div (d*(b::nat)*(x*a)) = (uu::nat)";  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
527  | 
*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
528  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
529  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
530  | 
(*** Prepare linear arithmetic for nat numerals ***)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
531  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
532  | 
local  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
533  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
534  | 
(* reduce contradictory <= to False *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
535  | 
val add_rules =  | 
| 
14430
 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 
paulson 
parents: 
14390 
diff
changeset
 | 
536  | 
[thm "Let_number_of", Let_0, Let_1, nat_0, nat_1,  | 
| 
11868
 
56db9f3a6b3e
Numerals now work for the integers: the binary numerals for 0 and 1 rewrite
 
paulson 
parents: 
11704 
diff
changeset
 | 
537  | 
add_nat_number_of, diff_nat_number_of, mult_nat_number_of,  | 
| 
14365
 
3d4df8c166ae
replacing HOL/Real/PRat, PNat by the rational number development
 
paulson 
parents: 
14273 
diff
changeset
 | 
538  | 
eq_nat_number_of, less_nat_number_of, le_number_of_eq_not_less,  | 
| 11296 | 539  | 
le_Suc_number_of,le_number_of_Suc,  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
540  | 
less_Suc_number_of,less_number_of_Suc,  | 
| 11296 | 541  | 
Suc_eq_number_of,eq_number_of_Suc,  | 
| 14369 | 542  | 
mult_Suc, mult_Suc_right,  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
543  | 
eq_number_of_0, eq_0_number_of, less_0_number_of,  | 
| 14390 | 544  | 
of_int_number_of_eq, of_nat_number_of_eq, nat_number_of, if_True, if_False];  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
545  | 
|
| 
14387
 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 
paulson 
parents: 
14369 
diff
changeset
 | 
546  | 
val simprocs = [Nat_Numeral_Simprocs.combine_numerals]@  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
547  | 
Nat_Numeral_Simprocs.cancel_numerals;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
548  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
549  | 
in  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
550  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
551  | 
val nat_simprocs_setup =  | 
| 15921 | 552  | 
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
 | 
| 10693 | 553  | 
   {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms,
 | 
| 15921 | 554  | 
inj_thms = inj_thms, lessD = lessD, neqE = neqE,  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
555  | 
simpset = simpset addsimps add_rules  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
556  | 
addsimprocs simprocs})];  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
557  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
558  | 
end;  |