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