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