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