| author | paulson | 
| Wed, 08 May 2002 09:14:56 +0200 | |
| changeset 13117 | 0b233f430076 | 
| parent 12949 | b94843ffc0d1 | 
| child 13462 | 56610e2ba220 | 
| permissions | -rw-r--r-- | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/arith_data.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: Markus Wenzel, Stefan Berghofer and Tobias Nipkow  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
Various arithmetic proof procedures.  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
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  | 
(*---------------------------------------------------------------------------*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
(* 1. Cancellation of common terms *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
(*---------------------------------------------------------------------------*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
signature ARITH_DATA =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
sig  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
val nat_cancel_sums_add: simproc list  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
val nat_cancel_sums: simproc list  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
structure ArithData: ARITH_DATA =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
19  | 
struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
22  | 
(** abstract syntax of structure nat: 0, Suc, + **)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
(* mk_sum, mk_norm_sum *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
val one = HOLogic.mk_nat 1;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
val mk_plus = HOLogic.mk_binop "op +";  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
fun mk_sum [] = HOLogic.zero  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
| mk_sum [t] = t  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
| mk_sum (t :: ts) = mk_plus (t, mk_sum ts);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
(*normal form of sums: Suc (... (Suc (a + (b + ...))))*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
fun mk_norm_sum ts =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
35  | 
let val (ones, sums) = partition (equal one) ts in  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
funpow (length ones) HOLogic.mk_Suc (mk_sum sums)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
(* dest_sum *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
val dest_plus = HOLogic.dest_bin "op +" HOLogic.natT;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
fun dest_sum tm =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
if HOLogic.is_zero tm then []  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
else  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
(case try HOLogic.dest_Suc tm of  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
Some t => one :: dest_sum t  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
| None =>  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
(case try dest_plus tm of  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
Some (t, u) => dest_sum t @ dest_sum u  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
| None => [tm]));  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
|
| 
 
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  | 
(** generic proof tools **)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
(* prove conversions *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
val mk_eqv = HOLogic.mk_Trueprop o HOLogic.mk_eq;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
fun prove_conv expand_tac norm_tac sg (t, u) =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
mk_meta_eq (prove_goalw_cterm_nocheck [] (cterm_of sg (mk_eqv (t, u)))  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
(K [expand_tac, norm_tac]))  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
  handle ERROR => error ("The error(s) above occurred while trying to prove " ^
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
(string_of_cterm (cterm_of sg (mk_eqv (t, u)))));  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
val subst_equals = prove_goal HOL.thy "[| t = s; u = t |] ==> u = s"  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
68  | 
(fn prems => [cut_facts_tac prems 1, SIMPSET' asm_simp_tac 1]);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
(* rewriting *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
73  | 
fun simp_all rules = ALLGOALS (simp_tac (HOL_ss addsimps rules));  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
val add_rules = [add_Suc, add_Suc_right, add_0, add_0_right];  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
76  | 
val mult_rules = [mult_Suc, mult_Suc_right, mult_0, mult_0_right];  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
78  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
79  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
80  | 
(** cancel common summands **)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
81  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
structure Sum =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
83  | 
struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
84  | 
val mk_sum = mk_norm_sum;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
85  | 
val dest_sum = dest_sum;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
86  | 
val prove_conv = prove_conv;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
87  | 
val norm_tac = simp_all add_rules THEN simp_all add_ac;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
end;  | 
| 
 
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 gen_uncancel_tac rule ct =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
rtac (instantiate' [] [None, Some ct] (rule RS subst_equals)) 1;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
92  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
(* nat eq *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
structure EqCancelSums = CancelSumsFun  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
(struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
open Sum;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
val mk_bal = HOLogic.mk_eq;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
val dest_bal = HOLogic.dest_bin "op =" HOLogic.natT;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
val uncancel_tac = gen_uncancel_tac add_left_cancel;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
end);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
(* nat less *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
structure LessCancelSums = CancelSumsFun  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
(struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
109  | 
open Sum;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
val mk_bal = HOLogic.mk_binrel "op <";  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
111  | 
val dest_bal = HOLogic.dest_bin "op <" HOLogic.natT;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
112  | 
val uncancel_tac = gen_uncancel_tac add_left_cancel_less;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
end);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
(* nat le *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
117  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
118  | 
structure LeCancelSums = CancelSumsFun  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
(struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
120  | 
open Sum;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
121  | 
val mk_bal = HOLogic.mk_binrel "op <=";  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
122  | 
val dest_bal = HOLogic.dest_bin "op <=" HOLogic.natT;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
123  | 
val uncancel_tac = gen_uncancel_tac add_left_cancel_le;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
124  | 
end);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
127  | 
(* nat diff *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
structure DiffCancelSums = CancelSumsFun  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
(struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
131  | 
open Sum;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
132  | 
val mk_bal = HOLogic.mk_binop "op -";  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
val dest_bal = HOLogic.dest_bin "op -" HOLogic.natT;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
val uncancel_tac = gen_uncancel_tac diff_cancel;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
end);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
(** prepare nat_cancel simprocs **)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
|
| 
12338
 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 
wenzelm 
parents: 
12311 
diff
changeset
 | 
141  | 
fun prep_pat s = HOLogic.read_cterm (Theory.sign_of (the_context ())) s;  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
val prep_pats = map prep_pat;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
fun prep_simproc (name, pats, proc) = Simplifier.mk_simproc name pats proc;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
145  | 
|
| 10766 | 146  | 
val eq_pats = prep_pats ["(l::nat) + m = n", "(l::nat) = m + n", "Suc m = n",  | 
147  | 
"m = Suc n"];  | 
|
148  | 
val less_pats = prep_pats ["(l::nat) + m < n", "(l::nat) < m + n", "Suc m < n",  | 
|
149  | 
"m < Suc n"];  | 
|
150  | 
val le_pats = prep_pats ["(l::nat) + m <= n", "(l::nat) <= m + n",  | 
|
151  | 
"Suc m <= n", "m <= Suc n"];  | 
|
152  | 
val diff_pats = prep_pats ["((l::nat) + m) - n", "(l::nat) - (m + n)",  | 
|
153  | 
"Suc m - n", "m - Suc n"];  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
154  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
155  | 
val nat_cancel_sums_add = map prep_simproc  | 
| 10766 | 156  | 
  [("nateq_cancel_sums",   eq_pats,   EqCancelSums.proc),
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
157  | 
   ("natless_cancel_sums", less_pats, LessCancelSums.proc),
 | 
| 10766 | 158  | 
   ("natle_cancel_sums",   le_pats,   LeCancelSums.proc)];
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
159  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
160  | 
val nat_cancel_sums = nat_cancel_sums_add @  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
  [prep_simproc("natdiff_cancel_sums", diff_pats, DiffCancelSums.proc)];
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
162  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
163  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
164  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
165  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
166  | 
open ArithData;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
167  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
(*---------------------------------------------------------------------------*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
170  | 
(* 2. Linear arithmetic *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
171  | 
(*---------------------------------------------------------------------------*)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
172  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
173  | 
(* Parameters data for general linear arithmetic functor *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
174  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
175  | 
structure LA_Logic: LIN_ARITH_LOGIC =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
176  | 
struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
177  | 
val ccontr = ccontr;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
178  | 
val conjI = conjI;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
179  | 
val neqE = linorder_neqE;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
180  | 
val notI = notI;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
181  | 
val sym = sym;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
182  | 
val not_lessD = linorder_not_less RS iffD1;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
183  | 
val not_leD = linorder_not_le RS iffD1;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
184  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
185  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
186  | 
fun mk_Eq thm = (thm RS Eq_FalseI) handle THM _ => (thm RS Eq_TrueI);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
187  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
188  | 
val mk_Trueprop = HOLogic.mk_Trueprop;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
190  | 
fun neg_prop(TP$(Const("Not",_)$t)) = TP$t
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
  | neg_prop(TP$t) = TP $ (Const("Not",HOLogic.boolT-->HOLogic.boolT)$t);
 | 
| 
 
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  | 
fun is_False thm =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
194  | 
let val _ $ t = #prop(rep_thm thm)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
195  | 
  in t = Const("False",HOLogic.boolT) end;
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
196  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
197  | 
fun is_nat(t) = fastype_of1 t = HOLogic.natT;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
198  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
199  | 
fun mk_nat_thm sg t =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
  let val ct = cterm_of sg t  and cn = cterm_of sg (Var(("n",0),HOLogic.natT))
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
201  | 
in instantiate ([],[(cn,ct)]) le0 end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
203  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
204  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
205  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
206  | 
(* arith theory data *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
|
| 9593 | 208  | 
structure ArithTheoryDataArgs =  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
val name = "HOL/arith";  | 
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
211  | 
  type T = {splits: thm list, inj_consts: (string * typ)list, discrete: (string * bool) list};
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
212  | 
|
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
213  | 
  val empty = {splits = [], inj_consts = [], discrete = []};
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
214  | 
val copy = I;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
215  | 
val prep_ext = I;  | 
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
216  | 
  fun merge ({splits= splits1, inj_consts= inj_consts1, discrete= discrete1},
 | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
217  | 
             {splits= splits2, inj_consts= inj_consts2, discrete= discrete2}) =
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
218  | 
   {splits = Drule.merge_rules (splits1, splits2),
 | 
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
219  | 
inj_consts = merge_lists inj_consts1 inj_consts2,  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
discrete = merge_alists discrete1 discrete2};  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
221  | 
fun print _ _ = ();  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
222  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
223  | 
|
| 9593 | 224  | 
structure ArithTheoryData = TheoryDataFun(ArithTheoryDataArgs);  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
|
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
226  | 
fun arith_split_add (thy, thm) = (ArithTheoryData.map (fn {splits,inj_consts,discrete} =>
 | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
227  | 
  {splits= thm::splits, inj_consts= inj_consts, discrete= discrete}) thy, thm);
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
228  | 
|
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
229  | 
fun arith_discrete d = ArithTheoryData.map (fn {splits,inj_consts,discrete} =>
 | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
230  | 
  {splits = splits, inj_consts = inj_consts, discrete = d :: discrete});
 | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
231  | 
|
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
232  | 
fun arith_inj_const c = ArithTheoryData.map (fn {splits,inj_consts,discrete} =>
 | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
233  | 
  {splits = splits, inj_consts = c :: inj_consts, discrete = discrete});
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
234  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
235  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
236  | 
structure LA_Data_Ref: LIN_ARITH_DATA =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
237  | 
struct  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
238  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
239  | 
(* Decomposition of terms *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
240  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
241  | 
fun nT (Type("fun",[N,_])) = N = HOLogic.natT
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
242  | 
| nT _ = false;  | 
| 
 
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  | 
fun add_atom(t,m,(p,i)) = (case assoc(p,t) of None => ((t,m)::p,i)  | 
| 10693 | 245  | 
| Some n => (overwrite(p,(t,ratadd(n,m))), i));  | 
246  | 
||
247  | 
exception Zero;  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
248  | 
|
| 10693 | 249  | 
fun rat_of_term(numt,dent) =  | 
250  | 
let val num = HOLogic.dest_binum numt and den = HOLogic.dest_binum dent  | 
|
251  | 
in if den = 0 then raise Zero else int_ratdiv(num,den) end;  | 
|
| 10718 | 252  | 
|
253  | 
(* Warning: in rare cases number_of encloses a non-numeral,  | 
|
254  | 
in which case dest_binum raises TERM; hence all the handles below.  | 
|
| 
11334
 
a16eaf2a1edd
Allow Suc-numerals as coefficients in lin-arith formulae
 
nipkow 
parents: 
10906 
diff
changeset
 | 
255  | 
Same for Suc-terms that turn out not to be numerals -  | 
| 
 
a16eaf2a1edd
Allow Suc-numerals as coefficients in lin-arith formulae
 
nipkow 
parents: 
10906 
diff
changeset
 | 
256  | 
although the simplifier should eliminate those anyway...  | 
| 10718 | 257  | 
*)  | 
258  | 
||
| 
11334
 
a16eaf2a1edd
Allow Suc-numerals as coefficients in lin-arith formulae
 
nipkow 
parents: 
10906 
diff
changeset
 | 
259  | 
fun number_of_Sucs (Const("Suc",_) $ n) = number_of_Sucs n + 1
 | 
| 
 
a16eaf2a1edd
Allow Suc-numerals as coefficients in lin-arith formulae
 
nipkow 
parents: 
10906 
diff
changeset
 | 
260  | 
| number_of_Sucs t = if HOLogic.is_zero t then 0  | 
| 
 
a16eaf2a1edd
Allow Suc-numerals as coefficients in lin-arith formulae
 
nipkow 
parents: 
10906 
diff
changeset
 | 
261  | 
                       else raise TERM("number_of_Sucs",[])
 | 
| 
 
a16eaf2a1edd
Allow Suc-numerals as coefficients in lin-arith formulae
 
nipkow 
parents: 
10906 
diff
changeset
 | 
262  | 
|
| 10718 | 263  | 
(* decompose nested multiplications, bracketing them to the right and combining all  | 
264  | 
their coefficients  | 
|
265  | 
*)  | 
|
266  | 
||
267  | 
fun demult((mC as Const("op *",_)) $ s $ t,m) = ((case s of
 | 
|
268  | 
        Const("Numeral.number_of",_)$n
 | 
|
269  | 
=> demult(t,ratmul(m,rat_of_int(HOLogic.dest_binum n)))  | 
|
| 
12480
 
32e67277a4b9
tuned conversion from terms to "polynomials" for arith_tac: takes care
 
nipkow 
parents: 
12338 
diff
changeset
 | 
270  | 
      | Const("uminus",_)$(Const("Numeral.number_of",_)$n)
 | 
| 
 
32e67277a4b9
tuned conversion from terms to "polynomials" for arith_tac: takes care
 
nipkow 
parents: 
12338 
diff
changeset
 | 
271  | 
=> demult(t,ratmul(m,rat_of_int(~(HOLogic.dest_binum n))))  | 
| 
11334
 
a16eaf2a1edd
Allow Suc-numerals as coefficients in lin-arith formulae
 
nipkow 
parents: 
10906 
diff
changeset
 | 
272  | 
      | Const("Suc",_) $ _
 | 
| 
 
a16eaf2a1edd
Allow Suc-numerals as coefficients in lin-arith formulae
 
nipkow 
parents: 
10906 
diff
changeset
 | 
273  | 
=> demult(t,ratmul(m,rat_of_int(number_of_Sucs s)))  | 
| 10718 | 274  | 
      | Const("op *",_) $ s1 $ s2 => demult(mC $ s1 $ (mC $ s2 $ t),m)
 | 
275  | 
      | Const("HOL.divide",_) $ numt $ (Const("Numeral.number_of",_)$dent) =>
 | 
|
276  | 
let val den = HOLogic.dest_binum dent  | 
|
277  | 
in if den = 0 then raise Zero  | 
|
278  | 
else demult(mC $ numt $ t,ratmul(m, ratinv(rat_of_int den)))  | 
|
279  | 
end  | 
|
280  | 
| _ => atomult(mC,s,t,m)  | 
|
281  | 
) handle TERM _ => atomult(mC,s,t,m))  | 
|
282  | 
  | demult(atom as Const("HOL.divide",_) $ t $ (Const("Numeral.number_of",_)$dent), m) =
 | 
|
283  | 
(let val den = HOLogic.dest_binum dent  | 
|
284  | 
in if den = 0 then raise Zero else demult(t,ratmul(m, ratinv(rat_of_int den))) end  | 
|
285  | 
handle TERM _ => (Some atom,m))  | 
|
286  | 
  | demult(t as Const("Numeral.number_of",_)$n,m) =
 | 
|
287  | 
((None,ratmul(m,rat_of_int(HOLogic.dest_binum n)))  | 
|
288  | 
handle TERM _ => (Some t,m))  | 
|
| 
12480
 
32e67277a4b9
tuned conversion from terms to "polynomials" for arith_tac: takes care
 
nipkow 
parents: 
12338 
diff
changeset
 | 
289  | 
  | demult(Const("uminus",_)$t, m) = demult(t,ratmul(m,rat_of_int(~1)))
 | 
| 10718 | 290  | 
| demult(atom,m) = (Some atom,m)  | 
291  | 
||
292  | 
and atomult(mC,atom,t,m) = (case demult(t,m) of (None,m') => (Some atom,m')  | 
|
293  | 
| (Some t',m') => (Some(mC $ atom $ t'),m'))  | 
|
294  | 
||
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
295  | 
fun decomp2 inj_consts (rel,lhs,rhs) =  | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
296  | 
let  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
297  | 
(* Turn term into list of summand * multiplicity plus a constant *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
298  | 
fun poly(Const("op +",_) $ s $ t, m, pi) = poly(s,m,poly(t,m,pi))
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
299  | 
  | poly(all as Const("op -",T) $ s $ t, m, pi) =
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
300  | 
if nT T then add_atom(all,m,pi)  | 
| 10693 | 301  | 
else poly(s,m,poly(t,ratneg m,pi))  | 
302  | 
  | poly(Const("uminus",_) $ t, m, pi) = poly(t,ratneg m,pi)
 | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
303  | 
  | poly(Const("0",_), _, pi) = pi
 | 
| 11464 | 304  | 
  | poly(Const("1",_), m, (p,i)) = (p,ratadd(i,m))
 | 
| 10693 | 305  | 
  | poly(Const("Suc",_)$t, m, (p,i)) = poly(t, m, (p,ratadd(i,m)))
 | 
| 10718 | 306  | 
  | poly(t as Const("op *",_) $ _ $ _, m, pi as (p,i)) =
 | 
307  | 
(case demult(t,m) of  | 
|
308  | 
(None,m') => (p,ratadd(i,m))  | 
|
309  | 
| (Some u,m') => add_atom(u,m',pi))  | 
|
310  | 
  | poly(t as Const("HOL.divide",_) $ _ $ _, m, pi as (p,i)) =
 | 
|
311  | 
(case demult(t,m) of  | 
|
312  | 
(None,m') => (p,ratadd(i,m))  | 
|
313  | 
| (Some u,m') => add_atom(u,m',pi))  | 
|
314  | 
  | poly(all as (Const("Numeral.number_of",_)$t,m,(p,i))) =
 | 
|
315  | 
((p,ratadd(i,ratmul(m,rat_of_int(HOLogic.dest_binum t))))  | 
|
316  | 
handle TERM _ => add_atom all)  | 
|
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
317  | 
| poly(all as Const f $ x, m, pi) =  | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
318  | 
if f mem inj_consts then poly(x,m,pi) else add_atom(all,m,pi)  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
319  | 
| poly x = add_atom x;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
320  | 
|
| 10718 | 321  | 
val (p,i) = poly(lhs,rat_of_int 1,([],rat_of_int 0))  | 
322  | 
and (q,j) = poly(rhs,rat_of_int 1,([],rat_of_int 0))  | 
|
| 10693 | 323  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
324  | 
in case rel of  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
325  | 
"op <" => Some(p,i,"<",q,j)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
326  | 
| "op <=" => Some(p,i,"<=",q,j)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
327  | 
| "op =" => Some(p,i,"=",q,j)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
328  | 
| _ => None  | 
| 10693 | 329  | 
end handle Zero => None;  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
330  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
331  | 
fun negate(Some(x,i,rel,y,j,d)) = Some(x,i,"~"^rel,y,j,d)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
332  | 
| negate None = None;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
333  | 
|
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
334  | 
fun decomp1 (discrete,inj_consts) (T,xxx) =  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
335  | 
(case T of  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
336  | 
     Type("fun",[Type(D,[]),_]) =>
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
337  | 
(case assoc(discrete,D) of  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
338  | 
None => None  | 
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
339  | 
| Some d => (case decomp2 inj_consts xxx of  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
340  | 
None => None  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
341  | 
| Some(p,i,rel,q,j) => Some(p,i,rel,q,j,d)))  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
342  | 
| _ => None);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
343  | 
|
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
344  | 
fun decomp2 data (_$(Const(rel,T)$lhs$rhs)) = decomp1 data (T,(rel,lhs,rhs))  | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
345  | 
  | decomp2 data (_$(Const("Not",_)$(Const(rel,T)$lhs$rhs))) =
 | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
346  | 
negate(decomp1 data (T,(rel,lhs,rhs)))  | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
347  | 
| decomp2 data _ = None  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
348  | 
|
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
349  | 
fun decomp sg =  | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
350  | 
  let val {discrete, inj_consts, ...} = ArithTheoryData.get_sg sg
 | 
| 
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
351  | 
in decomp2 (discrete,inj_consts) end  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
352  | 
|
| 10693 | 353  | 
fun number_of(n,T) = HOLogic.number_of_const T $ (HOLogic.mk_bin n)  | 
354  | 
||
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
355  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
356  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
357  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
358  | 
structure Fast_Arith =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
359  | 
Fast_Lin_Arith(structure LA_Logic=LA_Logic and LA_Data=LA_Data_Ref);  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
360  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
361  | 
val fast_arith_tac = Fast_Arith.lin_arith_tac  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
362  | 
and trace_arith = Fast_Arith.trace;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
363  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
364  | 
local  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
365  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
366  | 
(* reduce contradictory <= to False.  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
367  | 
Most of the work is done by the cancel tactics.  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
368  | 
*)  | 
| 
12931
 
2c0251fada94
solved the problem that Larry's simproce cancle_numerals(?) returns
 
nipkow 
parents: 
12480 
diff
changeset
 | 
369  | 
val add_rules =  | 
| 
 
2c0251fada94
solved the problem that Larry's simproce cancle_numerals(?) returns
 
nipkow 
parents: 
12480 
diff
changeset
 | 
370  | 
[add_0,add_0_right,Zero_not_Suc,Suc_not_Zero,le_0_eq,  | 
| 
 
2c0251fada94
solved the problem that Larry's simproce cancle_numerals(?) returns
 
nipkow 
parents: 
12480 
diff
changeset
 | 
371  | 
One_nat_def];  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
372  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
373  | 
val add_mono_thms_nat = map (fn s => prove_goal (the_context ()) s  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
374  | 
(fn prems => [cut_facts_tac prems 1,  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
375  | 
blast_tac (claset() addIs [add_le_mono]) 1]))  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
376  | 
["(i <= j) & (k <= l) ==> i + k <= j + (l::nat)",  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
377  | 
"(i = j) & (k <= l) ==> i + k <= j + (l::nat)",  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
378  | 
"(i <= j) & (k = l) ==> i + k <= j + (l::nat)",  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
379  | 
"(i = j) & (k = l) ==> i + k = j + (l::nat)"  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
380  | 
];  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
381  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
382  | 
in  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
383  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
384  | 
val init_lin_arith_data =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
385  | 
Fast_Arith.setup @  | 
| 10693 | 386  | 
 [Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset = _} =>
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
387  | 
   {add_mono_thms = add_mono_thms @ add_mono_thms_nat,
 | 
| 10693 | 388  | 
mult_mono_thms = mult_mono_thms,  | 
| 
10574
 
8f98f0301d67
Linear arithmetic now copes with mixed nat/int formulae.
 
nipkow 
parents: 
10516 
diff
changeset
 | 
389  | 
inj_thms = inj_thms,  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
390  | 
lessD = lessD @ [Suc_leI],  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
391  | 
simpset = HOL_basic_ss addsimps add_rules addsimprocs nat_cancel_sums_add}),  | 
| 9593 | 392  | 
  ArithTheoryData.init, arith_discrete ("nat", true)];
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
393  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
394  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
395  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
396  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
397  | 
local  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
398  | 
val nat_arith_simproc_pats =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
399  | 
map (fn s => Thm.read_cterm (Theory.sign_of (the_context ())) (s, HOLogic.boolT))  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
400  | 
["(m::nat) < n","(m::nat) <= n", "(m::nat) = n"];  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
401  | 
in  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
402  | 
val fast_nat_arith_simproc = mk_simproc  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
403  | 
"fast_nat_arith" nat_arith_simproc_pats Fast_Arith.lin_arith_prover;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
404  | 
end;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
405  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
406  | 
(* Because of fast_nat_arith_simproc, the arithmetic solver is really only  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
407  | 
useful to detect inconsistencies among the premises for subgoals which are  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
408  | 
*not* themselves (in)equalities, because the latter activate  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
409  | 
fast_nat_arith_simproc anyway. However, it seems cheaper to activate the  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
410  | 
solver all the time rather than add the additional check. *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
411  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
412  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
413  | 
(* arith proof method *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
414  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
415  | 
(* FIXME: K true should be replaced by a sensible test to speed things up  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
416  | 
in case there are lots of irrelevant terms involved;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
417  | 
elimination of min/max can be optimized:  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
418  | 
(max m n + k <= r) = (m+k <= r & n+k <= r)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
419  | 
(l <= min m n + k) = (l <= m+k & l <= n+k)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
420  | 
*)  | 
| 10516 | 421  | 
local  | 
422  | 
||
423  | 
fun raw_arith_tac i st =  | 
|
| 9593 | 424  | 
refute_tac (K true) (REPEAT o split_tac (#splits (ArithTheoryData.get_sg (Thm.sign_of_thm st))))  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
425  | 
((REPEAT_DETERM o etac linorder_neqE) THEN' fast_arith_tac) i st;  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
426  | 
|
| 10516 | 427  | 
in  | 
428  | 
||
| 11753 | 429  | 
val arith_tac = fast_arith_tac ORELSE' (ObjectLogic.atomize_tac THEN' raw_arith_tac);  | 
| 10516 | 430  | 
|
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
431  | 
fun arith_method prems =  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
432  | 
Method.METHOD (fn facts => HEADGOAL (Method.insert_tac (prems @ facts) THEN' arith_tac));  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
433  | 
|
| 10516 | 434  | 
end;  | 
435  | 
||
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
436  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
437  | 
(* theory setup *)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
438  | 
|
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
439  | 
val arith_setup =  | 
| 10766 | 440  | 
[Simplifier.change_simpset_of (op addsimprocs) nat_cancel_sums] @  | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
441  | 
init_lin_arith_data @  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
442  | 
[Simplifier.change_simpset_of (op addSolver)  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
443  | 
(mk_solver "lin. arith." Fast_Arith.cut_lin_arith_tac),  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
444  | 
Simplifier.change_simpset_of (op addsimprocs) [fast_nat_arith_simproc],  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
445  | 
  Method.add_methods [("arith", (arith_method o #2) oo Method.syntax Args.bang_facts,
 | 
| 
9436
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
446  | 
"decide linear arithmethic")],  | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
447  | 
  Attrib.add_attributes [("arith_split",
 | 
| 
 
62bb04ab4b01
rearranged setup of arithmetic procedures, avoiding global reference values;
 
wenzelm 
parents:  
diff
changeset
 | 
448  | 
(Attrib.no_args arith_split_add, Attrib.no_args Attrib.undef_local_attribute),  | 
| 9893 | 449  | 
"declaration of split rules for arithmetic procedure")]];  |