| author | wenzelm | 
| Mon, 17 Oct 2005 23:10:13 +0200 | |
| changeset 17876 | b9c92f384109 | 
| parent 17465 | 93fc1211603f | 
| child 18155 | e5ab63ca6b0f | 
| permissions | -rw-r--r-- | 
| 13876 | 1  | 
(* Title: HOL/Integ/presburger.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Amine Chaieb and Stefan Berghofer, TU Muenchen  | 
|
4  | 
||
| 17465 | 5  | 
Tactic for solving arithmetical Goals in Presburger Arithmetic.  | 
| 13876 | 6  | 
|
| 17465 | 7  | 
This version of presburger deals with occurences of functional symbols  | 
8  | 
in the subgoal and abstract over them to try to prove the more general  | 
|
9  | 
formula. It then resolves with the subgoal. To enable this feature  | 
|
10  | 
call the procedure with the parameter abs.  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
11  | 
*)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
12  | 
|
| 13876 | 13  | 
signature PRESBURGER =  | 
14  | 
sig  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
15  | 
val presburger_tac : bool -> bool -> int -> tactic  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
16  | 
val presburger_method : bool -> bool -> int -> Proof.method  | 
| 13876 | 17  | 
val setup : (theory -> theory) list  | 
18  | 
val trace : bool ref  | 
|
19  | 
end;  | 
|
20  | 
||
21  | 
structure Presburger: PRESBURGER =  | 
|
22  | 
struct  | 
|
23  | 
||
24  | 
val trace = ref false;  | 
|
25  | 
fun trace_msg s = if !trace then tracing s else ();  | 
|
26  | 
||
27  | 
(*-----------------------------------------------------------------*)  | 
|
28  | 
(*cooper_pp: provefunction for the one-exstance quantifier elimination*)  | 
|
29  | 
(* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*)  | 
|
30  | 
(*-----------------------------------------------------------------*)  | 
|
31  | 
||
| 14941 | 32  | 
|
| 
14801
 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 
berghofe 
parents: 
14758 
diff
changeset
 | 
33  | 
val presburger_ss = simpset_of (theory "Presburger");  | 
| 
 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 
berghofe 
parents: 
14758 
diff
changeset
 | 
34  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
35  | 
fun cooper_pp sg (fm as e$Abs(xn,xT,p)) =  | 
| 13876 | 36  | 
let val (xn1,p1) = variant_abs (xn,xT,p)  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
37  | 
in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1) end;  | 
| 13876 | 38  | 
|
39  | 
fun mnnf_pp sg fm = CooperProof.proof_of_cnnf sg fm  | 
|
40  | 
(CooperProof.proof_of_evalc sg);  | 
|
41  | 
||
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
42  | 
fun tmproof_of_int_qelim sg fm =  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
43  | 
Qelim.tproof_of_mlift_qelim sg CooperDec.is_arith_rel  | 
| 13876 | 44  | 
(CooperProof.proof_of_linform sg) (mnnf_pp sg) (cooper_pp sg) fm;  | 
45  | 
||
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
46  | 
|
| 13876 | 47  | 
(* Theorems to be used in this tactic*)  | 
48  | 
||
49  | 
val zdvd_int = thm "zdvd_int";  | 
|
50  | 
val zdiff_int_split = thm "zdiff_int_split";  | 
|
51  | 
val all_nat = thm "all_nat";  | 
|
52  | 
val ex_nat = thm "ex_nat";  | 
|
53  | 
val number_of1 = thm "number_of1";  | 
|
54  | 
val number_of2 = thm "number_of2";  | 
|
55  | 
val split_zdiv = thm "split_zdiv";  | 
|
56  | 
val split_zmod = thm "split_zmod";  | 
|
57  | 
val mod_div_equality' = thm "mod_div_equality'";  | 
|
58  | 
val split_div' = thm "split_div'";  | 
|
59  | 
val Suc_plus1 = thm "Suc_plus1";  | 
|
60  | 
val imp_le_cong = thm "imp_le_cong";  | 
|
61  | 
val conj_le_cong = thm "conj_le_cong";  | 
|
62  | 
||
63  | 
(* extract all the constants in a term*)  | 
|
64  | 
fun add_term_typed_consts (Const (c, T), cs) = (c,T) ins cs  | 
|
65  | 
| add_term_typed_consts (t $ u, cs) =  | 
|
66  | 
add_term_typed_consts (t, add_term_typed_consts (u, cs))  | 
|
67  | 
| add_term_typed_consts (Abs (_, _, t), cs) = add_term_typed_consts (t, cs)  | 
|
68  | 
| add_term_typed_consts (_, cs) = cs;  | 
|
69  | 
||
70  | 
fun term_typed_consts t = add_term_typed_consts(t,[]);  | 
|
71  | 
||
| 
15661
 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 
wenzelm 
parents: 
15620 
diff
changeset
 | 
72  | 
(* Some Types*)  | 
| 13876 | 73  | 
val bT = HOLogic.boolT;  | 
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15574 
diff
changeset
 | 
74  | 
val bitT = HOLogic.bitT;  | 
| 13876 | 75  | 
val iT = HOLogic.intT;  | 
76  | 
val binT = HOLogic.binT;  | 
|
77  | 
val nT = HOLogic.natT;  | 
|
78  | 
||
79  | 
(* Allowed Consts in formulae for presburger tactic*)  | 
|
80  | 
||
81  | 
val allowed_consts =  | 
|
82  | 
  [("All", (iT --> bT) --> bT),
 | 
|
83  | 
   ("Ex", (iT --> bT) --> bT),
 | 
|
84  | 
   ("All", (nT --> bT) --> bT),
 | 
|
85  | 
   ("Ex", (nT --> bT) --> bT),
 | 
|
86  | 
||
87  | 
   ("op &", bT --> bT --> bT),
 | 
|
88  | 
   ("op |", bT --> bT --> bT),
 | 
|
89  | 
   ("op -->", bT --> bT --> bT),
 | 
|
90  | 
   ("op =", bT --> bT --> bT),
 | 
|
91  | 
   ("Not", bT --> bT),
 | 
|
92  | 
||
93  | 
   ("op <=", iT --> iT --> bT),
 | 
|
94  | 
   ("op =", iT --> iT --> bT),
 | 
|
95  | 
   ("op <", iT --> iT --> bT),
 | 
|
96  | 
   ("Divides.op dvd", iT --> iT --> bT),
 | 
|
97  | 
   ("Divides.op div", iT --> iT --> iT),
 | 
|
98  | 
   ("Divides.op mod", iT --> iT --> iT),
 | 
|
99  | 
   ("op +", iT --> iT --> iT),
 | 
|
100  | 
   ("op -", iT --> iT --> iT),
 | 
|
101  | 
   ("op *", iT --> iT --> iT), 
 | 
|
102  | 
   ("HOL.abs", iT --> iT),
 | 
|
103  | 
   ("uminus", iT --> iT),
 | 
|
| 
14801
 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 
berghofe 
parents: 
14758 
diff
changeset
 | 
104  | 
   ("HOL.max", iT --> iT --> iT),
 | 
| 
 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 
berghofe 
parents: 
14758 
diff
changeset
 | 
105  | 
   ("HOL.min", iT --> iT --> iT),
 | 
| 13876 | 106  | 
|
107  | 
   ("op <=", nT --> nT --> bT),
 | 
|
108  | 
   ("op =", nT --> nT --> bT),
 | 
|
109  | 
   ("op <", nT --> nT --> bT),
 | 
|
110  | 
   ("Divides.op dvd", nT --> nT --> bT),
 | 
|
111  | 
   ("Divides.op div", nT --> nT --> nT),
 | 
|
112  | 
   ("Divides.op mod", nT --> nT --> nT),
 | 
|
113  | 
   ("op +", nT --> nT --> nT),
 | 
|
114  | 
   ("op -", nT --> nT --> nT),
 | 
|
115  | 
   ("op *", nT --> nT --> nT), 
 | 
|
116  | 
   ("Suc", nT --> nT),
 | 
|
| 
14801
 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 
berghofe 
parents: 
14758 
diff
changeset
 | 
117  | 
   ("HOL.max", nT --> nT --> nT),
 | 
| 
 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 
berghofe 
parents: 
14758 
diff
changeset
 | 
118  | 
   ("HOL.min", nT --> nT --> nT),
 | 
| 13876 | 119  | 
|
| 
15620
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15574 
diff
changeset
 | 
120  | 
   ("Numeral.bit.B0", bitT),
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15574 
diff
changeset
 | 
121  | 
   ("Numeral.bit.B1", bitT),
 | 
| 
 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 
paulson 
parents: 
15574 
diff
changeset
 | 
122  | 
   ("Numeral.Bit", binT --> bitT --> binT),
 | 
| 15013 | 123  | 
   ("Numeral.Pls", binT),
 | 
124  | 
   ("Numeral.Min", binT),
 | 
|
| 13876 | 125  | 
   ("Numeral.number_of", binT --> iT),
 | 
126  | 
   ("Numeral.number_of", binT --> nT),
 | 
|
127  | 
   ("0", nT),
 | 
|
128  | 
   ("0", iT),
 | 
|
129  | 
   ("1", nT),
 | 
|
130  | 
   ("1", iT),
 | 
|
131  | 
   ("False", bT),
 | 
|
132  | 
   ("True", bT)];
 | 
|
133  | 
||
134  | 
(* Preparation of the formula to be sent to the Integer quantifier *)  | 
|
135  | 
(* elimination procedure *)  | 
|
136  | 
(* Transforms meta implications and meta quantifiers to object *)  | 
|
137  | 
(* implications and object quantifiers *)  | 
|
138  | 
||
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
139  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
140  | 
(*==================================*)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
141  | 
(* Abstracting on subterms ========*)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
142  | 
(*==================================*)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
143  | 
(* Returns occurences of terms that are function application of type int or nat*)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
144  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
145  | 
fun getfuncs fm = case strip_comb fm of  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
146  | 
(Free (_, T), ts as _ :: _) =>  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
147  | 
if body_type T mem [iT, nT]  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
148  | 
andalso not (ts = []) andalso forall (null o loose_bnos) ts  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
149  | 
then [fm]  | 
| 15570 | 150  | 
else Library.foldl op union ([], map getfuncs ts)  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
151  | 
| (Var (_, T), ts as _ :: _) =>  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
152  | 
if body_type T mem [iT, nT]  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
153  | 
andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm]  | 
| 15570 | 154  | 
else Library.foldl op union ([], map getfuncs ts)  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
155  | 
| (Const (s, T), ts) =>  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
156  | 
if (s, T) mem allowed_consts orelse not (body_type T mem [iT, nT])  | 
| 15570 | 157  | 
then Library.foldl op union ([], map getfuncs ts)  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
158  | 
else [fm]  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
159  | 
| (Abs (s, T, t), _) => getfuncs t  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
160  | 
| _ => [];  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
161  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
162  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
163  | 
fun abstract_pres sg fm =  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
164  | 
foldr (fn (t, u) =>  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
165  | 
let val T = fastype_of t  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
166  | 
      in all T $ Abs ("x", T, abstract_over (t, u)) end)
 | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
167  | 
fm (getfuncs fm);  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
168  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
169  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
170  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
171  | 
(* hasfuncs_on_bounds dont care of the type of the functions applied!  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
172  | 
It returns true if there is a subterm coresponding to the application of  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
173  | 
a function on a bounded variable.  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
174  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
175  | 
Function applications are allowed only for well predefined functions a  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
176  | 
consts*)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
177  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
178  | 
fun has_free_funcs fm = case strip_comb fm of  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
179  | 
(Free (_, T), ts as _ :: _) =>  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
180  | 
if (body_type T mem [iT,nT]) andalso (not (T mem [iT,nT]))  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
181  | 
then true  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
182  | 
else exists (fn x => x) (map has_free_funcs ts)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
183  | 
| (Var (_, T), ts as _ :: _) =>  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
184  | 
if (body_type T mem [iT,nT]) andalso not (T mem [iT,nT])  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
185  | 
then true  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
186  | 
else exists (fn x => x) (map has_free_funcs ts)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
187  | 
| (Const (s, T), ts) => exists (fn x => x) (map has_free_funcs ts)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
188  | 
| (Abs (s, T, t), _) => has_free_funcs t  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
189  | 
|_ => false;  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
190  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
191  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
192  | 
(*returns true if the formula is relevant for presburger arithmetic tactic  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
193  | 
The constants occuring in term t should be a subset of the allowed_consts  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
194  | 
There also should be no occurences of application of functions on bounded  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
195  | 
variables. Whenever this function will be used, it will be ensured that t  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
196  | 
will not contain subterms with function symbols that could have been  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
197  | 
abstracted over.*)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
198  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
199  | 
fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso  | 
| 15570 | 200  | 
map (fn i => snd (List.nth (ps, i))) (loose_bnos t) @  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
201  | 
map (snd o dest_Free) (term_frees t) @ map (snd o dest_Var) (term_vars t)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
202  | 
subset [iT, nT]  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
203  | 
andalso not (has_free_funcs t);  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
204  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
205  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
206  | 
fun prepare_for_presburger sg q fm =  | 
| 13876 | 207  | 
let  | 
208  | 
val ps = Logic.strip_params fm  | 
|
209  | 
val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm)  | 
|
210  | 
val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm)  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
211  | 
val _ = if relevant (rev ps) c then ()  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
212  | 
               else  (trace_msg ("Conclusion is not a presburger term:\n" ^
 | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
213  | 
Sign.string_of_term sg c); raise CooperDec.COOPER)  | 
| 13876 | 214  | 
fun mk_all ((s, T), (P,n)) =  | 
215  | 
if 0 mem loose_bnos P then  | 
|
216  | 
(HOLogic.all_const T $ Abs (s, T, P), n)  | 
|
217  | 
else (incr_boundvars ~1 P, n-1)  | 
|
218  | 
fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t;  | 
|
| 15570 | 219  | 
val (rhs,irhs) = List.partition (relevant (rev ps)) hs  | 
| 13876 | 220  | 
val np = length ps  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
221  | 
val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))  | 
| 
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
222  | 
(foldr HOLogic.mk_imp c rhs, np) ps  | 
| 15570 | 223  | 
val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT)  | 
| 13876 | 224  | 
(term_frees fm' @ term_vars fm');  | 
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
225  | 
val fm2 = foldr mk_all2 fm' vs  | 
| 13876 | 226  | 
in (fm2, np + length vs, length rhs) end;  | 
227  | 
||
228  | 
(*Object quantifier to meta --*)  | 
|
229  | 
fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ;  | 
|
230  | 
||
231  | 
(* object implication to meta---*)  | 
|
232  | 
fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp;  | 
|
233  | 
||
234  | 
(* the presburger tactic*)  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
235  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
236  | 
(* Parameters : q = flag for quantify ofer free variables ;  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
237  | 
a = flag for abstracting over function occurences  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
238  | 
i = subgoal *)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
239  | 
|
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
240  | 
fun presburger_tac q a i = ObjectLogic.atomize_tac i THEN (fn st =>  | 
| 13876 | 241  | 
let  | 
| 17465 | 242  | 
val g = List.nth (prems_of st, i - 1)  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
243  | 
val sg = sign_of_thm st  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
244  | 
(* The Abstraction step *)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
245  | 
val g' = if a then abstract_pres sg g else g  | 
| 13876 | 246  | 
(* Transform the term*)  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
247  | 
val (t,np,nh) = prepare_for_presburger sg q g'  | 
| 
15661
 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 
wenzelm 
parents: 
15620 
diff
changeset
 | 
248  | 
(* Some simpsets for dealing with mod div abs and nat*)  | 
| 13876 | 249  | 
val simpset0 = HOL_basic_ss  | 
250  | 
addsimps [mod_div_equality', Suc_plus1]  | 
|
| 13997 | 251  | 
addsplits [split_zdiv, split_zmod, split_div', split_min, split_max]  | 
| 13876 | 252  | 
(* Simp rules for changing (n::int) to int n *)  | 
253  | 
val simpset1 = HOL_basic_ss  | 
|
254  | 
addsimps [nat_number_of_def, zdvd_int] @ map (fn r => r RS sym)  | 
|
255  | 
[int_int_eq, zle_int, zless_int, zadd_int, zmult_int]  | 
|
256  | 
addsplits [zdiff_int_split]  | 
|
257  | 
(*simp rules for elimination of int n*)  | 
|
258  | 
||
259  | 
val simpset2 = HOL_basic_ss  | 
|
260  | 
addsimps [nat_0_le, all_nat, ex_nat, number_of1, number_of2, int_0, int_1]  | 
|
261  | 
addcongs [conj_le_cong, imp_le_cong]  | 
|
262  | 
(* simp rules for elimination of abs *)  | 
|
| 
14353
 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 
paulson 
parents: 
14130 
diff
changeset
 | 
263  | 
val simpset3 = HOL_basic_ss addsplits [abs_split]  | 
| 13876 | 264  | 
val ct = cterm_of sg (HOLogic.mk_Trueprop t)  | 
265  | 
(* Theorem for the nat --> int transformation *)  | 
|
266  | 
val pre_thm = Seq.hd (EVERY  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
267  | 
[simp_tac simpset0 1,  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
268  | 
TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1),  | 
| 
14801
 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 
berghofe 
parents: 
14758 
diff
changeset
 | 
269  | 
TRY (simp_tac simpset3 1), TRY (simp_tac presburger_ss 1)]  | 
| 13876 | 270  | 
(trivial ct))  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
271  | 
fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i)  | 
| 13876 | 272  | 
(* The result of the quantifier elimination *)  | 
273  | 
val (th, tac) = case (prop_of pre_thm) of  | 
|
274  | 
        Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
 | 
|
| 
14920
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
275  | 
let val pth =  | 
| 
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
276  | 
(* If quick_and_dirty then run without proof generation as oracle*)  | 
| 
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
277  | 
if !quick_and_dirty  | 
| 16836 | 278  | 
then presburger_oracle sg (Pattern.eta_long [] t1)  | 
| 14941 | 279  | 
(*  | 
280  | 
assume (cterm_of sg  | 
|
| 
14920
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
281  | 
(HOLogic.mk_Trueprop(HOLogic.mk_eq(t1,CooperDec.integer_qelim (Pattern.eta_long [] t1)))))  | 
| 14941 | 282  | 
*)  | 
| 
14920
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
283  | 
else tmproof_of_int_qelim sg (Pattern.eta_long [] t1)  | 
| 
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
284  | 
in  | 
| 13876 | 285  | 
          (trace_msg ("calling procedure with term:\n" ^
 | 
286  | 
Sign.string_of_term sg t1);  | 
|
| 
14920
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
287  | 
((pth RS iffD2) RS pre_thm,  | 
| 13876 | 288  | 
assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i)))  | 
| 
14920
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
289  | 
end  | 
| 13876 | 290  | 
| _ => (pre_thm, assm_tac i)  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
291  | 
in (rtac (((mp_step nh) o (spec_step np)) th) i  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
292  | 
THEN tac) st  | 
| 14130 | 293  | 
end handle Subscript => no_tac st | CooperDec.COOPER => no_tac st);  | 
| 13876 | 294  | 
|
295  | 
fun presburger_args meth =  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
296  | 
let val parse_flag =  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
297  | 
Args.$$$ "no_quantify" >> K (apfst (K false))  | 
| 
14811
 
9144ec118703
abstraction over functions is no default any more.
 
chaieb 
parents: 
14801 
diff
changeset
 | 
298  | 
|| Args.$$$ "abs" >> K (apsnd (K true));  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
299  | 
in  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
300  | 
Method.simple_args  | 
| 14882 | 301  | 
  (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >>
 | 
| 15570 | 302  | 
curry (Library.foldl op |>) (true, false))  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
303  | 
(fn (q,a) => fn _ => meth q a 1)  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
304  | 
end;  | 
| 13876 | 305  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
306  | 
fun presburger_method q a i = Method.METHOD (fn facts =>  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
14353 
diff
changeset
 | 
307  | 
Method.insert_tac facts 1 THEN presburger_tac q a i)  | 
| 13876 | 308  | 
|
309  | 
val setup =  | 
|
310  | 
  [Method.add_method ("presburger",
 | 
|
311  | 
presburger_args presburger_method,  | 
|
312  | 
"decision procedure for Presburger arithmetic"),  | 
|
313  | 
   ArithTheoryData.map (fn {splits, inj_consts, discrete, presburger} =>
 | 
|
314  | 
     {splits = splits, inj_consts = inj_consts, discrete = discrete,
 | 
|
| 15531 | 315  | 
presburger = SOME (presburger_tac true false)})];  | 
| 13876 | 316  | 
|
317  | 
end;  | 
|
318  | 
||
| 
14920
 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 
chaieb 
parents: 
14882 
diff
changeset
 | 
319  | 
val presburger_tac = Presburger.presburger_tac true false;  |