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