| author | haftmann | 
| Tue, 28 Jun 2005 11:58:56 +0200 | |
| changeset 16576 | 9ce0be075e6a | 
| parent 16398 | 7f0faa30f602 | 
| child 16837 | 416e86088931 | 
| permissions | -rw-r--r-- | 
| 13876 | 1  | 
(* Title: HOL/Integ/cooper_dec.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Amine Chaieb and Tobias Nipkow, TU Muenchen  | 
|
4  | 
||
5  | 
File containing the implementation of Cooper Algorithm  | 
|
6  | 
decision procedure (intensively inspired from J.Harrison)  | 
|
7  | 
*)  | 
|
| 
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: 
14877 
diff
changeset
 | 
8  | 
|
| 
 
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: 
14877 
diff
changeset
 | 
9  | 
|
| 13876 | 10  | 
signature COOPER_DEC =  | 
11  | 
sig  | 
|
12  | 
exception COOPER  | 
|
| 14941 | 13  | 
exception COOPER_ORACLE of term  | 
| 13876 | 14  | 
val is_arith_rel : term -> bool  | 
| 16389 | 15  | 
val mk_numeral : IntInf.int -> term  | 
16  | 
val dest_numeral : term -> IntInf.int  | 
|
| 15164 | 17  | 
val is_numeral : term -> bool  | 
| 13876 | 18  | 
val zero : term  | 
19  | 
val one : term  | 
|
| 16389 | 20  | 
val linear_cmul : IntInf.int -> term -> term  | 
| 13876 | 21  | 
val linear_add : string list -> term -> term -> term  | 
22  | 
val linear_sub : string list -> term -> term -> term  | 
|
23  | 
val linear_neg : term -> term  | 
|
24  | 
val lint : string list -> term -> term  | 
|
25  | 
val linform : string list -> term -> term  | 
|
| 16389 | 26  | 
val formlcm : term -> term -> IntInf.int  | 
27  | 
val adjustcoeff : term -> IntInf.int -> term -> term  | 
|
| 13876 | 28  | 
val unitycoeff : term -> term -> term  | 
| 16389 | 29  | 
val divlcm : term -> term -> IntInf.int  | 
| 13876 | 30  | 
val bset : term -> term -> term list  | 
31  | 
val aset : term -> term -> term list  | 
|
32  | 
val linrep : string list -> term -> term -> term -> term  | 
|
33  | 
val list_disj : term list -> 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: 
13876 
diff
changeset
 | 
34  | 
val list_conj : term list -> term  | 
| 13876 | 35  | 
val simpl : term -> term  | 
36  | 
val fv : term -> string list  | 
|
37  | 
val negate : term -> term  | 
|
| 16389 | 38  | 
val operations : (string * (IntInf.int * IntInf.int -> bool)) list  | 
| 
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: 
13876 
diff
changeset
 | 
39  | 
val conjuncts : term -> term list  | 
| 
 
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: 
13876 
diff
changeset
 | 
40  | 
val disjuncts : term -> term list  | 
| 
 
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: 
13876 
diff
changeset
 | 
41  | 
val has_bound : term -> bool  | 
| 
 
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: 
13876 
diff
changeset
 | 
42  | 
val minusinf : term -> term -> term  | 
| 
 
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: 
13876 
diff
changeset
 | 
43  | 
val plusinf : term -> term -> term  | 
| 14877 | 44  | 
val onatoms : (term -> term) -> term -> term  | 
45  | 
val evalc : term -> term  | 
|
| 15107 | 46  | 
val cooper_w : string list -> term -> (term option * term)  | 
| 
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: 
14877 
diff
changeset
 | 
47  | 
val integer_qelim : Term.term -> Term.term  | 
| 14941 | 48  | 
val mk_presburger_oracle : (Sign.sg * exn) -> term  | 
| 13876 | 49  | 
end;  | 
50  | 
||
51  | 
structure CooperDec : COOPER_DEC =  | 
|
52  | 
struct  | 
|
53  | 
||
54  | 
(* ========================================================================= *)  | 
|
55  | 
(* Cooper's algorithm for Presburger arithmetic. *)  | 
|
56  | 
(* ========================================================================= *)  | 
|
57  | 
exception COOPER;  | 
|
58  | 
||
| 14941 | 59  | 
(* Exception for the oracle *)  | 
60  | 
exception COOPER_ORACLE of term;  | 
|
61  | 
||
62  | 
||
| 13876 | 63  | 
(* ------------------------------------------------------------------------- *)  | 
64  | 
(* Lift operations up to numerals. *)  | 
|
65  | 
(* ------------------------------------------------------------------------- *)  | 
|
66  | 
||
67  | 
(*Assumption : The construction of atomar formulas in linearl arithmetic is based on  | 
|
| 16389 | 68  | 
relation operations of Type : [IntInf.int,IntInf.int]---> bool *)  | 
| 13876 | 69  | 
|
70  | 
(* ------------------------------------------------------------------------- *)  | 
|
71  | 
||
72  | 
(*Function is_arith_rel returns true if and only if the term is an atomar presburger  | 
|
73  | 
formula *)  | 
|
74  | 
fun is_arith_rel tm = case tm of  | 
|
75  | 
	 Const(p,Type ("fun",[Type ("Numeral.bin", []),Type ("fun",[Type ("Numeral.bin", 
 | 
|
76  | 
	 []),Type ("bool",[])] )])) $ _ $_ => true 
 | 
|
77  | 
	|Const(p,Type ("fun",[Type ("IntDef.int", []),Type ("fun",[Type ("IntDef.int", 
 | 
|
78  | 
	 []),Type ("bool",[])] )])) $ _ $_ => true 
 | 
|
79  | 
|_ => false;  | 
|
80  | 
||
81  | 
(*Function is_arith_rel returns true if and only if the term is an operation of the  | 
|
82  | 
form [int,int]---> int*)  | 
|
83  | 
||
84  | 
(*Transform a natural number to a term*)  | 
|
85  | 
||
86  | 
fun mk_numeral 0 = Const("0",HOLogic.intT)
 | 
|
87  | 
   |mk_numeral 1 = Const("1",HOLogic.intT)
 | 
|
| 16389 | 88  | 
|mk_numeral n = (HOLogic.number_of_const HOLogic.intT) $ (HOLogic.mk_bin n);  | 
| 13876 | 89  | 
|
90  | 
(*Transform an Term to an natural number*)  | 
|
91  | 
||
92  | 
fun dest_numeral (Const("0",Type ("IntDef.int", []))) = 0
 | 
|
93  | 
   |dest_numeral (Const("1",Type ("IntDef.int", []))) = 1
 | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
94  | 
   |dest_numeral (Const ("Numeral.number_of",_) $ n) = 
 | 
| 16389 | 95  | 
HOLogic.dest_binum n;  | 
| 
15661
 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 
wenzelm 
parents: 
15574 
diff
changeset
 | 
96  | 
(*Some terms often used for pattern matching*)  | 
| 13876 | 97  | 
|
98  | 
val zero = mk_numeral 0;  | 
|
99  | 
val one = mk_numeral 1;  | 
|
100  | 
||
101  | 
(*Tests if a Term is representing a number*)  | 
|
102  | 
||
103  | 
fun is_numeral t = (t = zero) orelse (t = one) orelse (can dest_numeral t);  | 
|
104  | 
||
105  | 
(*maps a unary natural function on a term containing an natural number*)  | 
|
106  | 
||
107  | 
fun numeral1 f n = mk_numeral (f(dest_numeral n));  | 
|
108  | 
||
109  | 
(*maps a binary natural function on 2 term containing natural numbers*)  | 
|
110  | 
||
111  | 
fun numeral2 f m n = mk_numeral(f(dest_numeral m) (dest_numeral n));  | 
|
112  | 
||
113  | 
(* ------------------------------------------------------------------------- *)  | 
|
114  | 
(* Operations on canonical linear terms c1 * x1 + ... + cn * xn + k *)  | 
|
115  | 
(* *)  | 
|
116  | 
(* Note that we're quite strict: the ci must be present even if 1 *)  | 
|
117  | 
(* (but if 0 we expect the monomial to be omitted) and k must be there *)  | 
|
118  | 
(* even if it's zero. Thus, it's a constant iff not an addition term. *)  | 
|
| 16389 | 119  | 
(* ------------------------------------------------------------------------- *)  | 
| 13876 | 120  | 
|
121  | 
||
122  | 
fun linear_cmul n tm = if n = 0 then zero else let fun times n k = n*k in  | 
|
123  | 
( case tm of  | 
|
124  | 
     (Const("op +",T)  $  (Const ("op *",T1 ) $c1 $  x1) $ rest) => 
 | 
|
125  | 
       Const("op +",T) $ ((Const("op *",T1) $ (numeral1 (times n) c1) $ x1)) $ (linear_cmul n rest) 
 | 
|
126  | 
|_ => numeral1 (times n) tm)  | 
|
127  | 
end ;  | 
|
128  | 
||
129  | 
||
130  | 
||
131  | 
||
132  | 
(* Whether the first of two items comes earlier in the list *)  | 
|
133  | 
fun earlier [] x y = false  | 
|
134  | 
|earlier (h::t) x y =if h = y then false  | 
|
135  | 
else if h = x then true  | 
|
136  | 
else earlier t x y ;  | 
|
137  | 
||
138  | 
fun earlierv vars (Bound i) (Bound j) = i < j  | 
|
139  | 
|earlierv vars (Bound _) _ = true  | 
|
140  | 
|earlierv vars _ (Bound _) = false  | 
|
141  | 
|earlierv vars (Free (x,_)) (Free (y,_)) = earlier vars x y;  | 
|
142  | 
||
143  | 
||
144  | 
fun linear_add vars tm1 tm2 =  | 
|
145  | 
let fun addwith x y = x + y in  | 
|
146  | 
(case (tm1,tm2) of  | 
|
147  | 
	((Const ("op +",T1) $ ( Const("op *",T2) $ c1 $  x1) $ rest1),(Const 
 | 
|
148  | 
	("op +",T3)$( Const("op *",T4) $ c2 $  x2) $ rest2)) => 
 | 
|
149  | 
if x1 = x2 then  | 
|
150  | 
let val c = (numeral2 (addwith) c1 c2)  | 
|
151  | 
in  | 
|
152  | 
if c = zero then (linear_add vars rest1 rest2)  | 
|
153  | 
	      else (Const("op +",T1) $ (Const("op *",T2) $ c $ x1) $ (linear_add vars  rest1 rest2)) 
 | 
|
154  | 
end  | 
|
155  | 
else  | 
|
156  | 
		if earlierv vars x1 x2 then (Const("op +",T1) $  
 | 
|
157  | 
		(Const("op *",T2)$ c1 $ x1) $ (linear_add vars rest1 tm2)) 
 | 
|
158  | 
    	       else (Const("op +",T1) $ (Const("op *",T2) $ c2 $ x2) $ (linear_add vars tm1 rest2)) 
 | 
|
159  | 
   	|((Const("op +",T1) $ (Const("op *",T2) $ c1 $ x1) $ rest1) ,_) => 
 | 
|
160  | 
    	  (Const("op +",T1)$ (Const("op *",T2) $ c1 $ x1) $ (linear_add vars 
 | 
|
161  | 
rest1 tm2))  | 
|
162  | 
   	|(_, (Const("op +",T1) $(Const("op *",T2) $ c2 $ x2) $ rest2)) => 
 | 
|
163  | 
      	  (Const("op +",T1) $ (Const("op *",T2) $ c2 $ x2) $ (linear_add vars tm1 
 | 
|
164  | 
rest2))  | 
|
165  | 
| (_,_) => numeral2 (addwith) tm1 tm2)  | 
|
166  | 
||
167  | 
end;  | 
|
168  | 
||
169  | 
(*To obtain the unary - applyed on a formula*)  | 
|
170  | 
||
171  | 
fun linear_neg tm = linear_cmul (0 - 1) tm;  | 
|
172  | 
||
173  | 
(*Substraction of two terms *)  | 
|
174  | 
||
175  | 
fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2);  | 
|
176  | 
||
177  | 
||
178  | 
(* ------------------------------------------------------------------------- *)  | 
|
179  | 
(* Linearize a term. *)  | 
|
180  | 
(* ------------------------------------------------------------------------- *)  | 
|
181  | 
||
182  | 
(* linearises a term from the point of view of Variable Free (x,T).  | 
|
183  | 
After this fuction the all expressions containig ths variable will have the form  | 
|
184  | 
c*Free(x,T) + t where c is a constant ant t is a Term which is not containing  | 
|
185  | 
Free(x,T)*)  | 
|
186  | 
||
187  | 
fun lint vars tm = if is_numeral tm then tm else case tm of  | 
|
188  | 
(Free (x,T)) => (HOLogic.mk_binop "op +" ((HOLogic.mk_binop "op *" ((mk_numeral 1),Free (x,T))), zero))  | 
|
189  | 
  |(Bound i) =>  (Const("op +",HOLogic.intT -->HOLogic.intT -->HOLogic.intT) $ 
 | 
|
190  | 
  (Const("op *",HOLogic.intT -->HOLogic.intT -->HOLogic.intT) $ (mk_numeral 1) $ (Bound i)) $ zero) 
 | 
|
191  | 
  |(Const("uminus",_) $ t ) => (linear_neg (lint vars t)) 
 | 
|
192  | 
  |(Const("op +",_) $ s $ t) => (linear_add vars (lint vars s) (lint vars t)) 
 | 
|
193  | 
  |(Const("op -",_) $ s $ t) => (linear_sub vars (lint vars s) (lint vars t)) 
 | 
|
194  | 
  |(Const ("op *",_) $ s $ t) => 
 | 
|
195  | 
let val s' = lint vars s  | 
|
196  | 
val t' = lint vars t  | 
|
197  | 
in  | 
|
198  | 
if is_numeral s' then (linear_cmul (dest_numeral s') t')  | 
|
199  | 
else if is_numeral t' then (linear_cmul (dest_numeral t') s')  | 
|
200  | 
||
| 
16299
 
872ad146bb14
Some error messages have been eliminated as suggested by Tobias Nipkow
 
chaieb 
parents: 
15965 
diff
changeset
 | 
201  | 
else raise COOPER  | 
| 13876 | 202  | 
end  | 
| 
16299
 
872ad146bb14
Some error messages have been eliminated as suggested by Tobias Nipkow
 
chaieb 
parents: 
15965 
diff
changeset
 | 
203  | 
|_ => raise COOPER;  | 
| 13876 | 204  | 
|
205  | 
||
206  | 
||
207  | 
(* ------------------------------------------------------------------------- *)  | 
|
208  | 
(* Linearize the atoms in a formula, and eliminate non-strict inequalities. *)  | 
|
209  | 
(* ------------------------------------------------------------------------- *)  | 
|
210  | 
||
211  | 
fun mkatom vars p t = Const(p,HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ zero $ (lint vars t);  | 
|
212  | 
||
| 15164 | 213  | 
fun linform vars (Const ("Divides.op dvd",_) $ c $ t) =
 | 
214  | 
if is_numeral c then  | 
|
| 13876 | 215  | 
let val c' = (mk_numeral(abs(dest_numeral c)))  | 
216  | 
in (HOLogic.mk_binrel "Divides.op dvd" (c,lint vars t))  | 
|
217  | 
end  | 
|
| 15164 | 218  | 
else (warning "Nonlinear term --- Non numeral leftside at dvd"  | 
219  | 
;raise COOPER)  | 
|
| 13876 | 220  | 
  |linform vars  (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ s $ t ) = (mkatom vars "op =" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ t $ s) ) 
 | 
221  | 
  |linform vars  (Const("op <",_)$ s $t ) = (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ t $ s))
 | 
|
222  | 
  |linform vars  (Const("op >",_) $ s $ t ) = (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ s $ t)) 
 | 
|
223  | 
  |linform vars  (Const("op <=",_)$ s $ t ) = 
 | 
|
224  | 
        (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ (Const("op +",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $t $(mk_numeral 1)) $ s)) 
 | 
|
225  | 
  |linform vars  (Const("op >=",_)$ s $ t ) = 
 | 
|
226  | 
        (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> 
 | 
|
227  | 
	HOLogic.intT) $ (Const("op +",HOLogic.intT --> HOLogic.intT --> 
 | 
|
228  | 
HOLogic.intT) $s $(mk_numeral 1)) $ t))  | 
|
229  | 
||
230  | 
|linform vars fm = fm;  | 
|
231  | 
||
232  | 
(* ------------------------------------------------------------------------- *)  | 
|
233  | 
(* Post-NNF transformation eliminating negated inequalities. *)  | 
|
234  | 
(* ------------------------------------------------------------------------- *)  | 
|
235  | 
||
236  | 
fun posineq fm = case fm of  | 
|
237  | 
 (Const ("Not",_)$(Const("op <",_)$ c $ t)) =>
 | 
|
238  | 
(HOLogic.mk_binrel "op <" (zero , (linear_sub [] (mk_numeral 1) (linear_add [] c t ) )))  | 
|
239  | 
  | ( Const ("op &",_) $ p $ q)  => HOLogic.mk_conj (posineq p,posineq q)
 | 
|
240  | 
  | ( Const ("op |",_) $ p $ q ) => HOLogic.mk_disj (posineq p,posineq q)
 | 
|
241  | 
| _ => fm;  | 
|
242  | 
||
243  | 
||
244  | 
(* ------------------------------------------------------------------------- *)  | 
|
245  | 
(* Find the LCM of the coefficients of x. *)  | 
|
246  | 
(* ------------------------------------------------------------------------- *)  | 
|
247  | 
(*gcd calculates gcd (a,b) and helps lcm_num calculating lcm (a,b)*)  | 
|
248  | 
||
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
249  | 
(*BEWARE: replaces Library.gcd!! There is also Library.lcm!*)  | 
| 16398 | 250  | 
fun gcd (a:IntInf.int) b = if a=0 then b else gcd (b mod a) a ;  | 
| 13876 | 251  | 
fun lcm_num a b = (abs a*b) div (gcd (abs a) (abs b));  | 
252  | 
||
253  | 
fun formlcm x fm = case fm of  | 
|
254  | 
    (Const (p,_)$ _ $(Const ("op +", _)$(Const ("op *",_)$ c $ y ) $z ) ) =>  if 
 | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
255  | 
(is_arith_rel fm) andalso (x = y) then (abs(dest_numeral c)) else 1  | 
| 13876 | 256  | 
  | ( Const ("Not", _) $p) => formlcm x p 
 | 
257  | 
  | ( Const ("op &",_) $ p $ q) => lcm_num (formlcm x p) (formlcm x q) 
 | 
|
258  | 
  | ( Const ("op |",_) $ p $ q )=> lcm_num (formlcm x p) (formlcm x q) 
 | 
|
259  | 
| _ => 1;  | 
|
260  | 
||
261  | 
(* ------------------------------------------------------------------------- *)  | 
|
262  | 
(* Adjust all coefficients of x in formula; fold in reduction to +/- 1. *)  | 
|
263  | 
(* ------------------------------------------------------------------------- *)  | 
|
264  | 
||
265  | 
fun adjustcoeff x l fm =  | 
|
266  | 
case fm of  | 
|
267  | 
      (Const(p,_) $d $( Const ("op +", _)$(Const ("op *",_) $ 
 | 
|
268  | 
c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then  | 
|
269  | 
let val m = l div (dest_numeral c)  | 
|
270  | 
val n = (if p = "op <" then abs(m) else m)  | 
|
271  | 
val xtm = HOLogic.mk_binop "op *" ((mk_numeral (m div n)), x)  | 
|
272  | 
in  | 
|
273  | 
(HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) ))))  | 
|
274  | 
end  | 
|
275  | 
else fm  | 
|
276  | 
  |( Const ("Not", _) $ p) => HOLogic.Not $ (adjustcoeff x l p) 
 | 
|
277  | 
  |( Const ("op &",_) $ p $ q) => HOLogic.conj$(adjustcoeff x l p) $(adjustcoeff x l q) 
 | 
|
278  | 
  |( Const ("op |",_) $ p $ q) => HOLogic.disj $(adjustcoeff x l p)$ (adjustcoeff x l q) 
 | 
|
279  | 
|_ => fm;  | 
|
280  | 
||
281  | 
(* ------------------------------------------------------------------------- *)  | 
|
282  | 
(* Hence make coefficient of x one in existential formula. *)  | 
|
283  | 
(* ------------------------------------------------------------------------- *)  | 
|
284  | 
||
285  | 
fun unitycoeff x fm =  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
286  | 
let val l = formlcm x fm  | 
| 13876 | 287  | 
val fm' = adjustcoeff x l fm in  | 
| 15267 | 288  | 
if l = 1 then fm'  | 
289  | 
else  | 
|
| 13876 | 290  | 
let val xp = (HOLogic.mk_binop "op +"  | 
| 15267 | 291  | 
((HOLogic.mk_binop "op *" ((mk_numeral 1), x )), zero))  | 
292  | 
in  | 
|
| 13876 | 293  | 
HOLogic.conj $(HOLogic.mk_binrel "Divides.op dvd" ((mk_numeral l) , xp )) $ (adjustcoeff x l fm)  | 
294  | 
end  | 
|
295  | 
end;  | 
|
296  | 
||
297  | 
(* adjustcoeffeq l fm adjusts the coeffitients c_i of x overall in fm to l*)  | 
|
298  | 
(* Here l must be a multiple of all c_i otherwise the obtained formula is not equivalent*)  | 
|
299  | 
(*  | 
|
300  | 
fun adjustcoeffeq x l fm =  | 
|
301  | 
case fm of  | 
|
302  | 
      (Const(p,_) $d $( Const ("op +", _)$(Const ("op *",_) $ 
 | 
|
303  | 
c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then  | 
|
304  | 
let val m = l div (dest_numeral c)  | 
|
305  | 
val n = (if p = "op <" then abs(m) else m)  | 
|
306  | 
val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x))  | 
|
307  | 
in (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) ))))  | 
|
308  | 
end  | 
|
309  | 
else fm  | 
|
310  | 
  |( Const ("Not", _) $ p) => HOLogic.Not $ (adjustcoeffeq x l p) 
 | 
|
311  | 
  |( Const ("op &",_) $ p $ q) => HOLogic.conj$(adjustcoeffeq x l p) $(adjustcoeffeq x l q) 
 | 
|
312  | 
  |( Const ("op |",_) $ p $ q) => HOLogic.disj $(adjustcoeffeq x l p)$ (adjustcoeffeq x l q) 
 | 
|
313  | 
|_ => fm;  | 
|
314  | 
||
315  | 
||
316  | 
*)  | 
|
317  | 
||
318  | 
(* ------------------------------------------------------------------------- *)  | 
|
319  | 
(* The "minus infinity" version. *)  | 
|
320  | 
(* ------------------------------------------------------------------------- *)  | 
|
321  | 
||
322  | 
fun minusinf x fm = case fm of  | 
|
323  | 
    (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ (c1 ) $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) => 
 | 
|
324  | 
if (is_arith_rel fm) andalso (x=y) andalso (c2 = one) andalso (c1 =zero) then HOLogic.false_const  | 
|
325  | 
else fm  | 
|
326  | 
||
327  | 
  |(Const("op <",_) $ c $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z 
 | 
|
| 15859 | 328  | 
)) => if (x = y)  | 
329  | 
then if (pm1 = one) andalso (c = zero) then HOLogic.false_const  | 
|
330  | 
else if (dest_numeral pm1 = ~1) andalso (c = zero) then HOLogic.true_const  | 
|
331  | 
else error "minusinf : term not in normal form!!!"  | 
|
332  | 
else fm  | 
|
| 13876 | 333  | 
|
334  | 
  |(Const ("Not", _) $ p) => HOLogic.Not $ (minusinf x p) 
 | 
|
335  | 
  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (minusinf x p) $ (minusinf x q) 
 | 
|
336  | 
  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (minusinf x p) $ (minusinf x q) 
 | 
|
337  | 
|_ => fm;  | 
|
338  | 
||
339  | 
(* ------------------------------------------------------------------------- *)  | 
|
340  | 
(* The "Plus infinity" version. *)  | 
|
341  | 
(* ------------------------------------------------------------------------- *)  | 
|
342  | 
||
343  | 
fun plusinf x fm = case fm of  | 
|
344  | 
    (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ (c1 ) $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
 | 
|
345  | 
if (is_arith_rel fm) andalso (x=y) andalso (c2 = one) andalso (c1 =zero) then HOLogic.false_const  | 
|
346  | 
else fm  | 
|
347  | 
||
348  | 
  |(Const("op <",_) $ c $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z
 | 
|
| 15859 | 349  | 
)) => if (x = y)  | 
350  | 
then if (pm1 = one) andalso (c = zero) then HOLogic.true_const  | 
|
351  | 
else if (dest_numeral pm1 = ~1) andalso (c = zero) then HOLogic.false_const  | 
|
352  | 
else error "plusinf : term not in normal form!!!"  | 
|
353  | 
else fm  | 
|
| 13876 | 354  | 
|
355  | 
  |(Const ("Not", _) $ p) => HOLogic.Not $ (plusinf x p)
 | 
|
356  | 
  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (plusinf x p) $ (plusinf x q)
 | 
|
357  | 
  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (plusinf x p) $ (plusinf x q)
 | 
|
358  | 
|_ => fm;  | 
|
359  | 
||
360  | 
(* ------------------------------------------------------------------------- *)  | 
|
361  | 
(* The LCM of all the divisors that involve x. *)  | 
|
362  | 
(* ------------------------------------------------------------------------- *)  | 
|
363  | 
||
364  | 
fun divlcm x (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z ) ) =  
 | 
|
365  | 
if x = y then abs(dest_numeral d) else 1  | 
|
366  | 
  |divlcm x ( Const ("Not", _) $ p) = divlcm x p 
 | 
|
367  | 
  |divlcm x ( Const ("op &",_) $ p $ q) = lcm_num (divlcm x p) (divlcm x q) 
 | 
|
368  | 
  |divlcm x ( Const ("op |",_) $ p $ q ) = lcm_num (divlcm x p) (divlcm x q) 
 | 
|
369  | 
|divlcm x _ = 1;  | 
|
370  | 
||
371  | 
(* ------------------------------------------------------------------------- *)  | 
|
372  | 
(* Construct the B-set. *)  | 
|
373  | 
(* ------------------------------------------------------------------------- *)  | 
|
374  | 
||
375  | 
fun bset x fm = case fm of  | 
|
376  | 
   (Const ("Not", _) $ p) => if (is_arith_rel p) then  
 | 
|
377  | 
(case p of  | 
|
378  | 
	      (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $c2 $y) $a ) )  
 | 
|
379  | 
=> if (is_arith_rel p) andalso (x= y) andalso (c2 = one) andalso (c1 = zero)  | 
|
380  | 
then [linear_neg a]  | 
|
381  | 
else bset x p  | 
|
382  | 
|_ =>[])  | 
|
383  | 
||
384  | 
else bset x p  | 
|
385  | 
  |(Const ("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +",_) $ (Const ("op *",_) $c2 $ x) $ a)) =>  if (c1 =zero) andalso (c2 = one) then [linear_neg(linear_add [] a (mk_numeral 1))]  else [] 
 | 
|
386  | 
  |(Const ("op <",_) $ c1$ (Const ("op +",_) $(Const ("op *",_)$ c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = one) then [linear_neg a] else [] 
 | 
|
387  | 
  |(Const ("op &",_) $ p $ q) => (bset x p) union (bset x q) 
 | 
|
388  | 
  |(Const ("op |",_) $ p $ q) => (bset x p) union (bset x q) 
 | 
|
389  | 
|_ => [];  | 
|
390  | 
||
391  | 
(* ------------------------------------------------------------------------- *)  | 
|
392  | 
(* Construct the A-set. *)  | 
|
393  | 
(* ------------------------------------------------------------------------- *)  | 
|
394  | 
||
395  | 
fun aset x fm = case fm of  | 
|
396  | 
   (Const ("Not", _) $ p) => if (is_arith_rel p) then
 | 
|
397  | 
(case p of  | 
|
398  | 
	      (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $c2 $y) $a ) )
 | 
|
399  | 
=> if (x= y) andalso (c2 = one) andalso (c1 = zero)  | 
|
400  | 
then [linear_neg a]  | 
|
401  | 
else []  | 
|
402  | 
|_ =>[])  | 
|
403  | 
||
404  | 
else aset x p  | 
|
405  | 
  |(Const ("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +",_) $ (Const ("op *",_) $c2 $ x) $ a)) =>  if (c1 =zero) andalso (c2 = one) then [linear_sub [] (mk_numeral 1) a]  else []
 | 
|
406  | 
  |(Const ("op <",_) $ c1$ (Const ("op +",_) $(Const ("op *",_)$ c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = (mk_numeral (~1))) then [a] else []
 | 
|
407  | 
  |(Const ("op &",_) $ p $ q) => (aset x p) union (aset x q)
 | 
|
408  | 
  |(Const ("op |",_) $ p $ q) => (aset x p) union (aset x q)
 | 
|
409  | 
|_ => [];  | 
|
410  | 
||
411  | 
||
412  | 
(* ------------------------------------------------------------------------- *)  | 
|
413  | 
(* Replace top variable with another linear form, retaining canonicality. *)  | 
|
414  | 
(* ------------------------------------------------------------------------- *)  | 
|
415  | 
||
416  | 
fun linrep vars x t fm = case fm of  | 
|
417  | 
   ((Const(p,_)$ d $ (Const("op +",_)$(Const("op *",_)$ c $ y) $ z))) => 
 | 
|
418  | 
if (x = y) andalso (is_arith_rel fm)  | 
|
419  | 
then  | 
|
420  | 
let val ct = linear_cmul (dest_numeral c) t  | 
|
421  | 
in (HOLogic.mk_binrel p (d, linear_add vars ct z))  | 
|
422  | 
end  | 
|
423  | 
else fm  | 
|
424  | 
  |(Const ("Not", _) $ p) => HOLogic.Not $ (linrep vars x t p) 
 | 
|
425  | 
  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (linrep vars x t p) $ (linrep vars x t q) 
 | 
|
426  | 
  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (linrep vars x t p) $ (linrep vars x t q) 
 | 
|
| 15267 | 427  | 
|_ => fm;  | 
| 13876 | 428  | 
|
429  | 
(* ------------------------------------------------------------------------- *)  | 
|
430  | 
(* Evaluation of constant expressions. *)  | 
|
431  | 
(* ------------------------------------------------------------------------- *)  | 
|
| 15107 | 432  | 
|
433  | 
(* An other implementation of divides, that covers more cases*)  | 
|
434  | 
||
435  | 
exception DVD_UNKNOWN  | 
|
436  | 
||
437  | 
fun dvd_op (d, t) =  | 
|
438  | 
if not(is_numeral d) then raise DVD_UNKNOWN  | 
|
439  | 
else let  | 
|
440  | 
val dn = dest_numeral d  | 
|
441  | 
fun coeffs_of x = case x of  | 
|
442  | 
Const(p,_) $ tl $ tr =>  | 
|
443  | 
if p = "op +" then (coeffs_of tl) union (coeffs_of tr)  | 
|
444  | 
else if p = "op *"  | 
|
445  | 
then if (is_numeral tr)  | 
|
446  | 
then [(dest_numeral tr) * (dest_numeral tl)]  | 
|
447  | 
else [dest_numeral tl]  | 
|
448  | 
else []  | 
|
449  | 
|_ => if (is_numeral t) then [dest_numeral t] else []  | 
|
450  | 
val ts = coeffs_of t  | 
|
451  | 
in case ts of  | 
|
452  | 
[] => raise DVD_UNKNOWN  | 
|
| 
15574
 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 
skalberg 
parents: 
15570 
diff
changeset
 | 
453  | 
|_ => foldr (fn(k,r) => r andalso (k mod dn = 0)) true ts  | 
| 15107 | 454  | 
end;  | 
455  | 
||
456  | 
||
| 13876 | 457  | 
val operations =  | 
| 16398 | 458  | 
  [("op =",op=), ("op <",IntInf.<), ("op >",IntInf.>), ("op <=",IntInf.<=) , 
 | 
459  | 
   ("op >=",IntInf.>=), 
 | 
|
460  | 
   ("Divides.op dvd",fn (x,y) =>((IntInf.mod(y, x)) = 0))]; 
 | 
|
| 13876 | 461  | 
|
| 15531 | 462  | 
fun applyoperation (SOME f) (a,b) = f (a, b)  | 
| 13876 | 463  | 
|applyoperation _ (_, _) = false;  | 
464  | 
||
465  | 
(*Evaluation of constant atomic formulas*)  | 
|
| 15107 | 466  | 
(*FIXME : This is an optimation but still incorrect !! *)  | 
467  | 
(*  | 
|
| 13876 | 468  | 
fun evalc_atom at = case at of  | 
| 15107 | 469  | 
(Const (p,_) $ s $ t) =>  | 
470  | 
(if p="Divides.op dvd" then  | 
|
471  | 
((if dvd_op(s,t) then HOLogic.true_const  | 
|
472  | 
else HOLogic.false_const)  | 
|
473  | 
handle _ => at)  | 
|
474  | 
else  | 
|
475  | 
case assoc (operations,p) of  | 
|
| 15531 | 476  | 
SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)  | 
| 15107 | 477  | 
handle _ => at)  | 
478  | 
| _ => at)  | 
|
479  | 
      |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
 | 
|
480  | 
case assoc (operations,p) of  | 
|
| 15531 | 481  | 
SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then  | 
| 15107 | 482  | 
HOLogic.false_const else HOLogic.true_const)  | 
483  | 
handle _ => at)  | 
|
484  | 
| _ => at)  | 
|
485  | 
| _ => at;  | 
|
486  | 
||
487  | 
*)  | 
|
488  | 
||
489  | 
fun evalc_atom at = case at of  | 
|
490  | 
(Const (p,_) $ s $ t) =>  | 
|
491  | 
( case assoc (operations,p) of  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
492  | 
SOME f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
493  | 
else HOLogic.false_const)  | 
| 15107 | 494  | 
handle _ => at)  | 
495  | 
| _ => at)  | 
|
496  | 
      |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
 | 
|
497  | 
case assoc (operations,p) of  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
498  | 
SOME f => ((if (f ((dest_numeral s),(dest_numeral t)))  | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
499  | 
then HOLogic.false_const else HOLogic.true_const)  | 
| 15107 | 500  | 
handle _ => at)  | 
501  | 
| _ => at)  | 
|
502  | 
| _ => at;  | 
|
503  | 
||
504  | 
(*Function onatoms apllys function f on the atomic formulas involved in a.*)  | 
|
| 13876 | 505  | 
|
506  | 
fun onatoms f a = if (is_arith_rel a) then f a else case a of  | 
|
507  | 
||
508  | 
  	(Const ("Not",_) $ p) => if is_arith_rel p then HOLogic.Not $ (f p) 
 | 
|
509  | 
||
510  | 
else HOLogic.Not $ (onatoms f p)  | 
|
511  | 
  	|(Const ("op &",_) $ p $ q) => HOLogic.conj $ (onatoms f p) $ (onatoms f q) 
 | 
|
512  | 
  	|(Const ("op |",_) $ p $ q) => HOLogic.disj $ (onatoms f p) $ (onatoms f q) 
 | 
|
513  | 
  	|(Const ("op -->",_) $ p $ q) => HOLogic.imp $ (onatoms f p) $ (onatoms f q) 
 | 
|
514  | 
  	|((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q) => (Const ("op =", [HOLogic.boolT, HOLogic.boolT] ---> HOLogic.boolT)) $ (onatoms f p) $ (onatoms f q) 
 | 
|
515  | 
  	|(Const("All",_) $ Abs(x,T,p)) => Const("All", [HOLogic.intT --> 
 | 
|
516  | 
HOLogic.boolT] ---> HOLogic.boolT)$ Abs (x ,T, (onatoms f p))  | 
|
517  | 
  	|(Const("Ex",_) $ Abs(x,T,p)) => Const("Ex", [HOLogic.intT --> HOLogic.boolT]---> HOLogic.boolT) $ Abs( x ,T, (onatoms f p)) 
 | 
|
518  | 
|_ => a;  | 
|
519  | 
||
520  | 
val evalc = onatoms evalc_atom;  | 
|
521  | 
||
522  | 
(* ------------------------------------------------------------------------- *)  | 
|
523  | 
(* Hence overall quantifier elimination. *)  | 
|
524  | 
(* ------------------------------------------------------------------------- *)  | 
|
525  | 
||
526  | 
(*Applyes a function iteratively on the list*)  | 
|
527  | 
||
528  | 
fun end_itlist f [] = error "end_itlist"  | 
|
529  | 
|end_itlist f [x] = x  | 
|
530  | 
|end_itlist f (h::t) = f h (end_itlist f t);  | 
|
531  | 
||
532  | 
||
533  | 
(*list_disj[conj] makes a disj[conj] of a given list. used with conjucts or disjuncts  | 
|
534  | 
it liearises iterated conj[disj]unctions. *)  | 
|
535  | 
||
536  | 
fun disj_help p q = HOLogic.disj $ p $ q ;  | 
|
537  | 
||
538  | 
fun list_disj l =  | 
|
539  | 
if l = [] then HOLogic.false_const else end_itlist disj_help l;  | 
|
540  | 
||
541  | 
fun conj_help p q = HOLogic.conj $ p $ q ;  | 
|
542  | 
||
543  | 
fun list_conj l =  | 
|
544  | 
if l = [] then HOLogic.true_const else end_itlist conj_help l;  | 
|
545  | 
||
546  | 
(*Simplification of Formulas *)  | 
|
547  | 
||
548  | 
(*Function q_bnd_chk checks if a quantified Formula makes sens : Means if in  | 
|
549  | 
the body of the existential quantifier there are bound variables to the  | 
|
550  | 
existential quantifier.*)  | 
|
551  | 
||
552  | 
fun has_bound fm =let fun has_boundh fm i = case fm of  | 
|
553  | 
Bound n => (i = n)  | 
|
554  | 
|Abs (_,_,p) => has_boundh p (i+1)  | 
|
555  | 
|t1 $ t2 => (has_boundh t1 i) orelse (has_boundh t2 i)  | 
|
556  | 
|_ =>false  | 
|
557  | 
||
558  | 
in case fm of  | 
|
559  | 
Bound _ => true  | 
|
560  | 
|Abs (_,_,p) => has_boundh p 0  | 
|
561  | 
|t1 $ t2 => (has_bound t1 ) orelse (has_bound t2 )  | 
|
562  | 
|_ =>false  | 
|
563  | 
end;  | 
|
564  | 
||
565  | 
(*has_sub_abs checks if in a given Formula there are subformulas which are quantifed  | 
|
566  | 
too. Is no used no more.*)  | 
|
567  | 
||
568  | 
fun has_sub_abs fm = case fm of  | 
|
569  | 
Abs (_,_,_) => true  | 
|
570  | 
|t1 $ t2 => (has_bound t1 ) orelse (has_bound t2 )  | 
|
571  | 
|_ =>false ;  | 
|
572  | 
||
573  | 
(*update_bounds called with i=0 udates the numeration of bounded variables because the  | 
|
574  | 
formula will not be quantified any more.*)  | 
|
575  | 
||
576  | 
fun update_bounds fm i = case fm of  | 
|
577  | 
Bound n => if n >= i then Bound (n-1) else fm  | 
|
578  | 
|Abs (x,T,p) => Abs(x,T,(update_bounds p (i+1)))  | 
|
579  | 
|t1 $ t2 => (update_bounds t1 i) $ (update_bounds t2 i)  | 
|
580  | 
|_ => fm ;  | 
|
581  | 
||
582  | 
(*psimpl : Simplification of propositions (general purpose)*)  | 
|
583  | 
fun psimpl1 fm = case fm of  | 
|
584  | 
    Const("Not",_) $ Const ("False",_) => HOLogic.true_const 
 | 
|
585  | 
  | Const("Not",_) $ Const ("True",_) => HOLogic.false_const 
 | 
|
586  | 
  | Const("op &",_) $ Const ("False",_) $ q => HOLogic.false_const 
 | 
|
587  | 
  | Const("op &",_) $ p $ Const ("False",_)  => HOLogic.false_const 
 | 
|
588  | 
  | Const("op &",_) $ Const ("True",_) $ q => q 
 | 
|
589  | 
  | Const("op &",_) $ p $ Const ("True",_) => p 
 | 
|
590  | 
  | Const("op |",_) $ Const ("False",_) $ q => q 
 | 
|
591  | 
  | Const("op |",_) $ p $ Const ("False",_)  => p 
 | 
|
592  | 
  | Const("op |",_) $ Const ("True",_) $ q => HOLogic.true_const 
 | 
|
593  | 
  | Const("op |",_) $ p $ Const ("True",_)  => HOLogic.true_const 
 | 
|
594  | 
  | Const("op -->",_) $ Const ("False",_) $ q => HOLogic.true_const 
 | 
|
595  | 
  | Const("op -->",_) $ Const ("True",_) $  q => q 
 | 
|
596  | 
  | Const("op -->",_) $ p $ Const ("True",_)  => HOLogic.true_const 
 | 
|
597  | 
  | Const("op -->",_) $ p $ Const ("False",_)  => HOLogic.Not $  p 
 | 
|
598  | 
  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ Const ("True",_) $ q => q 
 | 
|
599  | 
  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ Const ("True",_) => p 
 | 
|
600  | 
  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ Const ("False",_) $ q => HOLogic.Not $  q 
 | 
|
601  | 
  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ Const ("False",_)  => HOLogic.Not $  p 
 | 
|
602  | 
| _ => fm;  | 
|
603  | 
||
604  | 
fun psimpl fm = case fm of  | 
|
605  | 
   Const ("Not",_) $ p => psimpl1 (HOLogic.Not $ (psimpl p)) 
 | 
|
606  | 
  | Const("op &",_) $ p $ q => psimpl1 (HOLogic.mk_conj (psimpl p,psimpl q)) 
 | 
|
607  | 
  | Const("op |",_) $ p $ q => psimpl1 (HOLogic.mk_disj (psimpl p,psimpl q)) 
 | 
|
608  | 
  | Const("op -->",_) $ p $ q => psimpl1 (HOLogic.mk_imp(psimpl p,psimpl q)) 
 | 
|
| 15267 | 609  | 
  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q => psimpl1 (HOLogic.mk_eq(psimpl p,psimpl q))
 | 
| 13876 | 610  | 
| _ => fm;  | 
611  | 
||
612  | 
||
613  | 
(*simpl : Simplification of Terms involving quantifiers too.  | 
|
614  | 
This function is able to drop out some quantified expressions where there are no  | 
|
615  | 
bound varaibles.*)  | 
|
616  | 
||
617  | 
fun simpl1 fm =  | 
|
618  | 
case fm of  | 
|
619  | 
    Const("All",_) $Abs(x,_,p) => if (has_bound fm ) then fm  
 | 
|
620  | 
else (update_bounds p 0)  | 
|
621  | 
  | Const("Ex",_) $ Abs (x,_,p) => if has_bound fm then fm  
 | 
|
622  | 
else (update_bounds p 0)  | 
|
| 15267 | 623  | 
| _ => psimpl fm;  | 
| 13876 | 624  | 
|
625  | 
fun simpl fm = case fm of  | 
|
626  | 
    Const ("Not",_) $ p => simpl1 (HOLogic.Not $(simpl p))  
 | 
|
627  | 
  | Const ("op &",_) $ p $ q => simpl1 (HOLogic.mk_conj (simpl p ,simpl q))  
 | 
|
628  | 
  | Const ("op |",_) $ p $ q => simpl1 (HOLogic.mk_disj (simpl p ,simpl q ))  
 | 
|
629  | 
  | Const ("op -->",_) $ p $ q => simpl1 (HOLogic.mk_imp(simpl p ,simpl q ))  
 | 
|
630  | 
  | Const("op =", Type ("fun",[Type ("bool", []),_]))$ p $ q => simpl1 
 | 
|
631  | 
(HOLogic.mk_eq(simpl p ,simpl q ))  | 
|
| 
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: 
14877 
diff
changeset
 | 
632  | 
(*  | Const ("All",Ta) $ Abs(Vn,VT,p) => simpl1(Const("All",Ta) $ 
 | 
| 13876 | 633  | 
Abs(Vn,VT,simpl p ))  | 
634  | 
  | Const ("Ex",Ta)  $ Abs(Vn,VT,p) => simpl1(Const("Ex",Ta)  $ 
 | 
|
635  | 
Abs(Vn,VT,simpl p ))  | 
|
| 
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: 
14877 
diff
changeset
 | 
636  | 
*)  | 
| 13876 | 637  | 
| _ => fm;  | 
638  | 
||
639  | 
(* ------------------------------------------------------------------------- *)  | 
|
640  | 
||
641  | 
(* Puts fm into NNF*)  | 
|
642  | 
||
643  | 
fun nnf fm = if (is_arith_rel fm) then fm  | 
|
644  | 
else (case fm of  | 
|
645  | 
  ( Const ("op &",_) $ p $ q)  => HOLogic.conj $ (nnf p) $(nnf q) 
 | 
|
646  | 
  | (Const("op |",_) $ p $q) => HOLogic.disj $ (nnf p)$(nnf q) 
 | 
|
647  | 
  | (Const ("op -->",_)  $ p $ q) => HOLogic.disj $ (nnf (HOLogic.Not $ p)) $ (nnf q) 
 | 
|
648  | 
  | ((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q) =>(HOLogic.disj $ (HOLogic.conj $ (nnf p) $ (nnf q)) $ (HOLogic.conj $ (nnf (HOLogic.Not $ p) ) $ (nnf(HOLogic.Not $ q)))) 
 | 
|
649  | 
  | (Const ("Not",_)) $ ((Const ("Not",_)) $ p) => (nnf p) 
 | 
|
650  | 
  | (Const ("Not",_)) $ (( Const ("op &",_)) $ p $ q) =>HOLogic.disj $ (nnf(HOLogic.Not $ p)) $ (nnf(HOLogic.Not $q)) 
 | 
|
651  | 
  | (Const ("Not",_)) $ (( Const ("op |",_)) $ p $ q) =>HOLogic.conj $ (nnf(HOLogic.Not $ p)) $ (nnf(HOLogic.Not $ q)) 
 | 
|
652  | 
  | (Const ("Not",_)) $ (( Const ("op -->",_)) $ p $ q ) =>HOLogic.conj $ (nnf p) $(nnf(HOLogic.Not $ q)) 
 | 
|
653  | 
  | (Const ("Not",_)) $ ((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q ) =>(HOLogic.disj $ (HOLogic.conj $(nnf p) $ (nnf(HOLogic.Not $ q))) $ (HOLogic.conj $(nnf(HOLogic.Not $ p)) $ (nnf q))) 
 | 
|
654  | 
| _ => fm);  | 
|
655  | 
||
656  | 
||
657  | 
(* Function remred to remove redundancy in a list while keeping the order of appearance of the  | 
|
658  | 
elements. but VERY INEFFICIENT!! *)  | 
|
659  | 
||
660  | 
fun remred1 el [] = []  | 
|
661  | 
|remred1 el (h::t) = if el=h then (remred1 el t) else h::(remred1 el t);  | 
|
662  | 
||
663  | 
fun remred [] = []  | 
|
664  | 
|remred (x::l) = x::(remred1 x (remred l));  | 
|
665  | 
||
666  | 
(*Makes sure that all free Variables are of the type integer but this function is only  | 
|
667  | 
used temporarily, this job must be done by the parser later on.*)  | 
|
668  | 
||
669  | 
fun mk_uni_vars T (node $ rest) = (case node of  | 
|
670  | 
Free (name,_) => Free (name,T) $ (mk_uni_vars T rest)  | 
|
671  | 
|_=> (mk_uni_vars T node) $ (mk_uni_vars T rest ) )  | 
|
672  | 
|mk_uni_vars T (Free (v,_)) = Free (v,T)  | 
|
673  | 
|mk_uni_vars T tm = tm;  | 
|
674  | 
||
675  | 
fun mk_uni_int T (Const ("0",T2)) = if T = T2 then (mk_numeral 0) else (Const ("0",T2)) 
 | 
|
676  | 
    |mk_uni_int T (Const ("1",T2)) = if T = T2 then (mk_numeral 1) else (Const ("1",T2)) 
 | 
|
677  | 
|mk_uni_int T (node $ rest) = (mk_uni_int T node) $ (mk_uni_int T rest )  | 
|
678  | 
|mk_uni_int T (Abs(AV,AT,p)) = Abs(AV,AT,mk_uni_int T p)  | 
|
679  | 
|mk_uni_int T tm = tm;  | 
|
680  | 
||
681  | 
||
| 16398 | 682  | 
(* Minusinfinity Version*)  | 
683  | 
fun myupto (m:IntInf.int) n = if m > n then [] else m::(myupto (m+1) n)  | 
|
684  | 
||
| 13876 | 685  | 
fun coopermi vars1 fm =  | 
686  | 
case fm of  | 
|
| 
15965
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
687  | 
   Const ("Ex",_) $ Abs(x0,T,p0) => 
 | 
| 
 
f422f8283491
Use of IntInf.int instead of int in most numeric simprocs; avoids
 
paulson 
parents: 
15859 
diff
changeset
 | 
688  | 
let  | 
| 13876 | 689  | 
val (xn,p1) = variant_abs (x0,T,p0)  | 
690  | 
val x = Free (xn,T)  | 
|
691  | 
val vars = (xn::vars1)  | 
|
692  | 
val p = unitycoeff x (posineq (simpl p1))  | 
|
693  | 
val p_inf = simpl (minusinf x p)  | 
|
694  | 
val bset = bset x p  | 
|
| 16398 | 695  | 
val js = myupto 1 (divlcm x p)  | 
| 13876 | 696  | 
fun p_element j b = linrep vars x (linear_add vars b (mk_numeral j)) p  | 
697  | 
fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) bset)  | 
|
698  | 
in (list_disj (map stage js))  | 
|
699  | 
end  | 
|
700  | 
| _ => error "cooper: not an existential formula";  | 
|
701  | 
||
702  | 
||
703  | 
||
704  | 
(* The plusinfinity version of cooper*)  | 
|
705  | 
fun cooperpi vars1 fm =  | 
|
706  | 
case fm of  | 
|
707  | 
   Const ("Ex",_) $ Abs(x0,T,p0) => let 
 | 
|
708  | 
val (xn,p1) = variant_abs (x0,T,p0)  | 
|
709  | 
val x = Free (xn,T)  | 
|
710  | 
val vars = (xn::vars1)  | 
|
711  | 
val p = unitycoeff x (posineq (simpl p1))  | 
|
712  | 
val p_inf = simpl (plusinf x p)  | 
|
713  | 
val aset = aset x p  | 
|
| 16398 | 714  | 
val js = myupto 1 (divlcm x p)  | 
| 13876 | 715  | 
fun p_element j a = linrep vars x (linear_sub vars a (mk_numeral j)) p  | 
716  | 
fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) aset)  | 
|
717  | 
in (list_disj (map stage js))  | 
|
718  | 
end  | 
|
719  | 
| _ => error "cooper: not an existential formula";  | 
|
720  | 
||
721  | 
||
| 15107 | 722  | 
(* Try to find a withness for the formula *)  | 
723  | 
||
724  | 
fun inf_w mi d vars x p =  | 
|
725  | 
let val f = if mi then minusinf else plusinf in  | 
|
726  | 
case (simpl (minusinf x p)) of  | 
|
| 15531 | 727  | 
   Const("True",_)  => (SOME (mk_numeral 1), HOLogic.true_const)
 | 
728  | 
  |Const("False",_) => (NONE,HOLogic.false_const)
 | 
|
| 15107 | 729  | 
|F =>  | 
730  | 
let  | 
|
731  | 
fun h n =  | 
|
732  | 
case ((simpl o evalc) (linrep vars x (mk_numeral n) F)) of  | 
|
| 15531 | 733  | 
	Const("True",_) => (SOME (mk_numeral n),HOLogic.true_const)
 | 
734  | 
|F' => if n=1 then (NONE,F')  | 
|
| 15107 | 735  | 
else let val (rw,rf) = h (n-1) in  | 
736  | 
(rw,HOLogic.mk_disj(F',rf))  | 
|
737  | 
end  | 
|
738  | 
||
739  | 
in (h d)  | 
|
740  | 
end  | 
|
741  | 
end;  | 
|
742  | 
||
743  | 
fun set_w d b st vars x p = let  | 
|
744  | 
fun h ns = case ns of  | 
|
| 15531 | 745  | 
[] => (NONE,HOLogic.false_const)  | 
| 15107 | 746  | 
|n::nl => ( case ((simpl o evalc) (linrep vars x n p)) of  | 
| 15531 | 747  | 
      Const("True",_) => (SOME n,HOLogic.true_const)
 | 
| 15107 | 748  | 
|F' => let val (rw,rf) = h nl  | 
749  | 
in (rw,HOLogic.mk_disj(F',rf))  | 
|
750  | 
end)  | 
|
751  | 
val f = if b then linear_add else linear_sub  | 
|
| 16398 | 752  | 
val p_elements = foldr (fn (i,l) => l union (map (fn e => f [] e (mk_numeral i)) st)) [] (myupto 1 d)  | 
| 15107 | 753  | 
in h p_elements  | 
754  | 
end;  | 
|
755  | 
||
756  | 
fun withness d b st vars x p = case (inf_w b d vars x p) of  | 
|
| 15531 | 757  | 
(SOME n,_) => (SOME n,HOLogic.true_const)  | 
758  | 
|(NONE,Pinf) => (case (set_w d b st vars x p) of  | 
|
759  | 
(SOME n,_) => (SOME n,HOLogic.true_const)  | 
|
760  | 
|(_,Pst) => (NONE,HOLogic.mk_disj(Pinf,Pst)));  | 
|
| 15107 | 761  | 
|
762  | 
||
763  | 
||
| 13876 | 764  | 
|
765  | 
(*Cooper main procedure*)  | 
|
| 15267 | 766  | 
|
767  | 
exception STAGE_TRUE;  | 
|
768  | 
||
| 13876 | 769  | 
|
770  | 
fun cooper vars1 fm =  | 
|
771  | 
case fm of  | 
|
772  | 
   Const ("Ex",_) $ Abs(x0,T,p0) => let 
 | 
|
773  | 
val (xn,p1) = variant_abs (x0,T,p0)  | 
|
774  | 
val x = Free (xn,T)  | 
|
775  | 
val vars = (xn::vars1)  | 
|
| 
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: 
14877 
diff
changeset
 | 
776  | 
(* val p = unitycoeff x (posineq (simpl p1)) *)  | 
| 
 
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: 
14877 
diff
changeset
 | 
777  | 
val p = unitycoeff x p1  | 
| 13876 | 778  | 
val ast = aset x p  | 
779  | 
val bst = bset x p  | 
|
| 16398 | 780  | 
val js = myupto 1 (divlcm x p)  | 
| 13876 | 781  | 
val (p_inf,f,S ) =  | 
| 15267 | 782  | 
if (length bst) <= (length ast)  | 
783  | 
then (simpl (minusinf x p),linear_add,bst)  | 
|
784  | 
else (simpl (plusinf x p), linear_sub,ast)  | 
|
| 13876 | 785  | 
fun p_element j a = linrep vars x (f vars a (mk_numeral j)) p  | 
786  | 
fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) S)  | 
|
| 15267 | 787  | 
fun stageh n = ((if n = 0 then []  | 
788  | 
else  | 
|
789  | 
let  | 
|
790  | 
val nth_stage = simpl (evalc (stage n))  | 
|
791  | 
in  | 
|
792  | 
if (nth_stage = HOLogic.true_const)  | 
|
793  | 
then raise STAGE_TRUE  | 
|
794  | 
else if (nth_stage = HOLogic.false_const) then stageh (n-1)  | 
|
795  | 
else nth_stage::(stageh (n-1))  | 
|
796  | 
end )  | 
|
797  | 
handle STAGE_TRUE => [HOLogic.true_const])  | 
|
798  | 
val slist = stageh (divlcm x p)  | 
|
799  | 
in (list_disj slist)  | 
|
| 13876 | 800  | 
end  | 
801  | 
| _ => error "cooper: not an existential formula";  | 
|
802  | 
||
803  | 
||
| 15107 | 804  | 
(* A Version of cooper that returns a withness *)  | 
805  | 
fun cooper_w vars1 fm =  | 
|
806  | 
case fm of  | 
|
807  | 
   Const ("Ex",_) $ Abs(x0,T,p0) => let 
 | 
|
808  | 
val (xn,p1) = variant_abs (x0,T,p0)  | 
|
809  | 
val x = Free (xn,T)  | 
|
810  | 
val vars = (xn::vars1)  | 
|
811  | 
(* val p = unitycoeff x (posineq (simpl p1)) *)  | 
|
812  | 
val p = unitycoeff x p1  | 
|
813  | 
val ast = aset x p  | 
|
814  | 
val bst = bset x p  | 
|
815  | 
val d = divlcm x p  | 
|
816  | 
val (p_inf,S ) =  | 
|
817  | 
if (length bst) <= (length ast)  | 
|
818  | 
then (true,bst)  | 
|
819  | 
else (false,ast)  | 
|
820  | 
in withness d p_inf S vars x p  | 
|
821  | 
(* fun p_element j a = linrep vars x (f vars a (mk_numeral j)) p  | 
|
822  | 
fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) S)  | 
|
823  | 
in (list_disj (map stage js))  | 
|
824  | 
*)  | 
|
825  | 
end  | 
|
826  | 
| _ => error "cooper: not an existential formula";  | 
|
| 13876 | 827  | 
|
828  | 
||
829  | 
(*Function itlist applys a double parametred function f : 'a->'b->b iteratively to a List l : 'a  | 
|
830  | 
list With End condition b. ict calculates f(e1,f(f(e2,f(e3,...(...f(en,b))..)))))  | 
|
831  | 
assuming l = [e1,e2,...,en]*)  | 
|
832  | 
||
833  | 
fun itlist f l b = case l of  | 
|
834  | 
[] => b  | 
|
835  | 
| (h::t) => f h (itlist f t b);  | 
|
836  | 
||
837  | 
(* ------------------------------------------------------------------------- *)  | 
|
838  | 
(* Free variables in terms and formulas. *)  | 
|
839  | 
(* ------------------------------------------------------------------------- *)  | 
|
840  | 
||
841  | 
fun fvt tml = case tml of  | 
|
842  | 
[] => []  | 
|
843  | 
| Free(x,_)::r => x::(fvt r)  | 
|
844  | 
||
845  | 
fun fv fm = fvt (term_frees fm);  | 
|
846  | 
||
847  | 
||
848  | 
(* ========================================================================= *)  | 
|
849  | 
(* Quantifier elimination. *)  | 
|
850  | 
(* ========================================================================= *)  | 
|
851  | 
(*conj[/disj]uncts lists iterated conj[disj]unctions*)  | 
|
852  | 
||
853  | 
fun disjuncts fm = case fm of  | 
|
854  | 
    Const ("op |",_) $ p $ q => (disjuncts p) @ (disjuncts q) 
 | 
|
855  | 
| _ => [fm];  | 
|
856  | 
||
857  | 
fun conjuncts fm = case fm of  | 
|
858  | 
    Const ("op &",_) $p $ q => (conjuncts p) @ (conjuncts q) 
 | 
|
859  | 
| _ => [fm];  | 
|
860  | 
||
861  | 
||
862  | 
||
863  | 
(* ------------------------------------------------------------------------- *)  | 
|
864  | 
(* Lift procedure given literal modifier, formula normalizer & basic quelim. *)  | 
|
| 
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: 
14877 
diff
changeset
 | 
865  | 
(* ------------------------------------------------------------------------- *)  | 
| 15267 | 866  | 
|
| 
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: 
14877 
diff
changeset
 | 
867  | 
fun lift_qelim afn nfn qfn isat =  | 
| 
 
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: 
14877 
diff
changeset
 | 
868  | 
let  | 
| 
 
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: 
14877 
diff
changeset
 | 
869  | 
fun qelift vars fm = if (isat fm) then afn vars fm  | 
| 
 
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: 
14877 
diff
changeset
 | 
870  | 
else  | 
| 
 
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: 
14877 
diff
changeset
 | 
871  | 
case fm of  | 
| 
 
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: 
14877 
diff
changeset
 | 
872  | 
  Const ("Not",_) $ p => HOLogic.Not $ (qelift vars p) 
 | 
| 
 
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: 
14877 
diff
changeset
 | 
873  | 
  | Const ("op &",_) $ p $q => HOLogic.conj $ (qelift vars p) $ (qelift vars q) 
 | 
| 
 
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: 
14877 
diff
changeset
 | 
874  | 
  | Const ("op |",_) $ p $ q => HOLogic.disj $ (qelift vars p) $ (qelift vars q) 
 | 
| 
 
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: 
14877 
diff
changeset
 | 
875  | 
  | Const ("op -->",_) $ p $ q => HOLogic.imp $ (qelift vars p) $ (qelift vars q) 
 | 
| 
 
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: 
14877 
diff
changeset
 | 
876  | 
  | Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q => HOLogic.mk_eq ((qelift vars p),(qelift vars q)) 
 | 
| 
 
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: 
14877 
diff
changeset
 | 
877  | 
  | Const ("All",QT) $ Abs(x,T,p) => HOLogic.Not $(qelift vars (Const ("Ex",QT) $ Abs(x,T,(HOLogic.Not $ p)))) 
 | 
| 
 
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: 
14877 
diff
changeset
 | 
878  | 
  | (e as Const ("Ex",_)) $ Abs (x,T,p)  =>  qfn vars (e$Abs (x,T,(nfn(qelift (x::vars) p))))
 | 
| 
 
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: 
14877 
diff
changeset
 | 
879  | 
| _ => fm  | 
| 
 
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: 
14877 
diff
changeset
 | 
880  | 
|
| 
 
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: 
14877 
diff
changeset
 | 
881  | 
in (fn fm => qelift (fv fm) fm)  | 
| 
 
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: 
14877 
diff
changeset
 | 
882  | 
end;  | 
| 15267 | 883  | 
|
| 
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: 
14877 
diff
changeset
 | 
884  | 
|
| 15267 | 885  | 
(*  | 
| 13876 | 886  | 
fun lift_qelim afn nfn qfn isat =  | 
887  | 
let fun qelim x vars p =  | 
|
888  | 
let val cjs = conjuncts p  | 
|
| 15570 | 889  | 
val (ycjs,ncjs) = List.partition (has_bound) cjs in  | 
| 13876 | 890  | 
(if ycjs = [] then p else  | 
891  | 
let val q = (qfn vars ((HOLogic.exists_const HOLogic.intT  | 
|
892  | 
) $ Abs(x,HOLogic.intT,(list_conj ycjs)))) in  | 
|
893  | 
(itlist conj_help ncjs q)  | 
|
894  | 
end)  | 
|
895  | 
end  | 
|
896  | 
||
897  | 
fun qelift vars fm = if (isat fm) then afn vars fm  | 
|
898  | 
else  | 
|
899  | 
case fm of  | 
|
900  | 
      Const ("Not",_) $ p => HOLogic.Not $ (qelift vars p) 
 | 
|
901  | 
    | Const ("op &",_) $ p $q => HOLogic.conj $ (qelift vars p) $ (qelift vars q) 
 | 
|
902  | 
    | Const ("op |",_) $ p $ q => HOLogic.disj $ (qelift vars p) $ (qelift vars q) 
 | 
|
903  | 
    | Const ("op -->",_) $ p $ q => HOLogic.imp $ (qelift vars p) $ (qelift vars q) 
 | 
|
904  | 
    | Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q => HOLogic.mk_eq ((qelift vars p),(qelift vars q)) 
 | 
|
905  | 
    | Const ("All",QT) $ Abs(x,T,p) => HOLogic.Not $(qelift vars (Const ("Ex",QT) $ Abs(x,T,(HOLogic.Not $ p)))) 
 | 
|
906  | 
    | Const ("Ex",_) $ Abs (x,T,p)  => let  val djs = disjuncts(nfn(qelift (x::vars) p)) in 
 | 
|
907  | 
list_disj(map (qelim x vars) djs) end  | 
|
908  | 
| _ => fm  | 
|
909  | 
||
910  | 
in (fn fm => simpl(qelift (fv fm) fm))  | 
|
911  | 
end;  | 
|
| 15267 | 912  | 
*)  | 
| 13876 | 913  | 
|
914  | 
(* ------------------------------------------------------------------------- *)  | 
|
915  | 
(* Cleverer (proposisional) NNF with conditional and literal modification. *)  | 
|
916  | 
(* ------------------------------------------------------------------------- *)  | 
|
917  | 
||
918  | 
(*Function Negate used by cnnf, negates a formula p*)  | 
|
919  | 
||
920  | 
fun negate (Const ("Not",_) $ p) = p 
 | 
|
921  | 
|negate p = (HOLogic.Not $ p);  | 
|
922  | 
||
923  | 
fun cnnf lfn =  | 
|
924  | 
let fun cnnfh fm = case fm of  | 
|
925  | 
      (Const ("op &",_) $ p $ q) => HOLogic.mk_conj(cnnfh p,cnnfh q) 
 | 
|
926  | 
    | (Const ("op |",_) $ p $ q) => HOLogic.mk_disj(cnnfh p,cnnfh q) 
 | 
|
927  | 
    | (Const ("op -->",_) $ p $q) => HOLogic.mk_disj(cnnfh(HOLogic.Not $ p),cnnfh q) 
 | 
|
928  | 
    | (Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q) => HOLogic.mk_disj( 
 | 
|
929  | 
HOLogic.mk_conj(cnnfh p,cnnfh q),  | 
|
930  | 
HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $q)))  | 
|
931  | 
||
932  | 
    | (Const ("Not",_) $ (Const("Not",_) $ p)) => cnnfh p 
 | 
|
933  | 
    | (Const ("Not",_) $ (Const ("op &",_) $ p $ q)) => HOLogic.mk_disj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $ q)) 
 | 
|
934  | 
    | (Const ("Not",_) $(Const ("op |",_) $ (Const ("op &",_) $ p $ q) $  
 | 
|
935  | 
    			(Const ("op &",_) $ p1 $ r))) => if p1 = negate p then 
 | 
|
936  | 
HOLogic.mk_disj(  | 
|
937  | 
cnnfh (HOLogic.mk_conj(p,cnnfh(HOLogic.Not $ q))),  | 
|
938  | 
cnnfh (HOLogic.mk_conj(p1,cnnfh(HOLogic.Not $ r))))  | 
|
939  | 
else HOLogic.mk_conj(  | 
|
940  | 
cnnfh (HOLogic.mk_disj(cnnfh (HOLogic.Not $ p),cnnfh(HOLogic.Not $ q))),  | 
|
941  | 
cnnfh (HOLogic.mk_disj(cnnfh (HOLogic.Not $ p1),cnnfh(HOLogic.Not $ r)))  | 
|
942  | 
)  | 
|
943  | 
    | (Const ("Not",_) $ (Const ("op |",_) $ p $ q)) => HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $ q)) 
 | 
|
944  | 
    | (Const ("Not",_) $ (Const ("op -->",_) $ p $q)) => HOLogic.mk_conj(cnnfh p,cnnfh(HOLogic.Not $ q)) 
 | 
|
945  | 
    | (Const ("Not",_) $ (Const ("op =",Type ("fun",[Type ("bool", []),_]))  $ p $ q)) => HOLogic.mk_disj(HOLogic.mk_conj(cnnfh p,cnnfh(HOLogic.Not $ q)),HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh q)) 
 | 
|
946  | 
| _ => lfn fm  | 
|
| 
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: 
14877 
diff
changeset
 | 
947  | 
in cnnfh  | 
| 
 
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: 
14877 
diff
changeset
 | 
948  | 
end;  | 
| 13876 | 949  | 
|
950  | 
(*End- function the quantifierelimination an decion procedure of presburger formulas.*)  | 
|
| 
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: 
14877 
diff
changeset
 | 
951  | 
|
| 
 
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: 
14877 
diff
changeset
 | 
952  | 
(*  | 
| 13876 | 953  | 
val integer_qelim = simpl o evalc o (lift_qelim linform (simpl o (cnnf posineq o evalc)) cooper is_arith_rel) ;  | 
| 
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: 
14877 
diff
changeset
 | 
954  | 
*)  | 
| 15267 | 955  | 
|
956  | 
||
| 
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: 
14877 
diff
changeset
 | 
957  | 
val integer_qelim = simpl o evalc o (lift_qelim linform (cnnf posineq o evalc) cooper is_arith_rel) ;  | 
| 
 
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: 
14877 
diff
changeset
 | 
958  | 
|
| 14941 | 959  | 
fun mk_presburger_oracle (sg,COOPER_ORACLE t) =  | 
960  | 
if (!quick_and_dirty)  | 
|
961  | 
then HOLogic.mk_Trueprop (HOLogic.mk_eq(t,integer_qelim t))  | 
|
| 15107 | 962  | 
else raise COOPER_ORACLE t  | 
963  | 
|mk_presburger_oracle (sg,_) = error "Oops";  | 
|
| 13876 | 964  | 
end;  |