# HG changeset patch # User chaieb # Date 1126711552 -7200 # Node ID 105519771c67db045124ef02eedeb6cb0467586f # Parent afaa031ed4dad2592873d4fbbbd9cec55d33b2cc The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy comm_ring : a reflected Method for proving equalities in a commutative ring diff -r afaa031ed4da -r 105519771c67 src/HOL/Commutative_Ring.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Commutative_Ring.thy Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,310 @@ + +(* ID: $Id$ + Author: Bernhard Haeupler + + Proving equalities in a commutative Ring done "right" in Isabelle/HOL +*) + +theory Commutative_Ring +imports List +uses ("comm_ring.ML") +begin + + (* Syntax of multivariate polynomials (pol) and polynomial expressions*) +datatype 'a pol = + Pc 'a + | Pinj nat "'a pol" + | PX "'a pol" nat "'a pol" + +datatype 'a polex = + Pol "'a pol" + | Add "'a polex" "'a polex" + | Sub "'a polex" "'a polex" + | Mul "'a polex" "'a polex" + | Pow "'a polex" nat + | Neg "'a polex" + + (* Interpretation functions for the shadow syntax *) +consts + Ipol :: "('a::{comm_ring,recpower}) list \ 'a pol \ 'a" + Ipolex :: "('a::{comm_ring,recpower}) list \ 'a polex \ 'a" + +primrec + "Ipol l (Pc c) = c" + "Ipol l (Pinj i P) = Ipol (drop i l) P" + "Ipol l (PX P x Q) = (Ipol l P)*((hd l)^x) + (Ipol (drop 1 l) Q)" + +primrec + "Ipolex l (Pol P) = Ipol l P" + "Ipolex l (Add P Q) = (Ipolex l P) + (Ipolex l Q)" + "Ipolex l (Sub P Q) = (Ipolex l P) - (Ipolex l Q)" + "Ipolex l (Mul P Q) = (Ipolex l P) * (Ipolex l Q)" + "Ipolex l (Pow p n) = (Ipolex l p) ^ n" + "Ipolex l (Neg P) = -(Ipolex l P)" + + (* Create polynomial normalized polynomials given normalized inputs *) +constdefs mkPinj :: "nat \ 'a pol \ 'a pol" + mkPinj_def: "mkPinj x P \ (case P of + (Pc c) \ (Pc c) | + (Pinj y P) \ Pinj (x+y) P | + (PX p1 y p2) \ Pinj x P )" + +constdefs mkPX :: "('a::{comm_ring,recpower}) pol \ nat \ 'a pol \ 'a pol" + mkPX_def: "mkPX P i Q == (case P of + (Pc c) \ (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) | + (Pinj j R) \ (PX P i Q) | + (PX P2 i2 Q2) \ (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )" + + (* Defining the basic ring operations on normalized polynomials *) +consts +add :: "(('a::{comm_ring,recpower}) pol) \ ('a pol) \ 'a pol" +mul :: "(('a::{comm_ring,recpower}) pol) \ ('a pol) \ 'a pol" +neg :: "('a::{comm_ring,recpower}) pol \ 'a pol" +sqr :: "('a::{comm_ring,recpower}) pol \ 'a pol" +pow :: "('a::{comm_ring,recpower}) pol \ nat \ 'a pol" + + + (* Addition *) +recdef add "measure (\(x, y). size x + size y)" + "add (Pc a, Pc b) = Pc (a+b)" + "add (Pc c, Pinj i P) = Pinj i (add (P, Pc c))" + "add (Pinj i P, Pc c) = Pinj i (add (P, Pc c))" + "add (Pc c, PX P i Q) = PX P i (add (Q, Pc c))" + "add (PX P i Q, Pc c) = PX P i (add (Q, Pc c))" + "add (Pinj x P, Pinj y Q) = + (if x=y then mkPinj x (add (P, Q)) + else (if x>y then mkPinj y (add (Pinj (x-y) P, Q)) + else mkPinj x (add (Pinj (y-x) Q, P)) ))" + "add (Pinj x P, PX Q y R) = + (if x=0 then add(P, PX Q y R) + else (if x=1 then PX Q y (add (R, P)) + else PX Q y (add (R, Pinj (x - 1) P))))" + "add (PX P x R, Pinj y Q) = + (if y=0 then add(PX P x R, Q) + else (if y=1 then PX P x (add (R, Q)) + else PX P x (add (R, Pinj (y - 1) Q))))" + "add (PX P1 x P2, PX Q1 y Q2) = + (if x=y then mkPX (add (P1, Q1)) x (add (P2, Q2)) + else (if x>y then mkPX (add (PX P1 (x-y) (Pc 0), Q1)) y (add (P2,Q2)) + else mkPX (add (PX Q1 (y-x) (Pc 0), P1)) x (add (P2,Q2)) ))" + + (* Multiplication *) +recdef mul "measure (\(x, y). size x + size y)" + "mul (Pc a, Pc b) = Pc (a*b)" + "mul (Pc c, Pinj i P) = (if c=0 then Pc 0 else mkPinj i (mul (P, Pc c)))" + "mul (Pinj i P, Pc c) = (if c=0 then Pc 0 else mkPinj i (mul (P, Pc c)))" + "mul (Pc c, PX P i Q) = + (if c=0 then Pc 0 else mkPX (mul (P, Pc c)) i (mul (Q, Pc c)))" + "mul (PX P i Q, Pc c) = + (if c=0 then Pc 0 else mkPX (mul (P, Pc c)) i (mul (Q, Pc c)))" + "mul (Pinj x P, Pinj y Q) = + (if x=y then mkPinj x (mul (P, Q)) + else (if x>y then mkPinj y (mul (Pinj (x-y) P, Q)) + else mkPinj x (mul (Pinj (y-x) Q, P)) ))" + "mul (Pinj x P, PX Q y R) = + (if x=0 then mul(P, PX Q y R) + else (if x=1 then mkPX (mul (Pinj x P, Q)) y (mul (R, P)) + else mkPX (mul (Pinj x P, Q)) y (mul (R, Pinj (x - 1) P))))" + "mul (PX P x R, Pinj y Q) = + (if y=0 then mul(PX P x R, Q) + else (if y=1 then mkPX (mul (Pinj y Q, P)) x (mul (R, Q)) + else mkPX (mul (Pinj y Q, P)) x (mul (R, Pinj (y - 1) Q))))" + "mul (PX P1 x P2, PX Q1 y Q2) = + add (mkPX (mul (P1, Q1)) (x+y) (mul (P2, Q2)), + add (mkPX (mul (P1, mkPinj 1 Q2)) x (Pc 0), mkPX (mul (Q1, mkPinj 1 P2)) y (Pc 0)) )" +(hints simp add: mkPinj_def split: pol.split) + + (* Negation*) +primrec + "neg (Pc c) = (Pc (-c))" + "neg (Pinj i P) = Pinj i (neg P)" + "neg (PX P x Q) = PX (neg P) x (neg Q)" + + (* Substraction*) +constdefs sub :: "(('a::{comm_ring,recpower}) pol) \ ('a pol) \ 'a pol" + "sub p q \ add (p,neg q)" + + (* Square for Fast Exponentation *) +primrec + "sqr (Pc c) = Pc (c*c)" + "sqr (Pinj i P) = mkPinj i (sqr P)" + "sqr (PX A x B) = add (mkPX (sqr A) (x+x) (sqr B), mkPX (mul (mul (Pc (1+1), A), mkPinj 1 B)) x (Pc 0))" + + (* Fast Exponentation *) +lemma pow_wf:"odd n \ (n::nat) div 2 < n" by(cases n, auto) +recdef pow "measure (\(x, y). y)" + "pow (p, 0) = Pc 1" + "pow (p, n) = (if even n then (pow (sqr p, n div 2)) else mul (p, pow (sqr p, n div 2)))" +(hints simp add: pow_wf) + +lemma pow_if: "pow (p,n) = (if n=0 then Pc 1 else (if even n then (pow (sqr p, n div 2)) else mul (p, pow (sqr p, n div 2))))" +by (cases n) simp_all + +(* +lemma number_of_nat_B0: "(number_of (w BIT bit.B0) ::nat) = 2* (number_of w)" +by simp + +lemma number_of_nat_even: "even (number_of (w BIT bit.B0)::nat)" +by simp + +lemma pow_even : "pow (p, number_of(w BIT bit.B0)) = pow (sqr p, number_of w)" + ( is "pow(?p,?n) = pow (_,?n2)") +proof- + have "even ?n" by simp + hence "pow (p, ?n) = pow (sqr p, ?n div 2)" + apply simp + apply (cases "IntDef.neg (number_of w)") + apply simp + done +*) + (* Normalization of polynomial expressions *) + +consts norm :: "('a::{comm_ring,recpower}) polex \ 'a pol" +primrec + "norm (Pol P) = P" + "norm (Add P Q) = add (norm P, norm Q)" + "norm (Sub p q) = sub (norm p) (norm q)" + "norm (Mul P Q) = mul (norm P, norm Q)" + "norm (Pow p n) = pow (norm p, n)" + "norm (Neg P) = neg (norm P)" + + (* mkPinj preserve semantics *) +lemma mkPinj_ci: "ALL a l. Ipol l (mkPinj a B) = Ipol l (Pinj a B)" + by(induct B, auto simp add: mkPinj_def ring_eq_simps) + + (* mkPX preserves semantics *) +lemma mkPX_ci: "ALL b l. Ipol l (mkPX A b C) = Ipol l (PX A b C)" + by(case_tac A, auto simp add: mkPX_def mkPinj_ci power_add ring_eq_simps) + + (* Correctness theorems for the implemented operations *) + (* Negation *) +lemma neg_ci: "ALL l. Ipol l (neg P) = -(Ipol l P)" + by(induct P, auto) + + (* Addition *) +lemma add_ci: "ALL l. Ipol l (add (P, Q)) = Ipol l P + Ipol l Q" +proof(induct P Q rule: add.induct) + case (6 x P y Q) + have "x < y \ x = y \ x > y" by arith + moreover + { assume "xy" hence "EX d. d+y=x" by arith + then obtain d where "d+y=x".. + with prems have ?case by (auto simp add: mkPinj_ci ring_eq_simps) } + ultimately show ?case by blast +next + case (7 x P Q y R) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by auto } + moreover + { assume "x=1" with prems have ?case by (auto simp add: ring_eq_simps) } + moreover + { assume "x > 1" from prems have ?case by(cases x, auto) } + ultimately show ?case by blast +next + case (8 P x R y Q) + have "(y = 0) \ (y = 1) \ (y > 1)" by arith + moreover + {assume "y=0" with prems have ?case by simp} + moreover + {assume "y=1" with prems have ?case by simp} + moreover + {assume "y>1" hence "EX d. d+1=y" by arith + then obtain d where "d+1=y" .. + with prems have ?case by auto } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + have "y < x \ x = y \ x < y" by arith + moreover + {assume "y < x" hence "EX d. d+y=x" by arith + then obtain d where "d+y=x".. + with prems have ?case by (auto simp add: power_add mkPX_ci ring_eq_simps) } + moreover + {assume "x=y" with prems have ?case by(auto simp add: mkPX_ci ring_eq_simps) } + moreover + {assume "x l. Ipol l (sub p q) = (Ipol l p) - (Ipol l q)" + by (auto simp add: add_ci neg_ci sub_def) + + (* Square *) +lemma sqr_ci:"ALL ls. Ipol ls (sqr p) = Ipol ls p * Ipol ls p" +by(induct p, auto simp add: add_ci mkPinj_ci mkPX_ci mul_ci ring_eq_simps power_add) + + + (* Power *) +lemma even_pow:"even n \ pow (p, n) = pow (sqr p, n div 2)" by(induct n,auto) +lemma pow_ci: "ALL p. Ipol ls (pow (p, n)) = (Ipol ls p) ^ n" +proof(induct n rule: nat_less_induct) + case (1 k) + have two:"2=Suc( Suc 0)" by simp + from prems show ?case + proof(cases k) + case (Suc l) + hence KL:"k=Suc l" by simp + have "even l \ odd l" by (simp) + moreover + {assume EL:"even l" + have "Suc l div 2 = l div 2" by (simp add: nat_number even_nat_plus_one_div_two[OF EL]) + moreover + from KL have"l\mp. Ipol ls (pow (p, m)) = Ipol ls p ^ m; k = Suc l; odd l\ \ \p. Ipol ls (sqr p) ^ (Suc l div 2) = Ipol ls p ^ Suc l" + proof(cases l) + case (Suc w) + from prems have EW:"even w" by simp + from two have two_times:"(2 * (w div 2))= w" by (simp only: even_nat_div_two_times_two[OF EW]) + have A:"ALL p. (Ipol ls p * Ipol ls p) = (Ipol ls p) ^ (Suc (Suc 0))" by (simp add: power_Suc) + from A two[symmetric] have "ALL p.(Ipol ls p * Ipol ls p) = (Ipol ls p) ^ 2" by simp + with prems show ?thesis by (auto simp add: power_mult[symmetric, of _ 2 _] two_times mul_ci sqr_ci) + qed(simp) + with prems have ?thesis by simp } + ultimately show ?thesis by blast + qed(simp) +qed + + (* Normalization preserves semantics *) +lemma norm_ci:"Ipolex l Pe = Ipol l (norm Pe)" + by(induct Pe, simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci) + +(* Reflection lemma: Key to the (incomplete) decision procedure *) +lemma norm_eq: + assumes A:"norm P1 = norm P2" + shows "Ipolex l P1 = Ipolex l P2" +proof - + from A have "Ipol l (norm P1) = Ipol l (norm P2)" by simp + thus ?thesis by(simp only: norm_ci) +qed + + + (* Code generation *) +(* +Does not work, since no generic ring operations implementation is there +generate_code ("ring.ML") test = "norm"*) + + + use "comm_ring.ML" +setup "CommRing.setup" + +end \ No newline at end of file diff -r afaa031ed4da -r 105519771c67 src/HOL/Integ/Presburger.thy --- a/src/HOL/Integ/Presburger.thy Wed Sep 14 10:24:39 2005 +0200 +++ b/src/HOL/Integ/Presburger.thy Wed Sep 14 17:25:52 2005 +0200 @@ -10,7 +10,8 @@ theory Presburger imports NatSimprocs SetInterval -uses ("cooper_dec.ML") ("cooper_proof.ML") ("qelim.ML") ("presburger.ML") +uses ("cooper_dec.ML") ("cooper_proof.ML") ("qelim.ML") + ("reflected_presburger.ML") ("reflected_cooper.ML") ("presburger.ML") begin text {* Theorem for unitifying the coeffitients of @{text x} in an existential formula*} @@ -982,8 +983,10 @@ by (simp cong: conj_cong) use "cooper_dec.ML" +use "reflected_presburger.ML" +use "reflected_cooper.ML" oracle - presburger_oracle ("term") = CooperDec.presburger_oracle + presburger_oracle ("term") = ReflectedCooper.presburger_oracle use "cooper_proof.ML" use "qelim.ML" diff -r afaa031ed4da -r 105519771c67 src/HOL/Integ/cooper_dec.ML --- a/src/HOL/Integ/cooper_dec.ML Wed Sep 14 10:24:39 2005 +0200 +++ b/src/HOL/Integ/cooper_dec.ML Wed Sep 14 17:25:52 2005 +0200 @@ -44,7 +44,6 @@ val evalc : term -> term val cooper_w : string list -> term -> (term option * term) val integer_qelim : Term.term -> Term.term - val presburger_oracle : theory -> term -> term end; structure CooperDec : COOPER_DEC = @@ -938,8 +937,4 @@ val integer_qelim = simpl o evalc o (lift_qelim linform (cnnf posineq o evalc) cooper is_arith_rel) ; -fun presburger_oracle thy t = - if (!quick_and_dirty) - then HOLogic.mk_Trueprop (HOLogic.mk_eq(t,integer_qelim t)) - else error "Presburger oracle: not in quick_and_dirty mode" end; diff -r afaa031ed4da -r 105519771c67 src/HOL/Integ/reflected_cooper.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Integ/reflected_cooper.ML Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,124 @@ +(* $Id$ *) +(* The oracle for Presburger arithmetic based on the verified Code *) + (* in HOL/ex/Reflected_Presburger.thy *) + +structure ReflectedCooper = +struct + +open Generated; +(* pseudo reification : term -> intterm *) + +fun i_of_term vs t = + case t of + Free(xn,xT) => (case assoc(vs,t) of + NONE => error "Variable not found in the list!!" + | SOME n => Var n) + | Const("0",iT) => Cst 0 + | Const("1",iT) => Cst 1 + | Bound i => Var (nat i) + | Const("uminus",_)$t' => Neg (i_of_term vs t') + | Const ("op +",_)$t1$t2 => Add (i_of_term vs t1,i_of_term vs t2) + | Const ("op -",_)$t1$t2 => Sub (i_of_term vs t1,i_of_term vs t2) + | Const ("op *",_)$t1$t2 => Mult (i_of_term vs t1,i_of_term vs t2) + | Const ("Numeral.number_of",_)$t' => Cst (HOLogic.dest_binum t') + | _ => error "i_of_term: unknown term"; + + +(* pseudo reification : term -> QF *) +fun qf_of_term vs t = + case t of + Const("True",_) => T + | Const("False",_) => F + | Const("op <",_)$t1$t2 => Lt (i_of_term vs t1,i_of_term vs t2) + | Const("op <=",_)$t1$t2 => Le (i_of_term vs t1,i_of_term vs t2) + | Const ("Divides.op dvd",_)$t1$t2 => + Divides(i_of_term vs t1,i_of_term vs t2) + | Const("op =",eqT)$t1$t2 => + if (domain_type eqT = HOLogic.intT) + then let val i1 = i_of_term vs t1 + val i2 = i_of_term vs t2 + in Eq (i1,i2) + end + else Equ(qf_of_term vs t1,qf_of_term vs t2) + | Const("op &",_)$t1$t2 => And(qf_of_term vs t1,qf_of_term vs t2) + | Const("op |",_)$t1$t2 => Or(qf_of_term vs t1,qf_of_term vs t2) + | Const("op -->",_)$t1$t2 => Imp(qf_of_term vs t1,qf_of_term vs t2) + | Const("Not",_)$t' => NOT(qf_of_term vs t') + | Const("Ex",_)$Abs(xn,xT,p) => + QEx(qf_of_term (map (fn(v,n) => (v,n + 1)) vs) p) + | Const("All",_)$Abs(xn,xT,p) => + QAll(qf_of_term (map (fn(v,n) => (v,n + 1)) vs) p) + | _ => error "qf_of_term : unknown term!"; + +(* +fun parse s = term_of (read_cterm (sign_of Main.thy) (s,HOLogic.boolT)); + +val t = parse "ALL (i::int) (j::int). i < 8* j --> (i - 1 = j + 3 + 2*j) & (j <= -i + k ) | 4 = i | 5 dvd i"; +*) +fun zip [] [] = [] + | zip (x::xs) (y::ys) = (x,y)::(zip xs ys); + + +fun start_vs t = + let val fs = term_frees t + in zip fs (map nat (0 upto (length fs - 1))) + end ; + +(* transform intterm and QF back to terms *) +val iT = HOLogic.intT; +val bT = HOLogic.boolT; +fun myassoc2 l v = + case l of + [] => NONE + | (x,v')::xs => if v = v' then SOME x + else myassoc2 xs v; + +fun term_of_i vs t = + case t of + Cst i => CooperDec.mk_numeral i + | Var n => valOf (myassoc2 vs n) + | Neg t' => Const("uminus",iT --> iT)$(term_of_i vs t') + | Add(t1,t2) => Const("op +",[iT,iT] ---> iT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Sub(t1,t2) => Const("op -",[iT,iT] ---> iT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Mult(t1,t2) => Const("op *",[iT,iT] ---> iT)$ + (term_of_i vs t1)$(term_of_i vs t2); + +fun term_of_qf vs t = + case t of + T => HOLogic.true_const + | F => HOLogic.false_const + | Lt(t1,t2) => Const("op <",[iT,iT] ---> bT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Le(t1,t2) => Const("op <=",[iT,iT] ---> bT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Gt(t1,t2) => Const("op <",[iT,iT] ---> bT)$ + (term_of_i vs t2)$(term_of_i vs t1) + | Ge(t1,t2) => Const("op <=",[iT,iT] ---> bT)$ + (term_of_i vs t2)$(term_of_i vs t1) + | Eq(t1,t2) => Const("op =",[iT,iT] ---> bT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Divides(t1,t2) => Const("Divides.op dvd",[iT,iT] ---> bT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | NOT t' => HOLogic.Not$(term_of_qf vs t') + | And(t1,t2) => HOLogic.conj$(term_of_qf vs t1)$(term_of_qf vs t2) + | Or(t1,t2) => HOLogic.disj$(term_of_qf vs t1)$(term_of_qf vs t2) + | Imp(t1,t2) => HOLogic.imp$(term_of_qf vs t1)$(term_of_qf vs t2) + | Equ(t1,t2) => (HOLogic.eq_const bT)$(term_of_qf vs t1)$ + (term_of_qf vs t2) + | _ => error "If this is raised, Isabelle/HOL or generate_code is inconsistent!"; + +(* The oracle *) + exception COOPER; + +fun presburger_oracle thy t = + let val vs = start_vs t + val result = lift_un (term_of_qf vs) (pa (qf_of_term vs t)) + in + case result of + None => raise COOPER + | Some t' => HOLogic.mk_Trueprop (HOLogic.mk_eq(t,t')) + end ; + +end; \ No newline at end of file diff -r afaa031ed4da -r 105519771c67 src/HOL/Integ/reflected_presburger.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Integ/reflected_presburger.ML Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,2172 @@ +(* $Id$ *) + + (* Caution: This file should not be modified. *) + (* It is autmatically generated from HOL/ex/Reflected_Presburger.thy *) +fun nat i = if i < 0 then 0 else i; +structure Generated = +struct + +datatype intterm = Cst of int | Var of int | Neg of intterm + | Add of intterm * intterm | Sub of intterm * intterm + | Mult of intterm * intterm; + +datatype QF = Lt of intterm * intterm | Gt of intterm * intterm + | Le of intterm * intterm | Ge of intterm * intterm | Eq of intterm * intterm + | Divides of intterm * intterm | T | F | NOT of QF | And of QF * QF + | Or of QF * QF | Imp of QF * QF | Equ of QF * QF | QAll of QF | QEx of QF; + +datatype 'a option = None | Some of 'a; + +fun lift_un c None = None + | lift_un c (Some p) = Some (c p); + +fun lift_bin (c, (Some a, Some b)) = Some (c a b) + | lift_bin (c, (None, y)) = None + | lift_bin (c, (Some y, None)) = None; + +fun lift_qe qe None = None + | lift_qe qe (Some p) = qe p; + +fun qelim (qe, QAll p) = lift_un NOT (lift_qe qe (lift_un NOT (qelim (qe, p)))) + | qelim (qe, QEx p) = lift_qe qe (qelim (qe, p)) + | qelim (qe, And (p, q)) = + lift_bin ((fn x => fn xa => And (x, xa)), (qelim (qe, p), qelim (qe, q))) + | qelim (qe, Or (p, q)) = + lift_bin ((fn x => fn xa => Or (x, xa)), (qelim (qe, p), qelim (qe, q))) + | qelim (qe, Imp (p, q)) = + lift_bin ((fn x => fn xa => Imp (x, xa)), (qelim (qe, p), qelim (qe, q))) + | qelim (qe, Equ (p, q)) = + lift_bin ((fn x => fn xa => Equ (x, xa)), (qelim (qe, p), qelim (qe, q))) + | qelim (qe, NOT p) = lift_un NOT (qelim (qe, p)) + | qelim (qe, Lt (w, x)) = Some (Lt (w, x)) + | qelim (qe, Gt (y, z)) = Some (Gt (y, z)) + | qelim (qe, Le (aa, ab)) = Some (Le (aa, ab)) + | qelim (qe, Ge (ac, ad)) = Some (Ge (ac, ad)) + | qelim (qe, Eq (ae, af)) = Some (Eq (ae, af)) + | qelim (qe, Divides (ag, ah)) = Some (Divides (ag, ah)) + | qelim (qe, T) = Some T + | qelim (qe, F) = Some F; + +fun lin_mul (c, Cst i) = Cst (c * i) + | lin_mul (c, Add (Mult (Cst c', Var n), r)) = + (if (c = 0) then Cst 0 + else Add (Mult (Cst (c * c'), Var n), lin_mul (c, r))); + +fun op_60_def0 m n = ((m) < (n)); + +fun op_60_61_def0 m n = not (op_60_def0 n m); + +fun lin_add (Add (Mult (Cst c1, Var n1), r1), Add (Mult (Cst c2, Var n2), r2)) = + (if (n1 = n2) + then let val c = Cst (c1 + c2) + in (if ((c1 + c2) = 0) then lin_add (r1, r2) + else Add (Mult (c, Var n1), lin_add (r1, r2))) + end + else (if op_60_61_def0 n1 n2 + then Add (Mult (Cst c1, Var n1), + lin_add (r1, Add (Mult (Cst c2, Var n2), r2))) + else Add (Mult (Cst c2, Var n2), + lin_add (Add (Mult (Cst c1, Var n1), r1), r2)))) + | lin_add (Add (Mult (Cst c1, Var n1), r1), Cst b) = + Add (Mult (Cst c1, Var n1), lin_add (r1, Cst b)) + | lin_add (Cst x, Add (Mult (Cst c2, Var n2), r2)) = + Add (Mult (Cst c2, Var n2), lin_add (Cst x, r2)) + | lin_add (Cst b1, Cst b2) = Cst (b1 + b2); + +fun lin_neg i = lin_mul (~1, i); + +fun linearize (Cst b) = Some (Cst b) + | linearize (Var n) = Some (Add (Mult (Cst 1, Var n), Cst 0)) + | linearize (Neg i) = lift_un lin_neg (linearize i) + | linearize (Add (i, j)) = + lift_bin ((fn x => fn y => lin_add (x, y)), (linearize i, linearize j)) + | linearize (Sub (i, j)) = + lift_bin + ((fn x => fn y => lin_add (x, lin_neg y)), (linearize i, linearize j)) + | linearize (Mult (i, j)) = + (case linearize i of None => None + | Some x => + (case x of + Cst xa => + (case linearize j of None => None + | Some x => Some (lin_mul (xa, x))) + | Var xa => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | Neg xa => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | Add (xa, xb) => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | Sub (xa, xb) => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | Mult (xa, xb) => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)))); + +fun linform (Le (it1, it2)) = + lift_bin + ((fn x => fn y => Le (lin_add (x, lin_neg y), Cst 0)), + (linearize it1, linearize it2)) + | linform (Eq (it1, it2)) = + lift_bin + ((fn x => fn y => Eq (lin_add (x, lin_neg y), Cst 0)), + (linearize it1, linearize it2)) + | linform (Divides (d, t)) = + (case linearize d of None => None + | Some x => + (case x of + Cst xa => + (if (xa = 0) then None + else (case linearize t of None => None + | Some xa => Some (Divides (x, xa)))) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | linform T = Some T + | linform F = Some F + | linform (NOT p) = lift_un NOT (linform p) + | linform (And (p, q)) = + lift_bin ((fn f => fn g => And (f, g)), (linform p, linform q)) + | linform (Or (p, q)) = + lift_bin ((fn f => fn g => Or (f, g)), (linform p, linform q)); + +fun nnf (Lt (it1, it2)) = Le (Sub (it1, it2), Cst (~ 1)) + | nnf (Gt (it1, it2)) = Le (Sub (it2, it1), Cst (~ 1)) + | nnf (Le (it1, it2)) = Le (it1, it2) + | nnf (Ge (it1, it2)) = Le (it2, it1) + | nnf (Eq (it1, it2)) = Eq (it2, it1) + | nnf (Divides (d, t)) = Divides (d, t) + | nnf T = T + | nnf F = F + | nnf (And (p, q)) = And (nnf p, nnf q) + | nnf (Or (p, q)) = Or (nnf p, nnf q) + | nnf (Imp (p, q)) = Or (nnf (NOT p), nnf q) + | nnf (Equ (p, q)) = Or (And (nnf p, nnf q), And (nnf (NOT p), nnf (NOT q))) + | nnf (NOT (Lt (it1, it2))) = Le (it2, it1) + | nnf (NOT (Gt (it1, it2))) = Le (it1, it2) + | nnf (NOT (Le (it1, it2))) = Le (Sub (it2, it1), Cst (~ 1)) + | nnf (NOT (Ge (it1, it2))) = Le (Sub (it1, it2), Cst (~ 1)) + | nnf (NOT (Eq (it1, it2))) = NOT (Eq (it1, it2)) + | nnf (NOT (Divides (d, t))) = NOT (Divides (d, t)) + | nnf (NOT T) = F + | nnf (NOT F) = T + | nnf (NOT (NOT p)) = nnf p + | nnf (NOT (And (p, q))) = Or (nnf (NOT p), nnf (NOT q)) + | nnf (NOT (Or (p, q))) = And (nnf (NOT p), nnf (NOT q)) + | nnf (NOT (Imp (p, q))) = And (nnf p, nnf (NOT q)) + | nnf (NOT (Equ (p, q))) = + Or (And (nnf p, nnf (NOT q)), And (nnf (NOT p), nnf q)); + +fun op_45_def2 z w = (z + ~ w); + +fun op_45_def0 m n = nat (op_45_def2 (m) (n)); + +val id_1_def0 : int = (0 + 1); + +fun decrvarsI (Cst i) = Cst i + | decrvarsI (Var n) = Var (op_45_def0 n id_1_def0) + | decrvarsI (Neg a) = Neg (decrvarsI a) + | decrvarsI (Add (a, b)) = Add (decrvarsI a, decrvarsI b) + | decrvarsI (Sub (a, b)) = Sub (decrvarsI a, decrvarsI b) + | decrvarsI (Mult (a, b)) = Mult (decrvarsI a, decrvarsI b); + +fun decrvars (Lt (a, b)) = Lt (decrvarsI a, decrvarsI b) + | decrvars (Gt (a, b)) = Gt (decrvarsI a, decrvarsI b) + | decrvars (Le (a, b)) = Le (decrvarsI a, decrvarsI b) + | decrvars (Ge (a, b)) = Ge (decrvarsI a, decrvarsI b) + | decrvars (Eq (a, b)) = Eq (decrvarsI a, decrvarsI b) + | decrvars (Divides (a, b)) = Divides (decrvarsI a, decrvarsI b) + | decrvars T = T + | decrvars F = F + | decrvars (NOT p) = NOT (decrvars p) + | decrvars (And (p, q)) = And (decrvars p, decrvars q) + | decrvars (Or (p, q)) = Or (decrvars p, decrvars q) + | decrvars (Imp (p, q)) = Imp (decrvars p, decrvars q) + | decrvars (Equ (p, q)) = Equ (decrvars p, decrvars q); + +fun op_64 [] ys = ys + | op_64 (x :: xs) ys = (x :: op_64 xs ys); + +fun map f [] = [] + | map f (x :: xs) = (f x :: map f xs); + +fun iupto (i, j) = (if (j < i) then [] else (i :: iupto ((i + 1), j))); + +fun all_sums (j, []) = [] + | all_sums (j, (i :: is)) = + op_64 (map (fn x => lin_add (i, Cst x)) (iupto (1, j))) (all_sums (j, is)); + +fun split x = (fn p => x (fst p) (snd p)); + +fun negateSnd x = split (fn q => fn r => (q, ~ r)) x; + +fun adjust b = + (fn (q, r) => + (if (0 <= op_45_def2 r b) then (((2 * q) + 1), op_45_def2 r b) + else ((2 * q), r))); + +fun negDivAlg (a, b) = + (if ((0 <= (a + b)) orelse (b <= 0)) then (~1, (a + b)) + else adjust b (negDivAlg (a, (2 * b)))); + +fun posDivAlg (a, b) = + (if ((a < b) orelse (b <= 0)) then (0, a) + else adjust b (posDivAlg (a, (2 * b)))); + +fun divAlg x = + split (fn a => fn b => + (if (0 <= a) + then (if (0 <= b) then posDivAlg (a, b) + else (if (a = 0) then (0, 0) + else negateSnd (negDivAlg (~ a, ~ b)))) + else (if (0 < b) then negDivAlg (a, b) + else negateSnd (posDivAlg (~ a, ~ b))))) + x; + +fun op_mod_def1 a b = snd (divAlg (a, b)); + +fun op_dvd m n = (op_mod_def1 n m = 0); + +fun psimpl (Le (l, r)) = + (case lift_bin + ((fn x => fn y => lin_add (x, lin_neg y)), + (linearize l, linearize r)) of + None => Le (l, r) + | Some x => + (case x of Cst xa => (if (xa <= 0) then T else F) + | Var xa => Le (x, Cst 0) | Neg xa => Le (x, Cst 0) + | Add (xa, xb) => Le (x, Cst 0) | Sub (xa, xb) => Le (x, Cst 0) + | Mult (xa, xb) => Le (x, Cst 0))) + | psimpl (Eq (l, r)) = + (case lift_bin + ((fn x => fn y => lin_add (x, lin_neg y)), + (linearize l, linearize r)) of + None => Eq (l, r) + | Some x => + (case x of Cst xa => (if (xa = 0) then T else F) + | Var xa => Eq (x, Cst 0) | Neg xa => Eq (x, Cst 0) + | Add (xa, xb) => Eq (x, Cst 0) | Sub (xa, xb) => Eq (x, Cst 0) + | Mult (xa, xb) => Eq (x, Cst 0))) + | psimpl (Divides (Cst d, t)) = + (case linearize t of None => Divides (Cst d, t) + | Some x => + (case x of Cst xa => (if op_dvd d xa then T else F) + | Var xa => Divides (Cst d, x) | Neg xa => Divides (Cst d, x) + | Add (xa, xb) => Divides (Cst d, x) + | Sub (xa, xb) => Divides (Cst d, x) + | Mult (xa, xb) => Divides (Cst d, x))) + | psimpl (Equ (p, q)) = + let val p' = psimpl p; val q' = psimpl q + in (case p' of + Lt (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Gt (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Le (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Ge (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Eq (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Divides (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | T => q' + | F => (case q' of Lt (x, xa) => NOT q' | Gt (x, xa) => NOT q' + | Le (x, xa) => NOT q' | Ge (x, xa) => NOT q' + | Eq (x, xa) => NOT q' | Divides (x, xa) => NOT q' | T => F + | F => T | NOT x => x | And (x, xa) => NOT q' + | Or (x, xa) => NOT q' | Imp (x, xa) => NOT q' + | Equ (x, xa) => NOT q' | QAll x => NOT q' | QEx x => NOT q') + | NOT x => + (case q' of Lt (xa, xb) => Equ (p', q') + | Gt (xa, xb) => Equ (p', q') | Le (xa, xb) => Equ (p', q') + | Ge (xa, xb) => Equ (p', q') | Eq (xa, xb) => Equ (p', q') + | Divides (xa, xb) => Equ (p', q') | T => p' | F => x + | NOT xa => Equ (x, xa) | And (xa, xb) => Equ (p', q') + | Or (xa, xb) => Equ (p', q') | Imp (xa, xb) => Equ (p', q') + | Equ (xa, xb) => Equ (p', q') | QAll xa => Equ (p', q') + | QEx xa => Equ (p', q')) + | And (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Or (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Imp (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Equ (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | QAll x => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | QEx x => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q'))) + end + | psimpl (NOT p) = + let val p' = psimpl p + in (case p' of Lt (x, xa) => NOT p' | Gt (x, xa) => NOT p' + | Le (x, xa) => NOT p' | Ge (x, xa) => NOT p' | Eq (x, xa) => NOT p' + | Divides (x, xa) => NOT p' | T => F | F => T | NOT x => x + | And (x, xa) => NOT p' | Or (x, xa) => NOT p' | Imp (x, xa) => NOT p' + | Equ (x, xa) => NOT p' | QAll x => NOT p' | QEx x => NOT p') + end + | psimpl (Lt (u, v)) = Lt (u, v) + | psimpl (Gt (w, x)) = Gt (w, x) + | psimpl (Ge (aa, ab)) = Ge (aa, ab) + | psimpl (Divides (Var bp, af)) = Divides (Var bp, af) + | psimpl (Divides (Neg bq, af)) = Divides (Neg bq, af) + | psimpl (Divides (Add (br, bs), af)) = Divides (Add (br, bs), af) + | psimpl (Divides (Sub (bt, bu), af)) = Divides (Sub (bt, bu), af) + | psimpl (Divides (Mult (bv, bw), af)) = Divides (Mult (bv, bw), af) + | psimpl T = T + | psimpl F = F + | psimpl (QAll ap) = QAll ap + | psimpl (QEx aq) = QEx aq + | psimpl (And (p, q)) = + let val p' = psimpl p + in (case p' of + Lt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Gt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Le (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Ge (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Eq (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Divides (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | T => psimpl q | F => F + | NOT x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | And (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Or (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Imp (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Equ (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | QAll x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | QEx x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end) + end + | psimpl (Or (p, q)) = + let val p' = psimpl p + in (case p' of + Lt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') | Gt (x, xa) => Or (p', q') + | Le (x, xa) => Or (p', q') | Ge (x, xa) => Or (p', q') + | Eq (x, xa) => Or (p', q') | Divides (x, xa) => Or (p', q') + | T => T | F => p' | NOT x => Or (p', q') + | And (x, xa) => Or (p', q') | Or (x, xa) => Or (p', q') + | Imp (x, xa) => Or (p', q') | Equ (x, xa) => Or (p', q') + | QAll x => Or (p', q') | QEx x => Or (p', q')) + end + | Gt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Le (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Ge (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Eq (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Divides (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | T => T | F => psimpl q + | NOT x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | And (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Or (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Imp (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Equ (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | QAll x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | QEx x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end) + end + | psimpl (Imp (p, q)) = + let val p' = psimpl p + in (case p' of + Lt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Gt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Le (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Ge (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Eq (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Divides (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | T => psimpl q | F => T + | NOT x => + let val q' = psimpl q + in (case q' of Lt (xa, xb) => Or (x, q') + | Gt (xa, xb) => Or (x, q') | Le (xa, xb) => Or (x, q') + | Ge (xa, xb) => Or (x, q') | Eq (xa, xb) => Or (x, q') + | Divides (xa, xb) => Or (x, q') | T => T | F => x + | NOT xa => Or (x, q') | And (xa, xb) => Or (x, q') + | Or (xa, xb) => Or (x, q') | Imp (xa, xb) => Or (x, q') + | Equ (xa, xb) => Or (x, q') | QAll xa => Or (x, q') + | QEx xa => Or (x, q')) + end + | And (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Or (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Imp (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Equ (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | QAll x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | QEx x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end) + end; + +fun subst_it i (Cst b) = Cst b + | subst_it i (Var n) = (if (n = 0) then i else Var n) + | subst_it i (Neg it) = Neg (subst_it i it) + | subst_it i (Add (it1, it2)) = Add (subst_it i it1, subst_it i it2) + | subst_it i (Sub (it1, it2)) = Sub (subst_it i it1, subst_it i it2) + | subst_it i (Mult (it1, it2)) = Mult (subst_it i it1, subst_it i it2); + +fun subst_p i (Le (it1, it2)) = Le (subst_it i it1, subst_it i it2) + | subst_p i (Lt (it1, it2)) = Lt (subst_it i it1, subst_it i it2) + | subst_p i (Ge (it1, it2)) = Ge (subst_it i it1, subst_it i it2) + | subst_p i (Gt (it1, it2)) = Gt (subst_it i it1, subst_it i it2) + | subst_p i (Eq (it1, it2)) = Eq (subst_it i it1, subst_it i it2) + | subst_p i (Divides (d, t)) = Divides (subst_it i d, subst_it i t) + | subst_p i T = T + | subst_p i F = F + | subst_p i (And (p, q)) = And (subst_p i p, subst_p i q) + | subst_p i (Or (p, q)) = Or (subst_p i p, subst_p i q) + | subst_p i (Imp (p, q)) = Imp (subst_p i p, subst_p i q) + | subst_p i (Equ (p, q)) = Equ (subst_p i p, subst_p i q) + | subst_p i (NOT p) = NOT (subst_p i p); + +fun explode_disj ([], p) = F + | explode_disj ((i :: is), p) = + let val pi = psimpl (subst_p i p) + in (case pi of + Lt (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Gt (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Le (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Ge (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Eq (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Divides (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | T => T | F => explode_disj (is, p) + | NOT x => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | And (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Or (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Imp (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Equ (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | QAll x => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | QEx x => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end) + end; + +fun minusinf (And (p, q)) = And (minusinf p, minusinf q) + | minusinf (Or (p, q)) = Or (minusinf p, minusinf q) + | minusinf (Lt (u, v)) = Lt (u, v) + | minusinf (Gt (w, x)) = Gt (w, x) + | minusinf (Le (Cst bo, z)) = Le (Cst bo, z) + | minusinf (Le (Var bp, z)) = Le (Var bp, z) + | minusinf (Le (Neg bq, z)) = Le (Neg bq, z) + | minusinf (Le (Add (Cst cg, bs), z)) = Le (Add (Cst cg, bs), z) + | minusinf (Le (Add (Var ch, bs), z)) = Le (Add (Var ch, bs), z) + | minusinf (Le (Add (Neg ci, bs), z)) = Le (Add (Neg ci, bs), z) + | minusinf (Le (Add (Add (cj, ck), bs), z)) = Le (Add (Add (cj, ck), bs), z) + | minusinf (Le (Add (Sub (cl, cm), bs), z)) = Le (Add (Sub (cl, cm), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = + Le (Add (Mult (Cst cy, Cst dq), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Var ei), bs), z)) = + (if (ei = 0) then (if (cy < 0) then F else T) + else Le (Add (Mult (Cst cy, Var (op_45_def0 ei id_1_def0 + 1)), bs), z)) + | minusinf (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = + Le (Add (Mult (Cst cy, Neg ds), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = + Le (Add (Mult (Cst cy, Add (dt, du)), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = + Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = + Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z) + | minusinf (Le (Add (Mult (Var cz, co), bs), z)) = + Le (Add (Mult (Var cz, co), bs), z) + | minusinf (Le (Add (Mult (Neg da, co), bs), z)) = + Le (Add (Mult (Neg da, co), bs), z) + | minusinf (Le (Add (Mult (Add (db, dc), co), bs), z)) = + Le (Add (Mult (Add (db, dc), co), bs), z) + | minusinf (Le (Add (Mult (Sub (dd, de), co), bs), z)) = + Le (Add (Mult (Sub (dd, de), co), bs), z) + | minusinf (Le (Add (Mult (Mult (df, dg), co), bs), z)) = + Le (Add (Mult (Mult (df, dg), co), bs), z) + | minusinf (Le (Sub (bt, bu), z)) = Le (Sub (bt, bu), z) + | minusinf (Le (Mult (bv, bw), z)) = Le (Mult (bv, bw), z) + | minusinf (Ge (aa, ab)) = Ge (aa, ab) + | minusinf (Eq (Cst ek, ad)) = Eq (Cst ek, ad) + | minusinf (Eq (Var el, ad)) = Eq (Var el, ad) + | minusinf (Eq (Neg em, ad)) = Eq (Neg em, ad) + | minusinf (Eq (Add (Cst fc, eo), ad)) = Eq (Add (Cst fc, eo), ad) + | minusinf (Eq (Add (Var fd, eo), ad)) = Eq (Add (Var fd, eo), ad) + | minusinf (Eq (Add (Neg fe, eo), ad)) = Eq (Add (Neg fe, eo), ad) + | minusinf (Eq (Add (Add (ff, fg), eo), ad)) = Eq (Add (Add (ff, fg), eo), ad) + | minusinf (Eq (Add (Sub (fh, fi), eo), ad)) = Eq (Add (Sub (fh, fi), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = + Eq (Add (Mult (Cst fu, Cst gm), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = + (if (he = 0) then F + else Eq (Add (Mult (Cst fu, Var (op_45_def0 he id_1_def0 + 1)), eo), ad)) + | minusinf (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = + Eq (Add (Mult (Cst fu, Neg go), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = + Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = + Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = + Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad) + | minusinf (Eq (Add (Mult (Var fv, fk), eo), ad)) = + Eq (Add (Mult (Var fv, fk), eo), ad) + | minusinf (Eq (Add (Mult (Neg fw, fk), eo), ad)) = + Eq (Add (Mult (Neg fw, fk), eo), ad) + | minusinf (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = + Eq (Add (Mult (Add (fx, fy), fk), eo), ad) + | minusinf (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = + Eq (Add (Mult (Sub (fz, ga), fk), eo), ad) + | minusinf (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = + Eq (Add (Mult (Mult (gb, gc), fk), eo), ad) + | minusinf (Eq (Sub (ep, eq), ad)) = Eq (Sub (ep, eq), ad) + | minusinf (Eq (Mult (er, es), ad)) = Eq (Mult (er, es), ad) + | minusinf (Divides (ae, af)) = Divides (ae, af) + | minusinf T = T + | minusinf F = F + | minusinf (NOT (Lt (hg, hh))) = NOT (Lt (hg, hh)) + | minusinf (NOT (Gt (hi, hj))) = NOT (Gt (hi, hj)) + | minusinf (NOT (Le (hk, hl))) = NOT (Le (hk, hl)) + | minusinf (NOT (Ge (hm, hn))) = NOT (Ge (hm, hn)) + | minusinf (NOT (Eq (Cst ja, hp))) = NOT (Eq (Cst ja, hp)) + | minusinf (NOT (Eq (Var jb, hp))) = NOT (Eq (Var jb, hp)) + | minusinf (NOT (Eq (Neg jc, hp))) = NOT (Eq (Neg jc, hp)) + | minusinf (NOT (Eq (Add (Cst js, je), hp))) = NOT (Eq (Add (Cst js, je), hp)) + | minusinf (NOT (Eq (Add (Var jt, je), hp))) = NOT (Eq (Add (Var jt, je), hp)) + | minusinf (NOT (Eq (Add (Neg ju, je), hp))) = NOT (Eq (Add (Neg ju, je), hp)) + | minusinf (NOT (Eq (Add (Add (jv, jw), je), hp))) = + NOT (Eq (Add (Add (jv, jw), je), hp)) + | minusinf (NOT (Eq (Add (Sub (jx, jy), je), hp))) = + NOT (Eq (Add (Sub (jx, jy), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = + (if (lu = 0) then T + else NOT (Eq (Add (Mult (Cst kk, Var (op_45_def0 lu id_1_def0 + 1)), je), + hp))) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = + NOT (Eq (Add (Mult (Var kl, ka), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = + NOT (Eq (Add (Mult (Neg km, ka), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = + NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = + NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = + NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp)) + | minusinf (NOT (Eq (Sub (jf, jg), hp))) = NOT (Eq (Sub (jf, jg), hp)) + | minusinf (NOT (Eq (Mult (jh, ji), hp))) = NOT (Eq (Mult (jh, ji), hp)) + | minusinf (NOT (Divides (hq, hr))) = NOT (Divides (hq, hr)) + | minusinf (NOT T) = NOT T + | minusinf (NOT F) = NOT F + | minusinf (NOT (NOT hs)) = NOT (NOT hs) + | minusinf (NOT (And (ht, hu))) = NOT (And (ht, hu)) + | minusinf (NOT (Or (hv, hw))) = NOT (Or (hv, hw)) + | minusinf (NOT (Imp (hx, hy))) = NOT (Imp (hx, hy)) + | minusinf (NOT (Equ (hz, ia))) = NOT (Equ (hz, ia)) + | minusinf (NOT (QAll ib)) = NOT (QAll ib) + | minusinf (NOT (QEx ic)) = NOT (QEx ic) + | minusinf (Imp (al, am)) = Imp (al, am) + | minusinf (Equ (an, ao)) = Equ (an, ao) + | minusinf (QAll ap) = QAll ap + | minusinf (QEx aq) = QEx aq; + +fun abs i = (if (i < 0) then ~ i else i); + +fun op_div_def1 a b = fst (divAlg (a, b)); + +fun op_mod_def0 m n = nat (op_mod_def1 (m) (n)); + +fun ngcd (m, n) = (if (n = 0) then m else ngcd (n, op_mod_def0 m n)); + +fun igcd x = split (fn a => fn b => (ngcd (nat (abs a), nat (abs b)))) x; + +fun ilcm a b = op_div_def1 (a * b) (igcd (a, b)); + +fun divlcm (NOT p) = divlcm p + | divlcm (And (p, q)) = ilcm (divlcm p) (divlcm q) + | divlcm (Or (p, q)) = ilcm (divlcm p) (divlcm q) + | divlcm (Lt (u, v)) = 1 + | divlcm (Gt (w, x)) = 1 + | divlcm (Le (y, z)) = 1 + | divlcm (Ge (aa, ab)) = 1 + | divlcm (Eq (ac, ad)) = 1 + | divlcm (Divides (Cst bo, Cst cg)) = 1 + | divlcm (Divides (Cst bo, Var ch)) = 1 + | divlcm (Divides (Cst bo, Neg ci)) = 1 + | divlcm (Divides (Cst bo, Add (Cst cy, ck))) = 1 + | divlcm (Divides (Cst bo, Add (Var cz, ck))) = 1 + | divlcm (Divides (Cst bo, Add (Neg da, ck))) = 1 + | divlcm (Divides (Cst bo, Add (Add (db, dc), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Sub (dd, de), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Cst ei), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Var fa), ck))) = + (if (fa = 0) then abs bo else 1) + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Neg ek), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Add (el, em)), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Sub (en, eo)), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Mult (ep, eq)), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Var dr, dg), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Neg ds, dg), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Add (dt, du), dg), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Sub (dv, dw), dg), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Mult (dx, dy), dg), ck))) = 1 + | divlcm (Divides (Cst bo, Sub (cl, cm))) = 1 + | divlcm (Divides (Cst bo, Mult (cn, co))) = 1 + | divlcm (Divides (Var bp, af)) = 1 + | divlcm (Divides (Neg bq, af)) = 1 + | divlcm (Divides (Add (br, bs), af)) = 1 + | divlcm (Divides (Sub (bt, bu), af)) = 1 + | divlcm (Divides (Mult (bv, bw), af)) = 1 + | divlcm T = 1 + | divlcm F = 1 + | divlcm (Imp (al, am)) = 1 + | divlcm (Equ (an, ao)) = 1 + | divlcm (QAll ap) = 1 + | divlcm (QEx aq) = 1; + +fun explode_minf (q, B) = + let val d = divlcm q; val pm = minusinf q; + val dj1 = explode_disj (map (fn x => Cst x) (iupto (1, d)), pm) + in (case dj1 of + Lt (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Gt (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Le (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Ge (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Eq (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Divides (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | T => T | F => explode_disj (all_sums (d, B), q) + | NOT x => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | And (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Or (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Imp (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Equ (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | QAll x => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | QEx x => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end) + end; + +fun mirror (And (p, q)) = And (mirror p, mirror q) + | mirror (Or (p, q)) = Or (mirror p, mirror q) + | mirror (Lt (u, v)) = Lt (u, v) + | mirror (Gt (w, x)) = Gt (w, x) + | mirror (Le (Cst bp, aa)) = Le (Cst bp, aa) + | mirror (Le (Var bq, aa)) = Le (Var bq, aa) + | mirror (Le (Neg br, aa)) = Le (Neg br, aa) + | mirror (Le (Add (Cst ch, bt), aa)) = Le (Add (Cst ch, bt), aa) + | mirror (Le (Add (Var ci, bt), aa)) = Le (Add (Var ci, bt), aa) + | mirror (Le (Add (Neg cj, bt), aa)) = Le (Add (Neg cj, bt), aa) + | mirror (Le (Add (Add (ck, cl), bt), aa)) = Le (Add (Add (ck, cl), bt), aa) + | mirror (Le (Add (Sub (cm, cn), bt), aa)) = Le (Add (Sub (cm, cn), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Cst dr), bt), aa)) = + Le (Add (Mult (Cst cz, Cst dr), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Var ej), bt), aa)) = + (if (ej = 0) then Le (Add (Mult (Cst (~ cz), Var 0), bt), aa) + else Le (Add (Mult (Cst cz, Var (op_45_def0 ej id_1_def0 + 1)), bt), aa)) + | mirror (Le (Add (Mult (Cst cz, Neg dt), bt), aa)) = + Le (Add (Mult (Cst cz, Neg dt), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Add (du, dv)), bt), aa)) = + Le (Add (Mult (Cst cz, Add (du, dv)), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Sub (dw, dx)), bt), aa)) = + Le (Add (Mult (Cst cz, Sub (dw, dx)), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Mult (dy, dz)), bt), aa)) = + Le (Add (Mult (Cst cz, Mult (dy, dz)), bt), aa) + | mirror (Le (Add (Mult (Var da, cp), bt), aa)) = + Le (Add (Mult (Var da, cp), bt), aa) + | mirror (Le (Add (Mult (Neg db, cp), bt), aa)) = + Le (Add (Mult (Neg db, cp), bt), aa) + | mirror (Le (Add (Mult (Add (dc, dd), cp), bt), aa)) = + Le (Add (Mult (Add (dc, dd), cp), bt), aa) + | mirror (Le (Add (Mult (Sub (de, df), cp), bt), aa)) = + Le (Add (Mult (Sub (de, df), cp), bt), aa) + | mirror (Le (Add (Mult (Mult (dg, dh), cp), bt), aa)) = + Le (Add (Mult (Mult (dg, dh), cp), bt), aa) + | mirror (Le (Sub (bu, bv), aa)) = Le (Sub (bu, bv), aa) + | mirror (Le (Mult (bw, bx), aa)) = Le (Mult (bw, bx), aa) + | mirror (Ge (ab, ac)) = Ge (ab, ac) + | mirror (Eq (Cst el, ae)) = Eq (Cst el, ae) + | mirror (Eq (Var em, ae)) = Eq (Var em, ae) + | mirror (Eq (Neg en, ae)) = Eq (Neg en, ae) + | mirror (Eq (Add (Cst fd, ep), ae)) = Eq (Add (Cst fd, ep), ae) + | mirror (Eq (Add (Var fe, ep), ae)) = Eq (Add (Var fe, ep), ae) + | mirror (Eq (Add (Neg ff, ep), ae)) = Eq (Add (Neg ff, ep), ae) + | mirror (Eq (Add (Add (fg, fh), ep), ae)) = Eq (Add (Add (fg, fh), ep), ae) + | mirror (Eq (Add (Sub (fi, fj), ep), ae)) = Eq (Add (Sub (fi, fj), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Cst gn), ep), ae)) = + Eq (Add (Mult (Cst fv, Cst gn), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Var hf), ep), ae)) = + (if (hf = 0) then Eq (Add (Mult (Cst (~ fv), Var 0), ep), ae) + else Eq (Add (Mult (Cst fv, Var (op_45_def0 hf id_1_def0 + 1)), ep), ae)) + | mirror (Eq (Add (Mult (Cst fv, Neg gp), ep), ae)) = + Eq (Add (Mult (Cst fv, Neg gp), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Add (gq, gr)), ep), ae)) = + Eq (Add (Mult (Cst fv, Add (gq, gr)), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Sub (gs, gt)), ep), ae)) = + Eq (Add (Mult (Cst fv, Sub (gs, gt)), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Mult (gu, gv)), ep), ae)) = + Eq (Add (Mult (Cst fv, Mult (gu, gv)), ep), ae) + | mirror (Eq (Add (Mult (Var fw, fl), ep), ae)) = + Eq (Add (Mult (Var fw, fl), ep), ae) + | mirror (Eq (Add (Mult (Neg fx, fl), ep), ae)) = + Eq (Add (Mult (Neg fx, fl), ep), ae) + | mirror (Eq (Add (Mult (Add (fy, fz), fl), ep), ae)) = + Eq (Add (Mult (Add (fy, fz), fl), ep), ae) + | mirror (Eq (Add (Mult (Sub (ga, gb), fl), ep), ae)) = + Eq (Add (Mult (Sub (ga, gb), fl), ep), ae) + | mirror (Eq (Add (Mult (Mult (gc, gd), fl), ep), ae)) = + Eq (Add (Mult (Mult (gc, gd), fl), ep), ae) + | mirror (Eq (Sub (eq, er), ae)) = Eq (Sub (eq, er), ae) + | mirror (Eq (Mult (es, et), ae)) = Eq (Mult (es, et), ae) + | mirror (Divides (Cst hh, Cst hz)) = Divides (Cst hh, Cst hz) + | mirror (Divides (Cst hh, Var ia)) = Divides (Cst hh, Var ia) + | mirror (Divides (Cst hh, Neg ib)) = Divides (Cst hh, Neg ib) + | mirror (Divides (Cst hh, Add (Cst ir, id))) = + Divides (Cst hh, Add (Cst ir, id)) + | mirror (Divides (Cst hh, Add (Var is, id))) = + Divides (Cst hh, Add (Var is, id)) + | mirror (Divides (Cst hh, Add (Neg it, id))) = + Divides (Cst hh, Add (Neg it, id)) + | mirror (Divides (Cst hh, Add (Add (iu, iv), id))) = + Divides (Cst hh, Add (Add (iu, iv), id)) + | mirror (Divides (Cst hh, Add (Sub (iw, ix), id))) = + Divides (Cst hh, Add (Sub (iw, ix), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Cst kb), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Cst kb), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Var kt), id))) = + (if (kt = 0) then Divides (Cst hh, Add (Mult (Cst (~ jj), Var 0), id)) + else Divides + (Cst hh, + Add (Mult (Cst jj, Var (op_45_def0 kt id_1_def0 + 1)), id))) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Neg kd), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Neg kd), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Add (ke, kf)), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Add (ke, kf)), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Sub (kg, kh)), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Sub (kg, kh)), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Mult (ki, kj)), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Mult (ki, kj)), id)) + | mirror (Divides (Cst hh, Add (Mult (Var jk, iz), id))) = + Divides (Cst hh, Add (Mult (Var jk, iz), id)) + | mirror (Divides (Cst hh, Add (Mult (Neg jl, iz), id))) = + Divides (Cst hh, Add (Mult (Neg jl, iz), id)) + | mirror (Divides (Cst hh, Add (Mult (Add (jm, jn), iz), id))) = + Divides (Cst hh, Add (Mult (Add (jm, jn), iz), id)) + | mirror (Divides (Cst hh, Add (Mult (Sub (jo, jp), iz), id))) = + Divides (Cst hh, Add (Mult (Sub (jo, jp), iz), id)) + | mirror (Divides (Cst hh, Add (Mult (Mult (jq, jr), iz), id))) = + Divides (Cst hh, Add (Mult (Mult (jq, jr), iz), id)) + | mirror (Divides (Cst hh, Sub (ie, if'))) = Divides (Cst hh, Sub (ie, if')) + | mirror (Divides (Cst hh, Mult (ig, ih))) = Divides (Cst hh, Mult (ig, ih)) + | mirror (Divides (Var hi, ag)) = Divides (Var hi, ag) + | mirror (Divides (Neg hj, ag)) = Divides (Neg hj, ag) + | mirror (Divides (Add (hk, hl), ag)) = Divides (Add (hk, hl), ag) + | mirror (Divides (Sub (hm, hn), ag)) = Divides (Sub (hm, hn), ag) + | mirror (Divides (Mult (ho, hp), ag)) = Divides (Mult (ho, hp), ag) + | mirror T = T + | mirror F = F + | mirror (NOT (Lt (kv, kw))) = NOT (Lt (kv, kw)) + | mirror (NOT (Gt (kx, ky))) = NOT (Gt (kx, ky)) + | mirror (NOT (Le (kz, la))) = NOT (Le (kz, la)) + | mirror (NOT (Ge (lb, lc))) = NOT (Ge (lb, lc)) + | mirror (NOT (Eq (Cst mp, le))) = NOT (Eq (Cst mp, le)) + | mirror (NOT (Eq (Var mq, le))) = NOT (Eq (Var mq, le)) + | mirror (NOT (Eq (Neg mr, le))) = NOT (Eq (Neg mr, le)) + | mirror (NOT (Eq (Add (Cst nh, mt), le))) = NOT (Eq (Add (Cst nh, mt), le)) + | mirror (NOT (Eq (Add (Var ni, mt), le))) = NOT (Eq (Add (Var ni, mt), le)) + | mirror (NOT (Eq (Add (Neg nj, mt), le))) = NOT (Eq (Add (Neg nj, mt), le)) + | mirror (NOT (Eq (Add (Add (nk, nl), mt), le))) = + NOT (Eq (Add (Add (nk, nl), mt), le)) + | mirror (NOT (Eq (Add (Sub (nm, nn), mt), le))) = + NOT (Eq (Add (Sub (nm, nn), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Cst or), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Cst or), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Var pj), mt), le))) = + (if (pj = 0) then NOT (Eq (Add (Mult (Cst (~ nz), Var 0), mt), le)) + else NOT (Eq (Add (Mult (Cst nz, Var (op_45_def0 pj id_1_def0 + 1)), mt), + le))) + | mirror (NOT (Eq (Add (Mult (Cst nz, Neg ot), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Neg ot), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Add (ou, ov)), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Add (ou, ov)), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Sub (ow, ox)), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Sub (ow, ox)), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Mult (oy, oz)), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Mult (oy, oz)), mt), le)) + | mirror (NOT (Eq (Add (Mult (Var oa, np), mt), le))) = + NOT (Eq (Add (Mult (Var oa, np), mt), le)) + | mirror (NOT (Eq (Add (Mult (Neg ob, np), mt), le))) = + NOT (Eq (Add (Mult (Neg ob, np), mt), le)) + | mirror (NOT (Eq (Add (Mult (Add (oc, od), np), mt), le))) = + NOT (Eq (Add (Mult (Add (oc, od), np), mt), le)) + | mirror (NOT (Eq (Add (Mult (Sub (oe, of'), np), mt), le))) = + NOT (Eq (Add (Mult (Sub (oe, of'), np), mt), le)) + | mirror (NOT (Eq (Add (Mult (Mult (og, oh), np), mt), le))) = + NOT (Eq (Add (Mult (Mult (og, oh), np), mt), le)) + | mirror (NOT (Eq (Sub (mu, mv), le))) = NOT (Eq (Sub (mu, mv), le)) + | mirror (NOT (Eq (Mult (mw, mx), le))) = NOT (Eq (Mult (mw, mx), le)) + | mirror (NOT (Divides (Cst pl, Cst qd))) = NOT (Divides (Cst pl, Cst qd)) + | mirror (NOT (Divides (Cst pl, Var qe))) = NOT (Divides (Cst pl, Var qe)) + | mirror (NOT (Divides (Cst pl, Neg qf))) = NOT (Divides (Cst pl, Neg qf)) + | mirror (NOT (Divides (Cst pl, Add (Cst qv, qh)))) = + NOT (Divides (Cst pl, Add (Cst qv, qh))) + | mirror (NOT (Divides (Cst pl, Add (Var qw, qh)))) = + NOT (Divides (Cst pl, Add (Var qw, qh))) + | mirror (NOT (Divides (Cst pl, Add (Neg qx, qh)))) = + NOT (Divides (Cst pl, Add (Neg qx, qh))) + | mirror (NOT (Divides (Cst pl, Add (Add (qy, qz), qh)))) = + NOT (Divides (Cst pl, Add (Add (qy, qz), qh))) + | mirror (NOT (Divides (Cst pl, Add (Sub (ra, rb), qh)))) = + NOT (Divides (Cst pl, Add (Sub (ra, rb), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Cst sf), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Cst sf), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Var sx), qh)))) = + (if (sx = 0) + then NOT (Divides (Cst pl, Add (Mult (Cst (~ rn), Var 0), qh))) + else NOT (Divides + (Cst pl, + Add (Mult (Cst rn, Var (op_45_def0 sx id_1_def0 + 1)), + qh)))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Neg sh), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Neg sh), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Add (si, sj)), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Add (si, sj)), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Sub (sk, sl)), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Sub (sk, sl)), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Mult (sm, sn)), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Mult (sm, sn)), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Var ro, rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Var ro, rd), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Neg rp, rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Neg rp, rd), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Add (rq, rr), rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Add (rq, rr), rd), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Sub (rs, rt), rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Sub (rs, rt), rd), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Mult (ru, rv), rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Mult (ru, rv), rd), qh))) + | mirror (NOT (Divides (Cst pl, Sub (qi, qj)))) = + NOT (Divides (Cst pl, Sub (qi, qj))) + | mirror (NOT (Divides (Cst pl, Mult (qk, ql)))) = + NOT (Divides (Cst pl, Mult (qk, ql))) + | mirror (NOT (Divides (Var pm, lg))) = NOT (Divides (Var pm, lg)) + | mirror (NOT (Divides (Neg pn, lg))) = NOT (Divides (Neg pn, lg)) + | mirror (NOT (Divides (Add (po, pp), lg))) = NOT (Divides (Add (po, pp), lg)) + | mirror (NOT (Divides (Sub (pq, pr), lg))) = NOT (Divides (Sub (pq, pr), lg)) + | mirror (NOT (Divides (Mult (ps, pt), lg))) = + NOT (Divides (Mult (ps, pt), lg)) + | mirror (NOT T) = NOT T + | mirror (NOT F) = NOT F + | mirror (NOT (NOT lh)) = NOT (NOT lh) + | mirror (NOT (And (li, lj))) = NOT (And (li, lj)) + | mirror (NOT (Or (lk, ll))) = NOT (Or (lk, ll)) + | mirror (NOT (Imp (lm, ln))) = NOT (Imp (lm, ln)) + | mirror (NOT (Equ (lo, lp))) = NOT (Equ (lo, lp)) + | mirror (NOT (QAll lq)) = NOT (QAll lq) + | mirror (NOT (QEx lr)) = NOT (QEx lr) + | mirror (Imp (am, an)) = Imp (am, an) + | mirror (Equ (ao, ap)) = Equ (ao, ap) + | mirror (QAll aq) = QAll aq + | mirror (QEx ar) = QEx ar; + +fun op_43_def0 m n = nat ((m) + (n)); + +fun size_def1 [] = 0 + | size_def1 (a :: list) = op_43_def0 (size_def1 list) (0 + 1); + +fun aset (And (p, q)) = op_64 (aset p) (aset q) + | aset (Or (p, q)) = op_64 (aset p) (aset q) + | aset (Lt (u, v)) = [] + | aset (Gt (w, x)) = [] + | aset (Le (Cst bo, z)) = [] + | aset (Le (Var bp, z)) = [] + | aset (Le (Neg bq, z)) = [] + | aset (Le (Add (Cst cg, bs), z)) = [] + | aset (Le (Add (Var ch, bs), z)) = [] + | aset (Le (Add (Neg ci, bs), z)) = [] + | aset (Le (Add (Add (cj, ck), bs), z)) = [] + | aset (Le (Add (Sub (cl, cm), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Var ei), bs), z)) = + (if (ei = 0) + then (if (cy < 0) then [lin_add (bs, Cst 1)] + else [lin_neg bs, lin_add (lin_neg bs, Cst 1)]) + else []) + | aset (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = [] + | aset (Le (Add (Mult (Var cz, co), bs), z)) = [] + | aset (Le (Add (Mult (Neg da, co), bs), z)) = [] + | aset (Le (Add (Mult (Add (db, dc), co), bs), z)) = [] + | aset (Le (Add (Mult (Sub (dd, de), co), bs), z)) = [] + | aset (Le (Add (Mult (Mult (df, dg), co), bs), z)) = [] + | aset (Le (Sub (bt, bu), z)) = [] + | aset (Le (Mult (bv, bw), z)) = [] + | aset (Ge (aa, ab)) = [] + | aset (Eq (Cst ek, ad)) = [] + | aset (Eq (Var el, ad)) = [] + | aset (Eq (Neg em, ad)) = [] + | aset (Eq (Add (Cst fc, eo), ad)) = [] + | aset (Eq (Add (Var fd, eo), ad)) = [] + | aset (Eq (Add (Neg fe, eo), ad)) = [] + | aset (Eq (Add (Add (ff, fg), eo), ad)) = [] + | aset (Eq (Add (Sub (fh, fi), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = + (if (he = 0) + then (if (fu < 0) then [lin_add (eo, Cst 1)] + else [lin_add (lin_neg eo, Cst 1)]) + else []) + | aset (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = [] + | aset (Eq (Add (Mult (Var fv, fk), eo), ad)) = [] + | aset (Eq (Add (Mult (Neg fw, fk), eo), ad)) = [] + | aset (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = [] + | aset (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = [] + | aset (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = [] + | aset (Eq (Sub (ep, eq), ad)) = [] + | aset (Eq (Mult (er, es), ad)) = [] + | aset (Divides (ae, af)) = [] + | aset T = [] + | aset F = [] + | aset (NOT (Lt (hg, hh))) = [] + | aset (NOT (Gt (hi, hj))) = [] + | aset (NOT (Le (hk, hl))) = [] + | aset (NOT (Ge (hm, hn))) = [] + | aset (NOT (Eq (Cst ja, hp))) = [] + | aset (NOT (Eq (Var jb, hp))) = [] + | aset (NOT (Eq (Neg jc, hp))) = [] + | aset (NOT (Eq (Add (Cst js, je), hp))) = [] + | aset (NOT (Eq (Add (Var jt, je), hp))) = [] + | aset (NOT (Eq (Add (Neg ju, je), hp))) = [] + | aset (NOT (Eq (Add (Add (jv, jw), je), hp))) = [] + | aset (NOT (Eq (Add (Sub (jx, jy), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = + (if (lu = 0) then (if (kk < 0) then [je] else [lin_neg je]) else []) + | aset (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = [] + | aset (NOT (Eq (Sub (jf, jg), hp))) = [] + | aset (NOT (Eq (Mult (jh, ji), hp))) = [] + | aset (NOT (Divides (hq, hr))) = [] + | aset (NOT T) = [] + | aset (NOT F) = [] + | aset (NOT (NOT hs)) = [] + | aset (NOT (And (ht, hu))) = [] + | aset (NOT (Or (hv, hw))) = [] + | aset (NOT (Imp (hx, hy))) = [] + | aset (NOT (Equ (hz, ia))) = [] + | aset (NOT (QAll ib)) = [] + | aset (NOT (QEx ic)) = [] + | aset (Imp (al, am)) = [] + | aset (Equ (an, ao)) = [] + | aset (QAll ap) = [] + | aset (QEx aq) = []; + +fun op_mem x [] = false + | op_mem x (y :: ys) = (if (y = x) then true else op_mem x ys); + +fun list_insert x xs = (if op_mem x xs then xs else (x :: xs)); + +fun list_set [] = [] + | list_set (x :: xs) = list_insert x (list_set xs); + +fun bset (And (p, q)) = op_64 (bset p) (bset q) + | bset (Or (p, q)) = op_64 (bset p) (bset q) + | bset (Lt (u, v)) = [] + | bset (Gt (w, x)) = [] + | bset (Le (Cst bo, z)) = [] + | bset (Le (Var bp, z)) = [] + | bset (Le (Neg bq, z)) = [] + | bset (Le (Add (Cst cg, bs), z)) = [] + | bset (Le (Add (Var ch, bs), z)) = [] + | bset (Le (Add (Neg ci, bs), z)) = [] + | bset (Le (Add (Add (cj, ck), bs), z)) = [] + | bset (Le (Add (Sub (cl, cm), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Var ei), bs), z)) = + (if (ei = 0) + then (if (cy < 0) then [lin_add (bs, Cst ~1), bs] + else [lin_add (lin_neg bs, Cst ~1)]) + else []) + | bset (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = [] + | bset (Le (Add (Mult (Var cz, co), bs), z)) = [] + | bset (Le (Add (Mult (Neg da, co), bs), z)) = [] + | bset (Le (Add (Mult (Add (db, dc), co), bs), z)) = [] + | bset (Le (Add (Mult (Sub (dd, de), co), bs), z)) = [] + | bset (Le (Add (Mult (Mult (df, dg), co), bs), z)) = [] + | bset (Le (Sub (bt, bu), z)) = [] + | bset (Le (Mult (bv, bw), z)) = [] + | bset (Ge (aa, ab)) = [] + | bset (Eq (Cst ek, ad)) = [] + | bset (Eq (Var el, ad)) = [] + | bset (Eq (Neg em, ad)) = [] + | bset (Eq (Add (Cst fc, eo), ad)) = [] + | bset (Eq (Add (Var fd, eo), ad)) = [] + | bset (Eq (Add (Neg fe, eo), ad)) = [] + | bset (Eq (Add (Add (ff, fg), eo), ad)) = [] + | bset (Eq (Add (Sub (fh, fi), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = + (if (he = 0) + then (if (fu < 0) then [lin_add (eo, Cst ~1)] + else [lin_add (lin_neg eo, Cst ~1)]) + else []) + | bset (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = [] + | bset (Eq (Add (Mult (Var fv, fk), eo), ad)) = [] + | bset (Eq (Add (Mult (Neg fw, fk), eo), ad)) = [] + | bset (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = [] + | bset (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = [] + | bset (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = [] + | bset (Eq (Sub (ep, eq), ad)) = [] + | bset (Eq (Mult (er, es), ad)) = [] + | bset (Divides (ae, af)) = [] + | bset T = [] + | bset F = [] + | bset (NOT (Lt (hg, hh))) = [] + | bset (NOT (Gt (hi, hj))) = [] + | bset (NOT (Le (hk, hl))) = [] + | bset (NOT (Ge (hm, hn))) = [] + | bset (NOT (Eq (Cst ja, hp))) = [] + | bset (NOT (Eq (Var jb, hp))) = [] + | bset (NOT (Eq (Neg jc, hp))) = [] + | bset (NOT (Eq (Add (Cst js, je), hp))) = [] + | bset (NOT (Eq (Add (Var jt, je), hp))) = [] + | bset (NOT (Eq (Add (Neg ju, je), hp))) = [] + | bset (NOT (Eq (Add (Add (jv, jw), je), hp))) = [] + | bset (NOT (Eq (Add (Sub (jx, jy), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = + (if (lu = 0) then (if (kk < 0) then [je] else [lin_neg je]) else []) + | bset (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = [] + | bset (NOT (Eq (Sub (jf, jg), hp))) = [] + | bset (NOT (Eq (Mult (jh, ji), hp))) = [] + | bset (NOT (Divides (hq, hr))) = [] + | bset (NOT T) = [] + | bset (NOT F) = [] + | bset (NOT (NOT hs)) = [] + | bset (NOT (And (ht, hu))) = [] + | bset (NOT (Or (hv, hw))) = [] + | bset (NOT (Imp (hx, hy))) = [] + | bset (NOT (Equ (hz, ia))) = [] + | bset (NOT (QAll ib)) = [] + | bset (NOT (QEx ic)) = [] + | bset (Imp (al, am)) = [] + | bset (Equ (an, ao)) = [] + | bset (QAll ap) = [] + | bset (QEx aq) = []; + +fun adjustcoeff (l, Le (Add (Mult (Cst c, Var 0), r), Cst i)) = + (if (c <= 0) + then Le (Add (Mult (Cst ~1, Var 0), lin_mul (~ (op_div_def1 l c), r)), + Cst 0) + else Le (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0)) + | adjustcoeff (l, Eq (Add (Mult (Cst c, Var 0), r), Cst i)) = + Eq (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst c, Var 0), r), Cst i))) = + NOT (Eq (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0)) + | adjustcoeff (l, And (p, q)) = And (adjustcoeff (l, p), adjustcoeff (l, q)) + | adjustcoeff (l, Or (p, q)) = Or (adjustcoeff (l, p), adjustcoeff (l, q)) + | adjustcoeff (l, Lt (w, x)) = Lt (w, x) + | adjustcoeff (l, Gt (y, z)) = Gt (y, z) + | adjustcoeff (l, Le (Cst bq, ab)) = Le (Cst bq, ab) + | adjustcoeff (l, Le (Var br, ab)) = Le (Var br, ab) + | adjustcoeff (l, Le (Neg bs, ab)) = Le (Neg bs, ab) + | adjustcoeff (l, Le (Add (Cst ci, bu), ab)) = Le (Add (Cst ci, bu), ab) + | adjustcoeff (l, Le (Add (Var cj, bu), ab)) = Le (Add (Var cj, bu), ab) + | adjustcoeff (l, Le (Add (Neg ck, bu), ab)) = Le (Add (Neg ck, bu), ab) + | adjustcoeff (l, Le (Add (Add (cl, cm), bu), ab)) = + Le (Add (Add (cl, cm), bu), ab) + | adjustcoeff (l, Le (Add (Sub (cn, co), bu), ab)) = + Le (Add (Sub (cn, co), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Cst ds), bu), ab)) = + Le (Add (Mult (Cst da, Cst ds), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Var en)) = + Le (Add (Mult (Cst da, Var 0), bu), Var en) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Neg eo)) = + Le (Add (Mult (Cst da, Var 0), bu), Neg eo) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Add (ep, eq))) = + Le (Add (Mult (Cst da, Var 0), bu), Add (ep, eq)) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Sub (er, es))) = + Le (Add (Mult (Cst da, Var 0), bu), Sub (er, es)) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Mult (et, eu))) = + Le (Add (Mult (Cst da, Var 0), bu), Mult (et, eu)) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var ek), bu), ab)) = + Le (Add (Mult (Cst da, Var ek), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Neg du), bu), ab)) = + Le (Add (Mult (Cst da, Neg du), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Add (dv, dw)), bu), ab)) = + Le (Add (Mult (Cst da, Add (dv, dw)), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Sub (dx, dy)), bu), ab)) = + Le (Add (Mult (Cst da, Sub (dx, dy)), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Mult (dz, ea)), bu), ab)) = + Le (Add (Mult (Cst da, Mult (dz, ea)), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Var db, cq), bu), ab)) = + Le (Add (Mult (Var db, cq), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Neg dc, cq), bu), ab)) = + Le (Add (Mult (Neg dc, cq), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Add (dd, de), cq), bu), ab)) = + Le (Add (Mult (Add (dd, de), cq), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Sub (df, dg), cq), bu), ab)) = + Le (Add (Mult (Sub (df, dg), cq), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Mult (dh, di), cq), bu), ab)) = + Le (Add (Mult (Mult (dh, di), cq), bu), ab) + | adjustcoeff (l, Le (Sub (bv, bw), ab)) = Le (Sub (bv, bw), ab) + | adjustcoeff (l, Le (Mult (bx, by), ab)) = Le (Mult (bx, by), ab) + | adjustcoeff (l, Ge (ac, ad)) = Ge (ac, ad) + | adjustcoeff (l, Eq (Cst fe, af)) = Eq (Cst fe, af) + | adjustcoeff (l, Eq (Var ff, af)) = Eq (Var ff, af) + | adjustcoeff (l, Eq (Neg fg, af)) = Eq (Neg fg, af) + | adjustcoeff (l, Eq (Add (Cst fw, fi), af)) = Eq (Add (Cst fw, fi), af) + | adjustcoeff (l, Eq (Add (Var fx, fi), af)) = Eq (Add (Var fx, fi), af) + | adjustcoeff (l, Eq (Add (Neg fy, fi), af)) = Eq (Add (Neg fy, fi), af) + | adjustcoeff (l, Eq (Add (Add (fz, ga), fi), af)) = + Eq (Add (Add (fz, ga), fi), af) + | adjustcoeff (l, Eq (Add (Sub (gb, gc), fi), af)) = + Eq (Add (Sub (gb, gc), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Cst hg), fi), af)) = + Eq (Add (Mult (Cst go, Cst hg), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Var ib)) = + Eq (Add (Mult (Cst go, Var 0), fi), Var ib) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Neg ic)) = + Eq (Add (Mult (Cst go, Var 0), fi), Neg ic) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Add (id, ie))) = + Eq (Add (Mult (Cst go, Var 0), fi), Add (id, ie)) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Sub (if', ig))) = + Eq (Add (Mult (Cst go, Var 0), fi), Sub (if', ig)) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Mult (ih, ii))) = + Eq (Add (Mult (Cst go, Var 0), fi), Mult (ih, ii)) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var hy), fi), af)) = + Eq (Add (Mult (Cst go, Var hy), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Neg hi), fi), af)) = + Eq (Add (Mult (Cst go, Neg hi), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Add (hj, hk)), fi), af)) = + Eq (Add (Mult (Cst go, Add (hj, hk)), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Sub (hl, hm)), fi), af)) = + Eq (Add (Mult (Cst go, Sub (hl, hm)), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Mult (hn, ho)), fi), af)) = + Eq (Add (Mult (Cst go, Mult (hn, ho)), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Var gp, ge), fi), af)) = + Eq (Add (Mult (Var gp, ge), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Neg gq, ge), fi), af)) = + Eq (Add (Mult (Neg gq, ge), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Add (gr, gs), ge), fi), af)) = + Eq (Add (Mult (Add (gr, gs), ge), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Sub (gt, gu), ge), fi), af)) = + Eq (Add (Mult (Sub (gt, gu), ge), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Mult (gv, gw), ge), fi), af)) = + Eq (Add (Mult (Mult (gv, gw), ge), fi), af) + | adjustcoeff (l, Eq (Sub (fj, fk), af)) = Eq (Sub (fj, fk), af) + | adjustcoeff (l, Eq (Mult (fl, fm), af)) = Eq (Mult (fl, fm), af) + | adjustcoeff (l, Divides (Cst is, Cst jk)) = Divides (Cst is, Cst jk) + | adjustcoeff (l, Divides (Cst is, Var jl)) = Divides (Cst is, Var jl) + | adjustcoeff (l, Divides (Cst is, Neg jm)) = Divides (Cst is, Neg jm) + | adjustcoeff (l, Divides (Cst is, Add (Cst kc, jo))) = + Divides (Cst is, Add (Cst kc, jo)) + | adjustcoeff (l, Divides (Cst is, Add (Var kd, jo))) = + Divides (Cst is, Add (Var kd, jo)) + | adjustcoeff (l, Divides (Cst is, Add (Neg ke, jo))) = + Divides (Cst is, Add (Neg ke, jo)) + | adjustcoeff (l, Divides (Cst is, Add (Add (kf, kg), jo))) = + Divides (Cst is, Add (Add (kf, kg), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Sub (kh, ki), jo))) = + Divides (Cst is, Add (Sub (kh, ki), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Cst lm), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Cst lm), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Var me), jo))) = + (if (me = 0) + then Divides + (Cst (op_div_def1 l ku * is), + Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l ku, jo))) + else Divides + (Cst is, + Add (Mult (Cst ku, Var (op_45_def0 me id_1_def0 + 1)), jo))) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Neg lo), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Neg lo), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Add (lp, lq)), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Add (lp, lq)), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Sub (lr, ls)), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Sub (lr, ls)), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Mult (lt, lu)), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Mult (lt, lu)), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Var kv, kk), jo))) = + Divides (Cst is, Add (Mult (Var kv, kk), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Neg kw, kk), jo))) = + Divides (Cst is, Add (Mult (Neg kw, kk), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Add (kx, ky), kk), jo))) = + Divides (Cst is, Add (Mult (Add (kx, ky), kk), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Sub (kz, la), kk), jo))) = + Divides (Cst is, Add (Mult (Sub (kz, la), kk), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Mult (lb, lc), kk), jo))) = + Divides (Cst is, Add (Mult (Mult (lb, lc), kk), jo)) + | adjustcoeff (l, Divides (Cst is, Sub (jp, jq))) = + Divides (Cst is, Sub (jp, jq)) + | adjustcoeff (l, Divides (Cst is, Mult (jr, js))) = + Divides (Cst is, Mult (jr, js)) + | adjustcoeff (l, Divides (Var it, ah)) = Divides (Var it, ah) + | adjustcoeff (l, Divides (Neg iu, ah)) = Divides (Neg iu, ah) + | adjustcoeff (l, Divides (Add (iv, iw), ah)) = Divides (Add (iv, iw), ah) + | adjustcoeff (l, Divides (Sub (ix, iy), ah)) = Divides (Sub (ix, iy), ah) + | adjustcoeff (l, Divides (Mult (iz, ja), ah)) = Divides (Mult (iz, ja), ah) + | adjustcoeff (l, T) = T + | adjustcoeff (l, F) = F + | adjustcoeff (l, NOT (Lt (mg, mh))) = NOT (Lt (mg, mh)) + | adjustcoeff (l, NOT (Gt (mi, mj))) = NOT (Gt (mi, mj)) + | adjustcoeff (l, NOT (Le (mk, ml))) = NOT (Le (mk, ml)) + | adjustcoeff (l, NOT (Ge (mm, mn))) = NOT (Ge (mm, mn)) + | adjustcoeff (l, NOT (Eq (Cst oa, mp))) = NOT (Eq (Cst oa, mp)) + | adjustcoeff (l, NOT (Eq (Var ob, mp))) = NOT (Eq (Var ob, mp)) + | adjustcoeff (l, NOT (Eq (Neg oc, mp))) = NOT (Eq (Neg oc, mp)) + | adjustcoeff (l, NOT (Eq (Add (Cst os, oe), mp))) = + NOT (Eq (Add (Cst os, oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Var ot, oe), mp))) = + NOT (Eq (Add (Var ot, oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Neg ou, oe), mp))) = + NOT (Eq (Add (Neg ou, oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Add (ov, ow), oe), mp))) = + NOT (Eq (Add (Add (ov, ow), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Sub (ox, oy), oe), mp))) = + NOT (Eq (Add (Sub (ox, oy), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Cst qc), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Cst qc), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Var qx))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Var qx)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Neg qy))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Neg qy)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Add (qz, ra)))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Add (qz, ra))) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Sub (rb, rc)))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Sub (rb, rc))) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Mult (rd, re)))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Mult (rd, re))) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var qu), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Var qu), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Neg qe), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Neg qe), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Add (qf, qg)), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Add (qf, qg)), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Sub (qh, qi)), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Sub (qh, qi)), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Mult (qj, qk)), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Mult (qj, qk)), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Var pl, pa), oe), mp))) = + NOT (Eq (Add (Mult (Var pl, pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Neg pm, pa), oe), mp))) = + NOT (Eq (Add (Mult (Neg pm, pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Add (pn, po), pa), oe), mp))) = + NOT (Eq (Add (Mult (Add (pn, po), pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Sub (pp, pq), pa), oe), mp))) = + NOT (Eq (Add (Mult (Sub (pp, pq), pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Mult (pr, ps), pa), oe), mp))) = + NOT (Eq (Add (Mult (Mult (pr, ps), pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Sub (of', og), mp))) = NOT (Eq (Sub (of', og), mp)) + | adjustcoeff (l, NOT (Eq (Mult (oh, oi), mp))) = NOT (Eq (Mult (oh, oi), mp)) + | adjustcoeff (l, NOT (Divides (Cst ro, Cst sg))) = + NOT (Divides (Cst ro, Cst sg)) + | adjustcoeff (l, NOT (Divides (Cst ro, Var sh))) = + NOT (Divides (Cst ro, Var sh)) + | adjustcoeff (l, NOT (Divides (Cst ro, Neg si))) = + NOT (Divides (Cst ro, Neg si)) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Cst sy, sk)))) = + NOT (Divides (Cst ro, Add (Cst sy, sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Var sz, sk)))) = + NOT (Divides (Cst ro, Add (Var sz, sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Neg ta, sk)))) = + NOT (Divides (Cst ro, Add (Neg ta, sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Add (tb, tc), sk)))) = + NOT (Divides (Cst ro, Add (Add (tb, tc), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Sub (td, te), sk)))) = + NOT (Divides (Cst ro, Add (Sub (td, te), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Cst ui), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Cst ui), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Var va), sk)))) = + (if (va = 0) + then NOT (Divides + (Cst (op_div_def1 l tq * ro), + Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l tq, sk)))) + else NOT (Divides + (Cst ro, + Add (Mult (Cst tq, Var (op_45_def0 va id_1_def0 + 1)), + sk)))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Neg uk), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Neg uk), sk))) + | adjustcoeff + (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Add (ul, um)), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Add (ul, um)), sk))) + | adjustcoeff + (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Sub (un, uo)), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Sub (un, uo)), sk))) + | adjustcoeff + (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Mult (up, uq)), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Mult (up, uq)), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Var tr, tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Var tr, tg), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Neg ts, tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Neg ts, tg), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Add (tt, tu), tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Add (tt, tu), tg), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Sub (tv, tw), tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Sub (tv, tw), tg), sk))) + | adjustcoeff + (l, NOT (Divides (Cst ro, Add (Mult (Mult (tx, ty), tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Mult (tx, ty), tg), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Sub (sl, sm)))) = + NOT (Divides (Cst ro, Sub (sl, sm))) + | adjustcoeff (l, NOT (Divides (Cst ro, Mult (sn, so)))) = + NOT (Divides (Cst ro, Mult (sn, so))) + | adjustcoeff (l, NOT (Divides (Var rp, mr))) = NOT (Divides (Var rp, mr)) + | adjustcoeff (l, NOT (Divides (Neg rq, mr))) = NOT (Divides (Neg rq, mr)) + | adjustcoeff (l, NOT (Divides (Add (rr, rs), mr))) = + NOT (Divides (Add (rr, rs), mr)) + | adjustcoeff (l, NOT (Divides (Sub (rt, ru), mr))) = + NOT (Divides (Sub (rt, ru), mr)) + | adjustcoeff (l, NOT (Divides (Mult (rv, rw), mr))) = + NOT (Divides (Mult (rv, rw), mr)) + | adjustcoeff (l, NOT T) = NOT T + | adjustcoeff (l, NOT F) = NOT F + | adjustcoeff (l, NOT (NOT ms)) = NOT (NOT ms) + | adjustcoeff (l, NOT (And (mt, mu))) = NOT (And (mt, mu)) + | adjustcoeff (l, NOT (Or (mv, mw))) = NOT (Or (mv, mw)) + | adjustcoeff (l, NOT (Imp (mx, my))) = NOT (Imp (mx, my)) + | adjustcoeff (l, NOT (Equ (mz, na))) = NOT (Equ (mz, na)) + | adjustcoeff (l, NOT (QAll nb)) = NOT (QAll nb) + | adjustcoeff (l, NOT (QEx nc)) = NOT (QEx nc) + | adjustcoeff (l, Imp (an, ao)) = Imp (an, ao) + | adjustcoeff (l, Equ (ap, aq)) = Equ (ap, aq) + | adjustcoeff (l, QAll ar) = QAll ar + | adjustcoeff (l, QEx as') = QEx as'; + +fun formlcm (Le (Add (Mult (Cst c, Var 0), r), Cst i)) = abs c + | formlcm (Eq (Add (Mult (Cst c, Var 0), r), Cst i)) = abs c + | formlcm (NOT p) = formlcm p + | formlcm (And (p, q)) = ilcm (formlcm p) (formlcm q) + | formlcm (Or (p, q)) = ilcm (formlcm p) (formlcm q) + | formlcm (Lt (u, v)) = 1 + | formlcm (Gt (w, x)) = 1 + | formlcm (Le (Cst bo, z)) = 1 + | formlcm (Le (Var bp, z)) = 1 + | formlcm (Le (Neg bq, z)) = 1 + | formlcm (Le (Add (Cst cg, bs), z)) = 1 + | formlcm (Le (Add (Var ch, bs), z)) = 1 + | formlcm (Le (Add (Neg ci, bs), z)) = 1 + | formlcm (Le (Add (Add (cj, ck), bs), z)) = 1 + | formlcm (Le (Add (Sub (cl, cm), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Var el)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Neg em)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Add (en, eo))) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Sub (ep, eq))) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Mult (er, es))) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var ei ), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = 1 + | formlcm (Le (Add (Mult (Var cz, co), bs), z)) = 1 + | formlcm (Le (Add (Mult (Neg da, co), bs), z)) = 1 + | formlcm (Le (Add (Mult (Add (db, dc), co), bs), z)) = 1 + | formlcm (Le (Add (Mult (Sub (dd, de), co), bs), z)) = 1 + | formlcm (Le (Add (Mult (Mult (df, dg), co), bs), z)) = 1 + | formlcm (Le (Sub (bt, bu), z)) = 1 + | formlcm (Le (Mult (bv, bw), z)) = 1 + | formlcm (Ge (aa, ab)) = 1 + | formlcm (Eq (Cst fc, ad)) = 1 + | formlcm (Eq (Var fd, ad)) = 1 + | formlcm (Eq (Neg fe, ad)) = 1 + | formlcm (Eq (Add (Cst fu, fg), ad)) = 1 + | formlcm (Eq (Add (Var fv, fg), ad)) = 1 + | formlcm (Eq (Add (Neg fw, fg), ad)) = 1 + | formlcm (Eq (Add (Add (fx, fy), fg), ad)) = 1 + | formlcm (Eq (Add (Sub (fz, ga), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Cst he), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Var hz)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Neg ia)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Add (ib, ic))) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Sub (id, ie))) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Mult (if', ig))) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var hw), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Neg hg), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Add (hh, hi)), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Sub (hj, hk)), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Mult (hl, hm)), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Var gn, gc), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Neg go, gc), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Add (gp, gq), gc), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Sub (gr, gs), gc), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Mult (gt, gu), gc), fg), ad)) = 1 + | formlcm (Eq (Sub (fh, fi), ad)) = 1 + | formlcm (Eq (Mult (fj, fk), ad)) = 1 + | formlcm (Divides (Cst iq, Cst ji)) = 1 + | formlcm (Divides (Cst iq, Var jj)) = 1 + | formlcm (Divides (Cst iq, Neg jk)) = 1 + | formlcm (Divides (Cst iq, Add (Cst ka, jm))) = 1 + | formlcm (Divides (Cst iq, Add (Var kb, jm))) = 1 + | formlcm (Divides (Cst iq, Add (Neg kc, jm))) = 1 + | formlcm (Divides (Cst iq, Add (Add (kd, ke), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Sub (kf, kg), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Cst lk), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Var mc), jm))) = + (if (mc = 0) then abs ks else 1) + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Neg lm), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Add (ln, lo)), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Sub (lp, lq)), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Mult (lr, ls)), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Var kt, ki), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Neg ku, ki), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Add (kv, kw), ki), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Sub (kx, ky), ki), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Mult (kz, la), ki), jm))) = 1 + | formlcm (Divides (Cst iq, Sub (jn, jo))) = 1 + | formlcm (Divides (Cst iq, Mult (jp, jq))) = 1 + | formlcm (Divides (Var ir, af)) = 1 + | formlcm (Divides (Neg is, af)) = 1 + | formlcm (Divides (Add (it, iu), af)) = 1 + | formlcm (Divides (Sub (iv, iw), af)) = 1 + | formlcm (Divides (Mult (ix, iy), af)) = 1 + | formlcm T = 1 + | formlcm F = 1 + | formlcm (Imp (al, am)) = 1 + | formlcm (Equ (an, ao)) = 1 + | formlcm (QAll ap) = 1 + | formlcm (QEx aq) = 1; + +fun unitycoeff p = + let val l = formlcm p; val p' = adjustcoeff (l, p) + in (if (l = 1) then p' + else And (Divides (Cst l, Add (Mult (Cst 1, Var 0), Cst 0)), p')) + end; + +fun unify p = + let val q = unitycoeff p; val B = list_set (bset q); val A = list_set (aset q) + in (if op_60_61_def0 (size_def1 B) (size_def1 A) then (q, B) + else (mirror q, map lin_neg A)) + end; + +fun cooper p = + lift_un (fn q => decrvars (explode_minf (unify q))) (linform (nnf p)); + +fun pa p = lift_un psimpl (qelim (cooper, p)); + +val test = pa; + +end; diff -r afaa031ed4da -r 105519771c67 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Wed Sep 14 10:24:39 2005 +0200 +++ b/src/HOL/IsaMakefile Wed Sep 14 17:25:52 2005 +0200 @@ -78,13 +78,15 @@ $(SRC)/TFL/casesplit.ML $(SRC)/TFL/dcterm.ML $(SRC)/TFL/post.ML \ $(SRC)/TFL/rules.ML $(SRC)/TFL/tfl.ML $(SRC)/TFL/thms.ML \ $(SRC)/TFL/thry.ML $(SRC)/TFL/usyntax.ML $(SRC)/TFL/utils.ML \ - Binomial.thy Datatype.ML Datatype.thy Datatype_Universe.thy Divides.thy \ + Binomial.thy Commutative_Ring.thy Datatype.ML Datatype.thy \ + Datatype_Universe.thy Divides.thy \ Equiv_Relations.thy Extraction.thy Finite_Set.ML Finite_Set.thy \ FixedPoint.thy Fun.thy HOL.ML HOL.thy Hilbert_Choice.thy Inductive.thy \ Infinite_Set.thy Integ/IntArith.thy Integ/IntDef.thy Integ/IntDiv.thy \ Integ/NatBin.thy Integ/NatSimprocs.thy Integ/Numeral.thy \ Integ/Parity.thy Integ/Presburger.thy Integ/cooper_dec.ML \ - Integ/cooper_proof.ML Integ/int_arith1.ML Integ/int_factor_simprocs.ML \ + Integ/cooper_proof.ML Integ/reflected_presburger.ML \ + Integ/reflected_cooper.ML Integ/int_arith1.ML Integ/int_factor_simprocs.ML \ Integ/nat_simprocs.ML Integ/presburger.ML Integ/qelim.ML LOrder.thy \ Lattice_Locales.thy List.ML List.thy Main.ML Main.thy Map.thy \ Nat.ML Nat.thy NatArith.thy OrderedGroup.ML OrderedGroup.thy \ @@ -96,7 +98,7 @@ Tools/ATP/recon_order_clauses.ML Tools/ATP/recon_parse.ML \ Tools/ATP/recon_prelim.ML Tools/ATP/recon_transfer_proof.ML \ Tools/ATP/recon_translate_proof.ML Tools/ATP/res_clasimpset.ML \ - Tools/ATP/watcher.ML \ + Tools/ATP/watcher.ML Tools/comm_ring.ML \ Tools/datatype_abs_proofs.ML Tools/datatype_aux.ML \ Tools/datatype_codegen.ML Tools/datatype_package.ML \ Tools/datatype_prop.ML Tools/datatype_realizer.ML \ @@ -584,13 +586,15 @@ HOL-ex: HOL $(LOG)/HOL-ex.gz $(LOG)/HOL-ex.gz: $(OUT)/HOL ex/Antiquote.thy \ - ex/BT.thy ex/BinEx.thy ex/Higher_Order_Logic.thy \ + ex/BT.thy ex/BinEx.thy ex/Commutative_RingEx.thy \ + ex/Commutative_Ring_Complete.thy ex/Higher_Order_Logic.thy \ ex/Hilbert_Classical.thy ex/InSort.thy \ ex/InductiveInvariant.thy ex/InductiveInvariant_examples.thy\ ex/Intuitionistic.thy \ ex/Lagrange.thy ex/Locales.thy ex/MergeSort.thy \ ex/MT.ML ex/MT.thy ex/MonoidGroup.thy ex/Multiquote.thy \ - ex/NatSum.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy ex/Puzzle.thy \ + ex/NatSum.thy ex/PER.thy ex/PresburgerEx.thy ex/Reflected_Presburger.thy \ + ex/Primrec.thy ex/Puzzle.thy \ ex/Qsort.thy ex/Quickcheck_Examples.thy ex/ROOT.ML ex/Recdefs.thy ex/Records.thy \ ex/Refute_Examples.thy \ ex/StringEx.thy ex/SVC_Oracle.ML ex/SVC_Oracle.thy \ diff -r afaa031ed4da -r 105519771c67 src/HOL/Presburger.thy --- a/src/HOL/Presburger.thy Wed Sep 14 10:24:39 2005 +0200 +++ b/src/HOL/Presburger.thy Wed Sep 14 17:25:52 2005 +0200 @@ -10,7 +10,8 @@ theory Presburger imports NatSimprocs SetInterval -uses ("cooper_dec.ML") ("cooper_proof.ML") ("qelim.ML") ("presburger.ML") +uses ("cooper_dec.ML") ("cooper_proof.ML") ("qelim.ML") + ("reflected_presburger.ML") ("reflected_cooper.ML") ("presburger.ML") begin text {* Theorem for unitifying the coeffitients of @{text x} in an existential formula*} @@ -982,8 +983,10 @@ by (simp cong: conj_cong) use "cooper_dec.ML" +use "reflected_presburger.ML" +use "reflected_cooper.ML" oracle - presburger_oracle ("term") = CooperDec.presburger_oracle + presburger_oracle ("term") = ReflectedCooper.presburger_oracle use "cooper_proof.ML" use "qelim.ML" diff -r afaa031ed4da -r 105519771c67 src/HOL/Tools/Presburger/cooper_dec.ML --- a/src/HOL/Tools/Presburger/cooper_dec.ML Wed Sep 14 10:24:39 2005 +0200 +++ b/src/HOL/Tools/Presburger/cooper_dec.ML Wed Sep 14 17:25:52 2005 +0200 @@ -44,7 +44,6 @@ val evalc : term -> term val cooper_w : string list -> term -> (term option * term) val integer_qelim : Term.term -> Term.term - val presburger_oracle : theory -> term -> term end; structure CooperDec : COOPER_DEC = @@ -938,8 +937,4 @@ val integer_qelim = simpl o evalc o (lift_qelim linform (cnnf posineq o evalc) cooper is_arith_rel) ; -fun presburger_oracle thy t = - if (!quick_and_dirty) - then HOLogic.mk_Trueprop (HOLogic.mk_eq(t,integer_qelim t)) - else error "Presburger oracle: not in quick_and_dirty mode" end; diff -r afaa031ed4da -r 105519771c67 src/HOL/Tools/Presburger/reflected_cooper.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Presburger/reflected_cooper.ML Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,124 @@ +(* $Id$ *) +(* The oracle for Presburger arithmetic based on the verified Code *) + (* in HOL/ex/Reflected_Presburger.thy *) + +structure ReflectedCooper = +struct + +open Generated; +(* pseudo reification : term -> intterm *) + +fun i_of_term vs t = + case t of + Free(xn,xT) => (case assoc(vs,t) of + NONE => error "Variable not found in the list!!" + | SOME n => Var n) + | Const("0",iT) => Cst 0 + | Const("1",iT) => Cst 1 + | Bound i => Var (nat i) + | Const("uminus",_)$t' => Neg (i_of_term vs t') + | Const ("op +",_)$t1$t2 => Add (i_of_term vs t1,i_of_term vs t2) + | Const ("op -",_)$t1$t2 => Sub (i_of_term vs t1,i_of_term vs t2) + | Const ("op *",_)$t1$t2 => Mult (i_of_term vs t1,i_of_term vs t2) + | Const ("Numeral.number_of",_)$t' => Cst (HOLogic.dest_binum t') + | _ => error "i_of_term: unknown term"; + + +(* pseudo reification : term -> QF *) +fun qf_of_term vs t = + case t of + Const("True",_) => T + | Const("False",_) => F + | Const("op <",_)$t1$t2 => Lt (i_of_term vs t1,i_of_term vs t2) + | Const("op <=",_)$t1$t2 => Le (i_of_term vs t1,i_of_term vs t2) + | Const ("Divides.op dvd",_)$t1$t2 => + Divides(i_of_term vs t1,i_of_term vs t2) + | Const("op =",eqT)$t1$t2 => + if (domain_type eqT = HOLogic.intT) + then let val i1 = i_of_term vs t1 + val i2 = i_of_term vs t2 + in Eq (i1,i2) + end + else Equ(qf_of_term vs t1,qf_of_term vs t2) + | Const("op &",_)$t1$t2 => And(qf_of_term vs t1,qf_of_term vs t2) + | Const("op |",_)$t1$t2 => Or(qf_of_term vs t1,qf_of_term vs t2) + | Const("op -->",_)$t1$t2 => Imp(qf_of_term vs t1,qf_of_term vs t2) + | Const("Not",_)$t' => NOT(qf_of_term vs t') + | Const("Ex",_)$Abs(xn,xT,p) => + QEx(qf_of_term (map (fn(v,n) => (v,n + 1)) vs) p) + | Const("All",_)$Abs(xn,xT,p) => + QAll(qf_of_term (map (fn(v,n) => (v,n + 1)) vs) p) + | _ => error "qf_of_term : unknown term!"; + +(* +fun parse s = term_of (read_cterm (sign_of Main.thy) (s,HOLogic.boolT)); + +val t = parse "ALL (i::int) (j::int). i < 8* j --> (i - 1 = j + 3 + 2*j) & (j <= -i + k ) | 4 = i | 5 dvd i"; +*) +fun zip [] [] = [] + | zip (x::xs) (y::ys) = (x,y)::(zip xs ys); + + +fun start_vs t = + let val fs = term_frees t + in zip fs (map nat (0 upto (length fs - 1))) + end ; + +(* transform intterm and QF back to terms *) +val iT = HOLogic.intT; +val bT = HOLogic.boolT; +fun myassoc2 l v = + case l of + [] => NONE + | (x,v')::xs => if v = v' then SOME x + else myassoc2 xs v; + +fun term_of_i vs t = + case t of + Cst i => CooperDec.mk_numeral i + | Var n => valOf (myassoc2 vs n) + | Neg t' => Const("uminus",iT --> iT)$(term_of_i vs t') + | Add(t1,t2) => Const("op +",[iT,iT] ---> iT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Sub(t1,t2) => Const("op -",[iT,iT] ---> iT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Mult(t1,t2) => Const("op *",[iT,iT] ---> iT)$ + (term_of_i vs t1)$(term_of_i vs t2); + +fun term_of_qf vs t = + case t of + T => HOLogic.true_const + | F => HOLogic.false_const + | Lt(t1,t2) => Const("op <",[iT,iT] ---> bT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Le(t1,t2) => Const("op <=",[iT,iT] ---> bT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Gt(t1,t2) => Const("op <",[iT,iT] ---> bT)$ + (term_of_i vs t2)$(term_of_i vs t1) + | Ge(t1,t2) => Const("op <=",[iT,iT] ---> bT)$ + (term_of_i vs t2)$(term_of_i vs t1) + | Eq(t1,t2) => Const("op =",[iT,iT] ---> bT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | Divides(t1,t2) => Const("Divides.op dvd",[iT,iT] ---> bT)$ + (term_of_i vs t1)$(term_of_i vs t2) + | NOT t' => HOLogic.Not$(term_of_qf vs t') + | And(t1,t2) => HOLogic.conj$(term_of_qf vs t1)$(term_of_qf vs t2) + | Or(t1,t2) => HOLogic.disj$(term_of_qf vs t1)$(term_of_qf vs t2) + | Imp(t1,t2) => HOLogic.imp$(term_of_qf vs t1)$(term_of_qf vs t2) + | Equ(t1,t2) => (HOLogic.eq_const bT)$(term_of_qf vs t1)$ + (term_of_qf vs t2) + | _ => error "If this is raised, Isabelle/HOL or generate_code is inconsistent!"; + +(* The oracle *) + exception COOPER; + +fun presburger_oracle thy t = + let val vs = start_vs t + val result = lift_un (term_of_qf vs) (pa (qf_of_term vs t)) + in + case result of + None => raise COOPER + | Some t' => HOLogic.mk_Trueprop (HOLogic.mk_eq(t,t')) + end ; + +end; \ No newline at end of file diff -r afaa031ed4da -r 105519771c67 src/HOL/Tools/Presburger/reflected_presburger.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Presburger/reflected_presburger.ML Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,2172 @@ +(* $Id$ *) + + (* Caution: This file should not be modified. *) + (* It is autmatically generated from HOL/ex/Reflected_Presburger.thy *) +fun nat i = if i < 0 then 0 else i; +structure Generated = +struct + +datatype intterm = Cst of int | Var of int | Neg of intterm + | Add of intterm * intterm | Sub of intterm * intterm + | Mult of intterm * intterm; + +datatype QF = Lt of intterm * intterm | Gt of intterm * intterm + | Le of intterm * intterm | Ge of intterm * intterm | Eq of intterm * intterm + | Divides of intterm * intterm | T | F | NOT of QF | And of QF * QF + | Or of QF * QF | Imp of QF * QF | Equ of QF * QF | QAll of QF | QEx of QF; + +datatype 'a option = None | Some of 'a; + +fun lift_un c None = None + | lift_un c (Some p) = Some (c p); + +fun lift_bin (c, (Some a, Some b)) = Some (c a b) + | lift_bin (c, (None, y)) = None + | lift_bin (c, (Some y, None)) = None; + +fun lift_qe qe None = None + | lift_qe qe (Some p) = qe p; + +fun qelim (qe, QAll p) = lift_un NOT (lift_qe qe (lift_un NOT (qelim (qe, p)))) + | qelim (qe, QEx p) = lift_qe qe (qelim (qe, p)) + | qelim (qe, And (p, q)) = + lift_bin ((fn x => fn xa => And (x, xa)), (qelim (qe, p), qelim (qe, q))) + | qelim (qe, Or (p, q)) = + lift_bin ((fn x => fn xa => Or (x, xa)), (qelim (qe, p), qelim (qe, q))) + | qelim (qe, Imp (p, q)) = + lift_bin ((fn x => fn xa => Imp (x, xa)), (qelim (qe, p), qelim (qe, q))) + | qelim (qe, Equ (p, q)) = + lift_bin ((fn x => fn xa => Equ (x, xa)), (qelim (qe, p), qelim (qe, q))) + | qelim (qe, NOT p) = lift_un NOT (qelim (qe, p)) + | qelim (qe, Lt (w, x)) = Some (Lt (w, x)) + | qelim (qe, Gt (y, z)) = Some (Gt (y, z)) + | qelim (qe, Le (aa, ab)) = Some (Le (aa, ab)) + | qelim (qe, Ge (ac, ad)) = Some (Ge (ac, ad)) + | qelim (qe, Eq (ae, af)) = Some (Eq (ae, af)) + | qelim (qe, Divides (ag, ah)) = Some (Divides (ag, ah)) + | qelim (qe, T) = Some T + | qelim (qe, F) = Some F; + +fun lin_mul (c, Cst i) = Cst (c * i) + | lin_mul (c, Add (Mult (Cst c', Var n), r)) = + (if (c = 0) then Cst 0 + else Add (Mult (Cst (c * c'), Var n), lin_mul (c, r))); + +fun op_60_def0 m n = ((m) < (n)); + +fun op_60_61_def0 m n = not (op_60_def0 n m); + +fun lin_add (Add (Mult (Cst c1, Var n1), r1), Add (Mult (Cst c2, Var n2), r2)) = + (if (n1 = n2) + then let val c = Cst (c1 + c2) + in (if ((c1 + c2) = 0) then lin_add (r1, r2) + else Add (Mult (c, Var n1), lin_add (r1, r2))) + end + else (if op_60_61_def0 n1 n2 + then Add (Mult (Cst c1, Var n1), + lin_add (r1, Add (Mult (Cst c2, Var n2), r2))) + else Add (Mult (Cst c2, Var n2), + lin_add (Add (Mult (Cst c1, Var n1), r1), r2)))) + | lin_add (Add (Mult (Cst c1, Var n1), r1), Cst b) = + Add (Mult (Cst c1, Var n1), lin_add (r1, Cst b)) + | lin_add (Cst x, Add (Mult (Cst c2, Var n2), r2)) = + Add (Mult (Cst c2, Var n2), lin_add (Cst x, r2)) + | lin_add (Cst b1, Cst b2) = Cst (b1 + b2); + +fun lin_neg i = lin_mul (~1, i); + +fun linearize (Cst b) = Some (Cst b) + | linearize (Var n) = Some (Add (Mult (Cst 1, Var n), Cst 0)) + | linearize (Neg i) = lift_un lin_neg (linearize i) + | linearize (Add (i, j)) = + lift_bin ((fn x => fn y => lin_add (x, y)), (linearize i, linearize j)) + | linearize (Sub (i, j)) = + lift_bin + ((fn x => fn y => lin_add (x, lin_neg y)), (linearize i, linearize j)) + | linearize (Mult (i, j)) = + (case linearize i of None => None + | Some x => + (case x of + Cst xa => + (case linearize j of None => None + | Some x => Some (lin_mul (xa, x))) + | Var xa => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | Neg xa => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | Add (xa, xb) => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | Sub (xa, xb) => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | Mult (xa, xb) => + (case linearize j of None => None + | Some xa => + (case xa of Cst xa => Some (lin_mul (xa, x)) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)))); + +fun linform (Le (it1, it2)) = + lift_bin + ((fn x => fn y => Le (lin_add (x, lin_neg y), Cst 0)), + (linearize it1, linearize it2)) + | linform (Eq (it1, it2)) = + lift_bin + ((fn x => fn y => Eq (lin_add (x, lin_neg y), Cst 0)), + (linearize it1, linearize it2)) + | linform (Divides (d, t)) = + (case linearize d of None => None + | Some x => + (case x of + Cst xa => + (if (xa = 0) then None + else (case linearize t of None => None + | Some xa => Some (Divides (x, xa)))) + | Var xa => None | Neg xa => None | Add (xa, xb) => None + | Sub (xa, xb) => None | Mult (xa, xb) => None)) + | linform T = Some T + | linform F = Some F + | linform (NOT p) = lift_un NOT (linform p) + | linform (And (p, q)) = + lift_bin ((fn f => fn g => And (f, g)), (linform p, linform q)) + | linform (Or (p, q)) = + lift_bin ((fn f => fn g => Or (f, g)), (linform p, linform q)); + +fun nnf (Lt (it1, it2)) = Le (Sub (it1, it2), Cst (~ 1)) + | nnf (Gt (it1, it2)) = Le (Sub (it2, it1), Cst (~ 1)) + | nnf (Le (it1, it2)) = Le (it1, it2) + | nnf (Ge (it1, it2)) = Le (it2, it1) + | nnf (Eq (it1, it2)) = Eq (it2, it1) + | nnf (Divides (d, t)) = Divides (d, t) + | nnf T = T + | nnf F = F + | nnf (And (p, q)) = And (nnf p, nnf q) + | nnf (Or (p, q)) = Or (nnf p, nnf q) + | nnf (Imp (p, q)) = Or (nnf (NOT p), nnf q) + | nnf (Equ (p, q)) = Or (And (nnf p, nnf q), And (nnf (NOT p), nnf (NOT q))) + | nnf (NOT (Lt (it1, it2))) = Le (it2, it1) + | nnf (NOT (Gt (it1, it2))) = Le (it1, it2) + | nnf (NOT (Le (it1, it2))) = Le (Sub (it2, it1), Cst (~ 1)) + | nnf (NOT (Ge (it1, it2))) = Le (Sub (it1, it2), Cst (~ 1)) + | nnf (NOT (Eq (it1, it2))) = NOT (Eq (it1, it2)) + | nnf (NOT (Divides (d, t))) = NOT (Divides (d, t)) + | nnf (NOT T) = F + | nnf (NOT F) = T + | nnf (NOT (NOT p)) = nnf p + | nnf (NOT (And (p, q))) = Or (nnf (NOT p), nnf (NOT q)) + | nnf (NOT (Or (p, q))) = And (nnf (NOT p), nnf (NOT q)) + | nnf (NOT (Imp (p, q))) = And (nnf p, nnf (NOT q)) + | nnf (NOT (Equ (p, q))) = + Or (And (nnf p, nnf (NOT q)), And (nnf (NOT p), nnf q)); + +fun op_45_def2 z w = (z + ~ w); + +fun op_45_def0 m n = nat (op_45_def2 (m) (n)); + +val id_1_def0 : int = (0 + 1); + +fun decrvarsI (Cst i) = Cst i + | decrvarsI (Var n) = Var (op_45_def0 n id_1_def0) + | decrvarsI (Neg a) = Neg (decrvarsI a) + | decrvarsI (Add (a, b)) = Add (decrvarsI a, decrvarsI b) + | decrvarsI (Sub (a, b)) = Sub (decrvarsI a, decrvarsI b) + | decrvarsI (Mult (a, b)) = Mult (decrvarsI a, decrvarsI b); + +fun decrvars (Lt (a, b)) = Lt (decrvarsI a, decrvarsI b) + | decrvars (Gt (a, b)) = Gt (decrvarsI a, decrvarsI b) + | decrvars (Le (a, b)) = Le (decrvarsI a, decrvarsI b) + | decrvars (Ge (a, b)) = Ge (decrvarsI a, decrvarsI b) + | decrvars (Eq (a, b)) = Eq (decrvarsI a, decrvarsI b) + | decrvars (Divides (a, b)) = Divides (decrvarsI a, decrvarsI b) + | decrvars T = T + | decrvars F = F + | decrvars (NOT p) = NOT (decrvars p) + | decrvars (And (p, q)) = And (decrvars p, decrvars q) + | decrvars (Or (p, q)) = Or (decrvars p, decrvars q) + | decrvars (Imp (p, q)) = Imp (decrvars p, decrvars q) + | decrvars (Equ (p, q)) = Equ (decrvars p, decrvars q); + +fun op_64 [] ys = ys + | op_64 (x :: xs) ys = (x :: op_64 xs ys); + +fun map f [] = [] + | map f (x :: xs) = (f x :: map f xs); + +fun iupto (i, j) = (if (j < i) then [] else (i :: iupto ((i + 1), j))); + +fun all_sums (j, []) = [] + | all_sums (j, (i :: is)) = + op_64 (map (fn x => lin_add (i, Cst x)) (iupto (1, j))) (all_sums (j, is)); + +fun split x = (fn p => x (fst p) (snd p)); + +fun negateSnd x = split (fn q => fn r => (q, ~ r)) x; + +fun adjust b = + (fn (q, r) => + (if (0 <= op_45_def2 r b) then (((2 * q) + 1), op_45_def2 r b) + else ((2 * q), r))); + +fun negDivAlg (a, b) = + (if ((0 <= (a + b)) orelse (b <= 0)) then (~1, (a + b)) + else adjust b (negDivAlg (a, (2 * b)))); + +fun posDivAlg (a, b) = + (if ((a < b) orelse (b <= 0)) then (0, a) + else adjust b (posDivAlg (a, (2 * b)))); + +fun divAlg x = + split (fn a => fn b => + (if (0 <= a) + then (if (0 <= b) then posDivAlg (a, b) + else (if (a = 0) then (0, 0) + else negateSnd (negDivAlg (~ a, ~ b)))) + else (if (0 < b) then negDivAlg (a, b) + else negateSnd (posDivAlg (~ a, ~ b))))) + x; + +fun op_mod_def1 a b = snd (divAlg (a, b)); + +fun op_dvd m n = (op_mod_def1 n m = 0); + +fun psimpl (Le (l, r)) = + (case lift_bin + ((fn x => fn y => lin_add (x, lin_neg y)), + (linearize l, linearize r)) of + None => Le (l, r) + | Some x => + (case x of Cst xa => (if (xa <= 0) then T else F) + | Var xa => Le (x, Cst 0) | Neg xa => Le (x, Cst 0) + | Add (xa, xb) => Le (x, Cst 0) | Sub (xa, xb) => Le (x, Cst 0) + | Mult (xa, xb) => Le (x, Cst 0))) + | psimpl (Eq (l, r)) = + (case lift_bin + ((fn x => fn y => lin_add (x, lin_neg y)), + (linearize l, linearize r)) of + None => Eq (l, r) + | Some x => + (case x of Cst xa => (if (xa = 0) then T else F) + | Var xa => Eq (x, Cst 0) | Neg xa => Eq (x, Cst 0) + | Add (xa, xb) => Eq (x, Cst 0) | Sub (xa, xb) => Eq (x, Cst 0) + | Mult (xa, xb) => Eq (x, Cst 0))) + | psimpl (Divides (Cst d, t)) = + (case linearize t of None => Divides (Cst d, t) + | Some x => + (case x of Cst xa => (if op_dvd d xa then T else F) + | Var xa => Divides (Cst d, x) | Neg xa => Divides (Cst d, x) + | Add (xa, xb) => Divides (Cst d, x) + | Sub (xa, xb) => Divides (Cst d, x) + | Mult (xa, xb) => Divides (Cst d, x))) + | psimpl (Equ (p, q)) = + let val p' = psimpl p; val q' = psimpl q + in (case p' of + Lt (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Gt (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Le (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Ge (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Eq (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Divides (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | T => q' + | F => (case q' of Lt (x, xa) => NOT q' | Gt (x, xa) => NOT q' + | Le (x, xa) => NOT q' | Ge (x, xa) => NOT q' + | Eq (x, xa) => NOT q' | Divides (x, xa) => NOT q' | T => F + | F => T | NOT x => x | And (x, xa) => NOT q' + | Or (x, xa) => NOT q' | Imp (x, xa) => NOT q' + | Equ (x, xa) => NOT q' | QAll x => NOT q' | QEx x => NOT q') + | NOT x => + (case q' of Lt (xa, xb) => Equ (p', q') + | Gt (xa, xb) => Equ (p', q') | Le (xa, xb) => Equ (p', q') + | Ge (xa, xb) => Equ (p', q') | Eq (xa, xb) => Equ (p', q') + | Divides (xa, xb) => Equ (p', q') | T => p' | F => x + | NOT xa => Equ (x, xa) | And (xa, xb) => Equ (p', q') + | Or (xa, xb) => Equ (p', q') | Imp (xa, xb) => Equ (p', q') + | Equ (xa, xb) => Equ (p', q') | QAll xa => Equ (p', q') + | QEx xa => Equ (p', q')) + | And (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Or (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Imp (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | Equ (x, xa) => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | QAll x => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q')) + | QEx x => + (case q' of Lt (x, xa) => Equ (p', q') | Gt (x, xa) => Equ (p', q') + | Le (x, xa) => Equ (p', q') | Ge (x, xa) => Equ (p', q') + | Eq (x, xa) => Equ (p', q') | Divides (x, xa) => Equ (p', q') + | T => p' | F => NOT p' | NOT x => Equ (p', q') + | And (x, xa) => Equ (p', q') | Or (x, xa) => Equ (p', q') + | Imp (x, xa) => Equ (p', q') | Equ (x, xa) => Equ (p', q') + | QAll x => Equ (p', q') | QEx x => Equ (p', q'))) + end + | psimpl (NOT p) = + let val p' = psimpl p + in (case p' of Lt (x, xa) => NOT p' | Gt (x, xa) => NOT p' + | Le (x, xa) => NOT p' | Ge (x, xa) => NOT p' | Eq (x, xa) => NOT p' + | Divides (x, xa) => NOT p' | T => F | F => T | NOT x => x + | And (x, xa) => NOT p' | Or (x, xa) => NOT p' | Imp (x, xa) => NOT p' + | Equ (x, xa) => NOT p' | QAll x => NOT p' | QEx x => NOT p') + end + | psimpl (Lt (u, v)) = Lt (u, v) + | psimpl (Gt (w, x)) = Gt (w, x) + | psimpl (Ge (aa, ab)) = Ge (aa, ab) + | psimpl (Divides (Var bp, af)) = Divides (Var bp, af) + | psimpl (Divides (Neg bq, af)) = Divides (Neg bq, af) + | psimpl (Divides (Add (br, bs), af)) = Divides (Add (br, bs), af) + | psimpl (Divides (Sub (bt, bu), af)) = Divides (Sub (bt, bu), af) + | psimpl (Divides (Mult (bv, bw), af)) = Divides (Mult (bv, bw), af) + | psimpl T = T + | psimpl F = F + | psimpl (QAll ap) = QAll ap + | psimpl (QEx aq) = QEx aq + | psimpl (And (p, q)) = + let val p' = psimpl p + in (case p' of + Lt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Gt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Le (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Ge (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Eq (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Divides (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | T => psimpl q | F => F + | NOT x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | And (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Or (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Imp (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | Equ (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | QAll x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end + | QEx x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => And (p', q') + | Gt (x, xa) => And (p', q') | Le (x, xa) => And (p', q') + | Ge (x, xa) => And (p', q') | Eq (x, xa) => And (p', q') + | Divides (x, xa) => And (p', q') | T => p' | F => F + | NOT x => And (p', q') | And (x, xa) => And (p', q') + | Or (x, xa) => And (p', q') | Imp (x, xa) => And (p', q') + | Equ (x, xa) => And (p', q') | QAll x => And (p', q') + | QEx x => And (p', q')) + end) + end + | psimpl (Or (p, q)) = + let val p' = psimpl p + in (case p' of + Lt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') | Gt (x, xa) => Or (p', q') + | Le (x, xa) => Or (p', q') | Ge (x, xa) => Or (p', q') + | Eq (x, xa) => Or (p', q') | Divides (x, xa) => Or (p', q') + | T => T | F => p' | NOT x => Or (p', q') + | And (x, xa) => Or (p', q') | Or (x, xa) => Or (p', q') + | Imp (x, xa) => Or (p', q') | Equ (x, xa) => Or (p', q') + | QAll x => Or (p', q') | QEx x => Or (p', q')) + end + | Gt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Le (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Ge (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Eq (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Divides (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | T => T | F => psimpl q + | NOT x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | And (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Or (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Imp (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | Equ (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | QAll x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end + | QEx x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Or (p', q') + | Gt (x, xa) => Or (p', q') | Le (x, xa) => Or (p', q') + | Ge (x, xa) => Or (p', q') | Eq (x, xa) => Or (p', q') + | Divides (x, xa) => Or (p', q') | T => T | F => p' + | NOT x => Or (p', q') | And (x, xa) => Or (p', q') + | Or (x, xa) => Or (p', q') | Imp (x, xa) => Or (p', q') + | Equ (x, xa) => Or (p', q') | QAll x => Or (p', q') + | QEx x => Or (p', q')) + end) + end + | psimpl (Imp (p, q)) = + let val p' = psimpl p + in (case p' of + Lt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Gt (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Le (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Ge (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Eq (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Divides (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | T => psimpl q | F => T + | NOT x => + let val q' = psimpl q + in (case q' of Lt (xa, xb) => Or (x, q') + | Gt (xa, xb) => Or (x, q') | Le (xa, xb) => Or (x, q') + | Ge (xa, xb) => Or (x, q') | Eq (xa, xb) => Or (x, q') + | Divides (xa, xb) => Or (x, q') | T => T | F => x + | NOT xa => Or (x, q') | And (xa, xb) => Or (x, q') + | Or (xa, xb) => Or (x, q') | Imp (xa, xb) => Or (x, q') + | Equ (xa, xb) => Or (x, q') | QAll xa => Or (x, q') + | QEx xa => Or (x, q')) + end + | And (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Or (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Imp (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | Equ (x, xa) => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | QAll x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end + | QEx x => + let val q' = psimpl q + in (case q' of Lt (x, xa) => Imp (p', q') + | Gt (x, xa) => Imp (p', q') | Le (x, xa) => Imp (p', q') + | Ge (x, xa) => Imp (p', q') | Eq (x, xa) => Imp (p', q') + | Divides (x, xa) => Imp (p', q') | T => T | F => NOT p' + | NOT x => Imp (p', q') | And (x, xa) => Imp (p', q') + | Or (x, xa) => Imp (p', q') | Imp (x, xa) => Imp (p', q') + | Equ (x, xa) => Imp (p', q') | QAll x => Imp (p', q') + | QEx x => Imp (p', q')) + end) + end; + +fun subst_it i (Cst b) = Cst b + | subst_it i (Var n) = (if (n = 0) then i else Var n) + | subst_it i (Neg it) = Neg (subst_it i it) + | subst_it i (Add (it1, it2)) = Add (subst_it i it1, subst_it i it2) + | subst_it i (Sub (it1, it2)) = Sub (subst_it i it1, subst_it i it2) + | subst_it i (Mult (it1, it2)) = Mult (subst_it i it1, subst_it i it2); + +fun subst_p i (Le (it1, it2)) = Le (subst_it i it1, subst_it i it2) + | subst_p i (Lt (it1, it2)) = Lt (subst_it i it1, subst_it i it2) + | subst_p i (Ge (it1, it2)) = Ge (subst_it i it1, subst_it i it2) + | subst_p i (Gt (it1, it2)) = Gt (subst_it i it1, subst_it i it2) + | subst_p i (Eq (it1, it2)) = Eq (subst_it i it1, subst_it i it2) + | subst_p i (Divides (d, t)) = Divides (subst_it i d, subst_it i t) + | subst_p i T = T + | subst_p i F = F + | subst_p i (And (p, q)) = And (subst_p i p, subst_p i q) + | subst_p i (Or (p, q)) = Or (subst_p i p, subst_p i q) + | subst_p i (Imp (p, q)) = Imp (subst_p i p, subst_p i q) + | subst_p i (Equ (p, q)) = Equ (subst_p i p, subst_p i q) + | subst_p i (NOT p) = NOT (subst_p i p); + +fun explode_disj ([], p) = F + | explode_disj ((i :: is), p) = + let val pi = psimpl (subst_p i p) + in (case pi of + Lt (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Gt (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Le (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Ge (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Eq (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Divides (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | T => T | F => explode_disj (is, p) + | NOT x => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | And (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Or (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Imp (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | Equ (x, xa) => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | QAll x => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end + | QEx x => + let val r = explode_disj (is, p) + in (case r of Lt (x, xa) => Or (pi, r) | Gt (x, xa) => Or (pi, r) + | Le (x, xa) => Or (pi, r) | Ge (x, xa) => Or (pi, r) + | Eq (x, xa) => Or (pi, r) | Divides (x, xa) => Or (pi, r) + | T => T | F => pi | NOT x => Or (pi, r) + | And (x, xa) => Or (pi, r) | Or (x, xa) => Or (pi, r) + | Imp (x, xa) => Or (pi, r) | Equ (x, xa) => Or (pi, r) + | QAll x => Or (pi, r) | QEx x => Or (pi, r)) + end) + end; + +fun minusinf (And (p, q)) = And (minusinf p, minusinf q) + | minusinf (Or (p, q)) = Or (minusinf p, minusinf q) + | minusinf (Lt (u, v)) = Lt (u, v) + | minusinf (Gt (w, x)) = Gt (w, x) + | minusinf (Le (Cst bo, z)) = Le (Cst bo, z) + | minusinf (Le (Var bp, z)) = Le (Var bp, z) + | minusinf (Le (Neg bq, z)) = Le (Neg bq, z) + | minusinf (Le (Add (Cst cg, bs), z)) = Le (Add (Cst cg, bs), z) + | minusinf (Le (Add (Var ch, bs), z)) = Le (Add (Var ch, bs), z) + | minusinf (Le (Add (Neg ci, bs), z)) = Le (Add (Neg ci, bs), z) + | minusinf (Le (Add (Add (cj, ck), bs), z)) = Le (Add (Add (cj, ck), bs), z) + | minusinf (Le (Add (Sub (cl, cm), bs), z)) = Le (Add (Sub (cl, cm), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = + Le (Add (Mult (Cst cy, Cst dq), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Var ei), bs), z)) = + (if (ei = 0) then (if (cy < 0) then F else T) + else Le (Add (Mult (Cst cy, Var (op_45_def0 ei id_1_def0 + 1)), bs), z)) + | minusinf (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = + Le (Add (Mult (Cst cy, Neg ds), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = + Le (Add (Mult (Cst cy, Add (dt, du)), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = + Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z) + | minusinf (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = + Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z) + | minusinf (Le (Add (Mult (Var cz, co), bs), z)) = + Le (Add (Mult (Var cz, co), bs), z) + | minusinf (Le (Add (Mult (Neg da, co), bs), z)) = + Le (Add (Mult (Neg da, co), bs), z) + | minusinf (Le (Add (Mult (Add (db, dc), co), bs), z)) = + Le (Add (Mult (Add (db, dc), co), bs), z) + | minusinf (Le (Add (Mult (Sub (dd, de), co), bs), z)) = + Le (Add (Mult (Sub (dd, de), co), bs), z) + | minusinf (Le (Add (Mult (Mult (df, dg), co), bs), z)) = + Le (Add (Mult (Mult (df, dg), co), bs), z) + | minusinf (Le (Sub (bt, bu), z)) = Le (Sub (bt, bu), z) + | minusinf (Le (Mult (bv, bw), z)) = Le (Mult (bv, bw), z) + | minusinf (Ge (aa, ab)) = Ge (aa, ab) + | minusinf (Eq (Cst ek, ad)) = Eq (Cst ek, ad) + | minusinf (Eq (Var el, ad)) = Eq (Var el, ad) + | minusinf (Eq (Neg em, ad)) = Eq (Neg em, ad) + | minusinf (Eq (Add (Cst fc, eo), ad)) = Eq (Add (Cst fc, eo), ad) + | minusinf (Eq (Add (Var fd, eo), ad)) = Eq (Add (Var fd, eo), ad) + | minusinf (Eq (Add (Neg fe, eo), ad)) = Eq (Add (Neg fe, eo), ad) + | minusinf (Eq (Add (Add (ff, fg), eo), ad)) = Eq (Add (Add (ff, fg), eo), ad) + | minusinf (Eq (Add (Sub (fh, fi), eo), ad)) = Eq (Add (Sub (fh, fi), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = + Eq (Add (Mult (Cst fu, Cst gm), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = + (if (he = 0) then F + else Eq (Add (Mult (Cst fu, Var (op_45_def0 he id_1_def0 + 1)), eo), ad)) + | minusinf (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = + Eq (Add (Mult (Cst fu, Neg go), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = + Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = + Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad) + | minusinf (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = + Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad) + | minusinf (Eq (Add (Mult (Var fv, fk), eo), ad)) = + Eq (Add (Mult (Var fv, fk), eo), ad) + | minusinf (Eq (Add (Mult (Neg fw, fk), eo), ad)) = + Eq (Add (Mult (Neg fw, fk), eo), ad) + | minusinf (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = + Eq (Add (Mult (Add (fx, fy), fk), eo), ad) + | minusinf (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = + Eq (Add (Mult (Sub (fz, ga), fk), eo), ad) + | minusinf (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = + Eq (Add (Mult (Mult (gb, gc), fk), eo), ad) + | minusinf (Eq (Sub (ep, eq), ad)) = Eq (Sub (ep, eq), ad) + | minusinf (Eq (Mult (er, es), ad)) = Eq (Mult (er, es), ad) + | minusinf (Divides (ae, af)) = Divides (ae, af) + | minusinf T = T + | minusinf F = F + | minusinf (NOT (Lt (hg, hh))) = NOT (Lt (hg, hh)) + | minusinf (NOT (Gt (hi, hj))) = NOT (Gt (hi, hj)) + | minusinf (NOT (Le (hk, hl))) = NOT (Le (hk, hl)) + | minusinf (NOT (Ge (hm, hn))) = NOT (Ge (hm, hn)) + | minusinf (NOT (Eq (Cst ja, hp))) = NOT (Eq (Cst ja, hp)) + | minusinf (NOT (Eq (Var jb, hp))) = NOT (Eq (Var jb, hp)) + | minusinf (NOT (Eq (Neg jc, hp))) = NOT (Eq (Neg jc, hp)) + | minusinf (NOT (Eq (Add (Cst js, je), hp))) = NOT (Eq (Add (Cst js, je), hp)) + | minusinf (NOT (Eq (Add (Var jt, je), hp))) = NOT (Eq (Add (Var jt, je), hp)) + | minusinf (NOT (Eq (Add (Neg ju, je), hp))) = NOT (Eq (Add (Neg ju, je), hp)) + | minusinf (NOT (Eq (Add (Add (jv, jw), je), hp))) = + NOT (Eq (Add (Add (jv, jw), je), hp)) + | minusinf (NOT (Eq (Add (Sub (jx, jy), je), hp))) = + NOT (Eq (Add (Sub (jx, jy), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = + (if (lu = 0) then T + else NOT (Eq (Add (Mult (Cst kk, Var (op_45_def0 lu id_1_def0 + 1)), je), + hp))) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = + NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = + NOT (Eq (Add (Mult (Var kl, ka), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = + NOT (Eq (Add (Mult (Neg km, ka), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = + NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = + NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp)) + | minusinf (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = + NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp)) + | minusinf (NOT (Eq (Sub (jf, jg), hp))) = NOT (Eq (Sub (jf, jg), hp)) + | minusinf (NOT (Eq (Mult (jh, ji), hp))) = NOT (Eq (Mult (jh, ji), hp)) + | minusinf (NOT (Divides (hq, hr))) = NOT (Divides (hq, hr)) + | minusinf (NOT T) = NOT T + | minusinf (NOT F) = NOT F + | minusinf (NOT (NOT hs)) = NOT (NOT hs) + | minusinf (NOT (And (ht, hu))) = NOT (And (ht, hu)) + | minusinf (NOT (Or (hv, hw))) = NOT (Or (hv, hw)) + | minusinf (NOT (Imp (hx, hy))) = NOT (Imp (hx, hy)) + | minusinf (NOT (Equ (hz, ia))) = NOT (Equ (hz, ia)) + | minusinf (NOT (QAll ib)) = NOT (QAll ib) + | minusinf (NOT (QEx ic)) = NOT (QEx ic) + | minusinf (Imp (al, am)) = Imp (al, am) + | minusinf (Equ (an, ao)) = Equ (an, ao) + | minusinf (QAll ap) = QAll ap + | minusinf (QEx aq) = QEx aq; + +fun abs i = (if (i < 0) then ~ i else i); + +fun op_div_def1 a b = fst (divAlg (a, b)); + +fun op_mod_def0 m n = nat (op_mod_def1 (m) (n)); + +fun ngcd (m, n) = (if (n = 0) then m else ngcd (n, op_mod_def0 m n)); + +fun igcd x = split (fn a => fn b => (ngcd (nat (abs a), nat (abs b)))) x; + +fun ilcm a b = op_div_def1 (a * b) (igcd (a, b)); + +fun divlcm (NOT p) = divlcm p + | divlcm (And (p, q)) = ilcm (divlcm p) (divlcm q) + | divlcm (Or (p, q)) = ilcm (divlcm p) (divlcm q) + | divlcm (Lt (u, v)) = 1 + | divlcm (Gt (w, x)) = 1 + | divlcm (Le (y, z)) = 1 + | divlcm (Ge (aa, ab)) = 1 + | divlcm (Eq (ac, ad)) = 1 + | divlcm (Divides (Cst bo, Cst cg)) = 1 + | divlcm (Divides (Cst bo, Var ch)) = 1 + | divlcm (Divides (Cst bo, Neg ci)) = 1 + | divlcm (Divides (Cst bo, Add (Cst cy, ck))) = 1 + | divlcm (Divides (Cst bo, Add (Var cz, ck))) = 1 + | divlcm (Divides (Cst bo, Add (Neg da, ck))) = 1 + | divlcm (Divides (Cst bo, Add (Add (db, dc), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Sub (dd, de), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Cst ei), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Var fa), ck))) = + (if (fa = 0) then abs bo else 1) + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Neg ek), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Add (el, em)), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Sub (en, eo)), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Cst dq, Mult (ep, eq)), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Var dr, dg), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Neg ds, dg), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Add (dt, du), dg), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Sub (dv, dw), dg), ck))) = 1 + | divlcm (Divides (Cst bo, Add (Mult (Mult (dx, dy), dg), ck))) = 1 + | divlcm (Divides (Cst bo, Sub (cl, cm))) = 1 + | divlcm (Divides (Cst bo, Mult (cn, co))) = 1 + | divlcm (Divides (Var bp, af)) = 1 + | divlcm (Divides (Neg bq, af)) = 1 + | divlcm (Divides (Add (br, bs), af)) = 1 + | divlcm (Divides (Sub (bt, bu), af)) = 1 + | divlcm (Divides (Mult (bv, bw), af)) = 1 + | divlcm T = 1 + | divlcm F = 1 + | divlcm (Imp (al, am)) = 1 + | divlcm (Equ (an, ao)) = 1 + | divlcm (QAll ap) = 1 + | divlcm (QEx aq) = 1; + +fun explode_minf (q, B) = + let val d = divlcm q; val pm = minusinf q; + val dj1 = explode_disj (map (fn x => Cst x) (iupto (1, d)), pm) + in (case dj1 of + Lt (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Gt (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Le (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Ge (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Eq (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Divides (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | T => T | F => explode_disj (all_sums (d, B), q) + | NOT x => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | And (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Or (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Imp (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | Equ (x, xa) => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | QAll x => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end + | QEx x => + let val dj2 = explode_disj (all_sums (d, B), q) + in (case dj2 of Lt (x, xa) => Or (dj1, dj2) + | Gt (x, xa) => Or (dj1, dj2) | Le (x, xa) => Or (dj1, dj2) + | Ge (x, xa) => Or (dj1, dj2) | Eq (x, xa) => Or (dj1, dj2) + | Divides (x, xa) => Or (dj1, dj2) | T => T | F => dj1 + | NOT x => Or (dj1, dj2) | And (x, xa) => Or (dj1, dj2) + | Or (x, xa) => Or (dj1, dj2) | Imp (x, xa) => Or (dj1, dj2) + | Equ (x, xa) => Or (dj1, dj2) | QAll x => Or (dj1, dj2) + | QEx x => Or (dj1, dj2)) + end) + end; + +fun mirror (And (p, q)) = And (mirror p, mirror q) + | mirror (Or (p, q)) = Or (mirror p, mirror q) + | mirror (Lt (u, v)) = Lt (u, v) + | mirror (Gt (w, x)) = Gt (w, x) + | mirror (Le (Cst bp, aa)) = Le (Cst bp, aa) + | mirror (Le (Var bq, aa)) = Le (Var bq, aa) + | mirror (Le (Neg br, aa)) = Le (Neg br, aa) + | mirror (Le (Add (Cst ch, bt), aa)) = Le (Add (Cst ch, bt), aa) + | mirror (Le (Add (Var ci, bt), aa)) = Le (Add (Var ci, bt), aa) + | mirror (Le (Add (Neg cj, bt), aa)) = Le (Add (Neg cj, bt), aa) + | mirror (Le (Add (Add (ck, cl), bt), aa)) = Le (Add (Add (ck, cl), bt), aa) + | mirror (Le (Add (Sub (cm, cn), bt), aa)) = Le (Add (Sub (cm, cn), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Cst dr), bt), aa)) = + Le (Add (Mult (Cst cz, Cst dr), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Var ej), bt), aa)) = + (if (ej = 0) then Le (Add (Mult (Cst (~ cz), Var 0), bt), aa) + else Le (Add (Mult (Cst cz, Var (op_45_def0 ej id_1_def0 + 1)), bt), aa)) + | mirror (Le (Add (Mult (Cst cz, Neg dt), bt), aa)) = + Le (Add (Mult (Cst cz, Neg dt), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Add (du, dv)), bt), aa)) = + Le (Add (Mult (Cst cz, Add (du, dv)), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Sub (dw, dx)), bt), aa)) = + Le (Add (Mult (Cst cz, Sub (dw, dx)), bt), aa) + | mirror (Le (Add (Mult (Cst cz, Mult (dy, dz)), bt), aa)) = + Le (Add (Mult (Cst cz, Mult (dy, dz)), bt), aa) + | mirror (Le (Add (Mult (Var da, cp), bt), aa)) = + Le (Add (Mult (Var da, cp), bt), aa) + | mirror (Le (Add (Mult (Neg db, cp), bt), aa)) = + Le (Add (Mult (Neg db, cp), bt), aa) + | mirror (Le (Add (Mult (Add (dc, dd), cp), bt), aa)) = + Le (Add (Mult (Add (dc, dd), cp), bt), aa) + | mirror (Le (Add (Mult (Sub (de, df), cp), bt), aa)) = + Le (Add (Mult (Sub (de, df), cp), bt), aa) + | mirror (Le (Add (Mult (Mult (dg, dh), cp), bt), aa)) = + Le (Add (Mult (Mult (dg, dh), cp), bt), aa) + | mirror (Le (Sub (bu, bv), aa)) = Le (Sub (bu, bv), aa) + | mirror (Le (Mult (bw, bx), aa)) = Le (Mult (bw, bx), aa) + | mirror (Ge (ab, ac)) = Ge (ab, ac) + | mirror (Eq (Cst el, ae)) = Eq (Cst el, ae) + | mirror (Eq (Var em, ae)) = Eq (Var em, ae) + | mirror (Eq (Neg en, ae)) = Eq (Neg en, ae) + | mirror (Eq (Add (Cst fd, ep), ae)) = Eq (Add (Cst fd, ep), ae) + | mirror (Eq (Add (Var fe, ep), ae)) = Eq (Add (Var fe, ep), ae) + | mirror (Eq (Add (Neg ff, ep), ae)) = Eq (Add (Neg ff, ep), ae) + | mirror (Eq (Add (Add (fg, fh), ep), ae)) = Eq (Add (Add (fg, fh), ep), ae) + | mirror (Eq (Add (Sub (fi, fj), ep), ae)) = Eq (Add (Sub (fi, fj), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Cst gn), ep), ae)) = + Eq (Add (Mult (Cst fv, Cst gn), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Var hf), ep), ae)) = + (if (hf = 0) then Eq (Add (Mult (Cst (~ fv), Var 0), ep), ae) + else Eq (Add (Mult (Cst fv, Var (op_45_def0 hf id_1_def0 + 1)), ep), ae)) + | mirror (Eq (Add (Mult (Cst fv, Neg gp), ep), ae)) = + Eq (Add (Mult (Cst fv, Neg gp), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Add (gq, gr)), ep), ae)) = + Eq (Add (Mult (Cst fv, Add (gq, gr)), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Sub (gs, gt)), ep), ae)) = + Eq (Add (Mult (Cst fv, Sub (gs, gt)), ep), ae) + | mirror (Eq (Add (Mult (Cst fv, Mult (gu, gv)), ep), ae)) = + Eq (Add (Mult (Cst fv, Mult (gu, gv)), ep), ae) + | mirror (Eq (Add (Mult (Var fw, fl), ep), ae)) = + Eq (Add (Mult (Var fw, fl), ep), ae) + | mirror (Eq (Add (Mult (Neg fx, fl), ep), ae)) = + Eq (Add (Mult (Neg fx, fl), ep), ae) + | mirror (Eq (Add (Mult (Add (fy, fz), fl), ep), ae)) = + Eq (Add (Mult (Add (fy, fz), fl), ep), ae) + | mirror (Eq (Add (Mult (Sub (ga, gb), fl), ep), ae)) = + Eq (Add (Mult (Sub (ga, gb), fl), ep), ae) + | mirror (Eq (Add (Mult (Mult (gc, gd), fl), ep), ae)) = + Eq (Add (Mult (Mult (gc, gd), fl), ep), ae) + | mirror (Eq (Sub (eq, er), ae)) = Eq (Sub (eq, er), ae) + | mirror (Eq (Mult (es, et), ae)) = Eq (Mult (es, et), ae) + | mirror (Divides (Cst hh, Cst hz)) = Divides (Cst hh, Cst hz) + | mirror (Divides (Cst hh, Var ia)) = Divides (Cst hh, Var ia) + | mirror (Divides (Cst hh, Neg ib)) = Divides (Cst hh, Neg ib) + | mirror (Divides (Cst hh, Add (Cst ir, id))) = + Divides (Cst hh, Add (Cst ir, id)) + | mirror (Divides (Cst hh, Add (Var is, id))) = + Divides (Cst hh, Add (Var is, id)) + | mirror (Divides (Cst hh, Add (Neg it, id))) = + Divides (Cst hh, Add (Neg it, id)) + | mirror (Divides (Cst hh, Add (Add (iu, iv), id))) = + Divides (Cst hh, Add (Add (iu, iv), id)) + | mirror (Divides (Cst hh, Add (Sub (iw, ix), id))) = + Divides (Cst hh, Add (Sub (iw, ix), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Cst kb), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Cst kb), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Var kt), id))) = + (if (kt = 0) then Divides (Cst hh, Add (Mult (Cst (~ jj), Var 0), id)) + else Divides + (Cst hh, + Add (Mult (Cst jj, Var (op_45_def0 kt id_1_def0 + 1)), id))) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Neg kd), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Neg kd), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Add (ke, kf)), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Add (ke, kf)), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Sub (kg, kh)), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Sub (kg, kh)), id)) + | mirror (Divides (Cst hh, Add (Mult (Cst jj, Mult (ki, kj)), id))) = + Divides (Cst hh, Add (Mult (Cst jj, Mult (ki, kj)), id)) + | mirror (Divides (Cst hh, Add (Mult (Var jk, iz), id))) = + Divides (Cst hh, Add (Mult (Var jk, iz), id)) + | mirror (Divides (Cst hh, Add (Mult (Neg jl, iz), id))) = + Divides (Cst hh, Add (Mult (Neg jl, iz), id)) + | mirror (Divides (Cst hh, Add (Mult (Add (jm, jn), iz), id))) = + Divides (Cst hh, Add (Mult (Add (jm, jn), iz), id)) + | mirror (Divides (Cst hh, Add (Mult (Sub (jo, jp), iz), id))) = + Divides (Cst hh, Add (Mult (Sub (jo, jp), iz), id)) + | mirror (Divides (Cst hh, Add (Mult (Mult (jq, jr), iz), id))) = + Divides (Cst hh, Add (Mult (Mult (jq, jr), iz), id)) + | mirror (Divides (Cst hh, Sub (ie, if'))) = Divides (Cst hh, Sub (ie, if')) + | mirror (Divides (Cst hh, Mult (ig, ih))) = Divides (Cst hh, Mult (ig, ih)) + | mirror (Divides (Var hi, ag)) = Divides (Var hi, ag) + | mirror (Divides (Neg hj, ag)) = Divides (Neg hj, ag) + | mirror (Divides (Add (hk, hl), ag)) = Divides (Add (hk, hl), ag) + | mirror (Divides (Sub (hm, hn), ag)) = Divides (Sub (hm, hn), ag) + | mirror (Divides (Mult (ho, hp), ag)) = Divides (Mult (ho, hp), ag) + | mirror T = T + | mirror F = F + | mirror (NOT (Lt (kv, kw))) = NOT (Lt (kv, kw)) + | mirror (NOT (Gt (kx, ky))) = NOT (Gt (kx, ky)) + | mirror (NOT (Le (kz, la))) = NOT (Le (kz, la)) + | mirror (NOT (Ge (lb, lc))) = NOT (Ge (lb, lc)) + | mirror (NOT (Eq (Cst mp, le))) = NOT (Eq (Cst mp, le)) + | mirror (NOT (Eq (Var mq, le))) = NOT (Eq (Var mq, le)) + | mirror (NOT (Eq (Neg mr, le))) = NOT (Eq (Neg mr, le)) + | mirror (NOT (Eq (Add (Cst nh, mt), le))) = NOT (Eq (Add (Cst nh, mt), le)) + | mirror (NOT (Eq (Add (Var ni, mt), le))) = NOT (Eq (Add (Var ni, mt), le)) + | mirror (NOT (Eq (Add (Neg nj, mt), le))) = NOT (Eq (Add (Neg nj, mt), le)) + | mirror (NOT (Eq (Add (Add (nk, nl), mt), le))) = + NOT (Eq (Add (Add (nk, nl), mt), le)) + | mirror (NOT (Eq (Add (Sub (nm, nn), mt), le))) = + NOT (Eq (Add (Sub (nm, nn), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Cst or), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Cst or), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Var pj), mt), le))) = + (if (pj = 0) then NOT (Eq (Add (Mult (Cst (~ nz), Var 0), mt), le)) + else NOT (Eq (Add (Mult (Cst nz, Var (op_45_def0 pj id_1_def0 + 1)), mt), + le))) + | mirror (NOT (Eq (Add (Mult (Cst nz, Neg ot), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Neg ot), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Add (ou, ov)), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Add (ou, ov)), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Sub (ow, ox)), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Sub (ow, ox)), mt), le)) + | mirror (NOT (Eq (Add (Mult (Cst nz, Mult (oy, oz)), mt), le))) = + NOT (Eq (Add (Mult (Cst nz, Mult (oy, oz)), mt), le)) + | mirror (NOT (Eq (Add (Mult (Var oa, np), mt), le))) = + NOT (Eq (Add (Mult (Var oa, np), mt), le)) + | mirror (NOT (Eq (Add (Mult (Neg ob, np), mt), le))) = + NOT (Eq (Add (Mult (Neg ob, np), mt), le)) + | mirror (NOT (Eq (Add (Mult (Add (oc, od), np), mt), le))) = + NOT (Eq (Add (Mult (Add (oc, od), np), mt), le)) + | mirror (NOT (Eq (Add (Mult (Sub (oe, of'), np), mt), le))) = + NOT (Eq (Add (Mult (Sub (oe, of'), np), mt), le)) + | mirror (NOT (Eq (Add (Mult (Mult (og, oh), np), mt), le))) = + NOT (Eq (Add (Mult (Mult (og, oh), np), mt), le)) + | mirror (NOT (Eq (Sub (mu, mv), le))) = NOT (Eq (Sub (mu, mv), le)) + | mirror (NOT (Eq (Mult (mw, mx), le))) = NOT (Eq (Mult (mw, mx), le)) + | mirror (NOT (Divides (Cst pl, Cst qd))) = NOT (Divides (Cst pl, Cst qd)) + | mirror (NOT (Divides (Cst pl, Var qe))) = NOT (Divides (Cst pl, Var qe)) + | mirror (NOT (Divides (Cst pl, Neg qf))) = NOT (Divides (Cst pl, Neg qf)) + | mirror (NOT (Divides (Cst pl, Add (Cst qv, qh)))) = + NOT (Divides (Cst pl, Add (Cst qv, qh))) + | mirror (NOT (Divides (Cst pl, Add (Var qw, qh)))) = + NOT (Divides (Cst pl, Add (Var qw, qh))) + | mirror (NOT (Divides (Cst pl, Add (Neg qx, qh)))) = + NOT (Divides (Cst pl, Add (Neg qx, qh))) + | mirror (NOT (Divides (Cst pl, Add (Add (qy, qz), qh)))) = + NOT (Divides (Cst pl, Add (Add (qy, qz), qh))) + | mirror (NOT (Divides (Cst pl, Add (Sub (ra, rb), qh)))) = + NOT (Divides (Cst pl, Add (Sub (ra, rb), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Cst sf), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Cst sf), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Var sx), qh)))) = + (if (sx = 0) + then NOT (Divides (Cst pl, Add (Mult (Cst (~ rn), Var 0), qh))) + else NOT (Divides + (Cst pl, + Add (Mult (Cst rn, Var (op_45_def0 sx id_1_def0 + 1)), + qh)))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Neg sh), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Neg sh), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Add (si, sj)), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Add (si, sj)), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Sub (sk, sl)), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Sub (sk, sl)), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Cst rn, Mult (sm, sn)), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Cst rn, Mult (sm, sn)), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Var ro, rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Var ro, rd), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Neg rp, rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Neg rp, rd), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Add (rq, rr), rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Add (rq, rr), rd), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Sub (rs, rt), rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Sub (rs, rt), rd), qh))) + | mirror (NOT (Divides (Cst pl, Add (Mult (Mult (ru, rv), rd), qh)))) = + NOT (Divides (Cst pl, Add (Mult (Mult (ru, rv), rd), qh))) + | mirror (NOT (Divides (Cst pl, Sub (qi, qj)))) = + NOT (Divides (Cst pl, Sub (qi, qj))) + | mirror (NOT (Divides (Cst pl, Mult (qk, ql)))) = + NOT (Divides (Cst pl, Mult (qk, ql))) + | mirror (NOT (Divides (Var pm, lg))) = NOT (Divides (Var pm, lg)) + | mirror (NOT (Divides (Neg pn, lg))) = NOT (Divides (Neg pn, lg)) + | mirror (NOT (Divides (Add (po, pp), lg))) = NOT (Divides (Add (po, pp), lg)) + | mirror (NOT (Divides (Sub (pq, pr), lg))) = NOT (Divides (Sub (pq, pr), lg)) + | mirror (NOT (Divides (Mult (ps, pt), lg))) = + NOT (Divides (Mult (ps, pt), lg)) + | mirror (NOT T) = NOT T + | mirror (NOT F) = NOT F + | mirror (NOT (NOT lh)) = NOT (NOT lh) + | mirror (NOT (And (li, lj))) = NOT (And (li, lj)) + | mirror (NOT (Or (lk, ll))) = NOT (Or (lk, ll)) + | mirror (NOT (Imp (lm, ln))) = NOT (Imp (lm, ln)) + | mirror (NOT (Equ (lo, lp))) = NOT (Equ (lo, lp)) + | mirror (NOT (QAll lq)) = NOT (QAll lq) + | mirror (NOT (QEx lr)) = NOT (QEx lr) + | mirror (Imp (am, an)) = Imp (am, an) + | mirror (Equ (ao, ap)) = Equ (ao, ap) + | mirror (QAll aq) = QAll aq + | mirror (QEx ar) = QEx ar; + +fun op_43_def0 m n = nat ((m) + (n)); + +fun size_def1 [] = 0 + | size_def1 (a :: list) = op_43_def0 (size_def1 list) (0 + 1); + +fun aset (And (p, q)) = op_64 (aset p) (aset q) + | aset (Or (p, q)) = op_64 (aset p) (aset q) + | aset (Lt (u, v)) = [] + | aset (Gt (w, x)) = [] + | aset (Le (Cst bo, z)) = [] + | aset (Le (Var bp, z)) = [] + | aset (Le (Neg bq, z)) = [] + | aset (Le (Add (Cst cg, bs), z)) = [] + | aset (Le (Add (Var ch, bs), z)) = [] + | aset (Le (Add (Neg ci, bs), z)) = [] + | aset (Le (Add (Add (cj, ck), bs), z)) = [] + | aset (Le (Add (Sub (cl, cm), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Var ei), bs), z)) = + (if (ei = 0) + then (if (cy < 0) then [lin_add (bs, Cst 1)] + else [lin_neg bs, lin_add (lin_neg bs, Cst 1)]) + else []) + | aset (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = [] + | aset (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = [] + | aset (Le (Add (Mult (Var cz, co), bs), z)) = [] + | aset (Le (Add (Mult (Neg da, co), bs), z)) = [] + | aset (Le (Add (Mult (Add (db, dc), co), bs), z)) = [] + | aset (Le (Add (Mult (Sub (dd, de), co), bs), z)) = [] + | aset (Le (Add (Mult (Mult (df, dg), co), bs), z)) = [] + | aset (Le (Sub (bt, bu), z)) = [] + | aset (Le (Mult (bv, bw), z)) = [] + | aset (Ge (aa, ab)) = [] + | aset (Eq (Cst ek, ad)) = [] + | aset (Eq (Var el, ad)) = [] + | aset (Eq (Neg em, ad)) = [] + | aset (Eq (Add (Cst fc, eo), ad)) = [] + | aset (Eq (Add (Var fd, eo), ad)) = [] + | aset (Eq (Add (Neg fe, eo), ad)) = [] + | aset (Eq (Add (Add (ff, fg), eo), ad)) = [] + | aset (Eq (Add (Sub (fh, fi), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = + (if (he = 0) + then (if (fu < 0) then [lin_add (eo, Cst 1)] + else [lin_add (lin_neg eo, Cst 1)]) + else []) + | aset (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = [] + | aset (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = [] + | aset (Eq (Add (Mult (Var fv, fk), eo), ad)) = [] + | aset (Eq (Add (Mult (Neg fw, fk), eo), ad)) = [] + | aset (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = [] + | aset (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = [] + | aset (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = [] + | aset (Eq (Sub (ep, eq), ad)) = [] + | aset (Eq (Mult (er, es), ad)) = [] + | aset (Divides (ae, af)) = [] + | aset T = [] + | aset F = [] + | aset (NOT (Lt (hg, hh))) = [] + | aset (NOT (Gt (hi, hj))) = [] + | aset (NOT (Le (hk, hl))) = [] + | aset (NOT (Ge (hm, hn))) = [] + | aset (NOT (Eq (Cst ja, hp))) = [] + | aset (NOT (Eq (Var jb, hp))) = [] + | aset (NOT (Eq (Neg jc, hp))) = [] + | aset (NOT (Eq (Add (Cst js, je), hp))) = [] + | aset (NOT (Eq (Add (Var jt, je), hp))) = [] + | aset (NOT (Eq (Add (Neg ju, je), hp))) = [] + | aset (NOT (Eq (Add (Add (jv, jw), je), hp))) = [] + | aset (NOT (Eq (Add (Sub (jx, jy), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = + (if (lu = 0) then (if (kk < 0) then [je] else [lin_neg je]) else []) + | aset (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = [] + | aset (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = [] + | aset (NOT (Eq (Sub (jf, jg), hp))) = [] + | aset (NOT (Eq (Mult (jh, ji), hp))) = [] + | aset (NOT (Divides (hq, hr))) = [] + | aset (NOT T) = [] + | aset (NOT F) = [] + | aset (NOT (NOT hs)) = [] + | aset (NOT (And (ht, hu))) = [] + | aset (NOT (Or (hv, hw))) = [] + | aset (NOT (Imp (hx, hy))) = [] + | aset (NOT (Equ (hz, ia))) = [] + | aset (NOT (QAll ib)) = [] + | aset (NOT (QEx ic)) = [] + | aset (Imp (al, am)) = [] + | aset (Equ (an, ao)) = [] + | aset (QAll ap) = [] + | aset (QEx aq) = []; + +fun op_mem x [] = false + | op_mem x (y :: ys) = (if (y = x) then true else op_mem x ys); + +fun list_insert x xs = (if op_mem x xs then xs else (x :: xs)); + +fun list_set [] = [] + | list_set (x :: xs) = list_insert x (list_set xs); + +fun bset (And (p, q)) = op_64 (bset p) (bset q) + | bset (Or (p, q)) = op_64 (bset p) (bset q) + | bset (Lt (u, v)) = [] + | bset (Gt (w, x)) = [] + | bset (Le (Cst bo, z)) = [] + | bset (Le (Var bp, z)) = [] + | bset (Le (Neg bq, z)) = [] + | bset (Le (Add (Cst cg, bs), z)) = [] + | bset (Le (Add (Var ch, bs), z)) = [] + | bset (Le (Add (Neg ci, bs), z)) = [] + | bset (Le (Add (Add (cj, ck), bs), z)) = [] + | bset (Le (Add (Sub (cl, cm), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Var ei), bs), z)) = + (if (ei = 0) + then (if (cy < 0) then [lin_add (bs, Cst ~1), bs] + else [lin_add (lin_neg bs, Cst ~1)]) + else []) + | bset (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = [] + | bset (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = [] + | bset (Le (Add (Mult (Var cz, co), bs), z)) = [] + | bset (Le (Add (Mult (Neg da, co), bs), z)) = [] + | bset (Le (Add (Mult (Add (db, dc), co), bs), z)) = [] + | bset (Le (Add (Mult (Sub (dd, de), co), bs), z)) = [] + | bset (Le (Add (Mult (Mult (df, dg), co), bs), z)) = [] + | bset (Le (Sub (bt, bu), z)) = [] + | bset (Le (Mult (bv, bw), z)) = [] + | bset (Ge (aa, ab)) = [] + | bset (Eq (Cst ek, ad)) = [] + | bset (Eq (Var el, ad)) = [] + | bset (Eq (Neg em, ad)) = [] + | bset (Eq (Add (Cst fc, eo), ad)) = [] + | bset (Eq (Add (Var fd, eo), ad)) = [] + | bset (Eq (Add (Neg fe, eo), ad)) = [] + | bset (Eq (Add (Add (ff, fg), eo), ad)) = [] + | bset (Eq (Add (Sub (fh, fi), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Cst gm), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Var he), eo), ad)) = + (if (he = 0) + then (if (fu < 0) then [lin_add (eo, Cst ~1)] + else [lin_add (lin_neg eo, Cst ~1)]) + else []) + | bset (Eq (Add (Mult (Cst fu, Neg go), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Add (gp, gq)), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Sub (gr, gs)), eo), ad)) = [] + | bset (Eq (Add (Mult (Cst fu, Mult (gt, gu)), eo), ad)) = [] + | bset (Eq (Add (Mult (Var fv, fk), eo), ad)) = [] + | bset (Eq (Add (Mult (Neg fw, fk), eo), ad)) = [] + | bset (Eq (Add (Mult (Add (fx, fy), fk), eo), ad)) = [] + | bset (Eq (Add (Mult (Sub (fz, ga), fk), eo), ad)) = [] + | bset (Eq (Add (Mult (Mult (gb, gc), fk), eo), ad)) = [] + | bset (Eq (Sub (ep, eq), ad)) = [] + | bset (Eq (Mult (er, es), ad)) = [] + | bset (Divides (ae, af)) = [] + | bset T = [] + | bset F = [] + | bset (NOT (Lt (hg, hh))) = [] + | bset (NOT (Gt (hi, hj))) = [] + | bset (NOT (Le (hk, hl))) = [] + | bset (NOT (Ge (hm, hn))) = [] + | bset (NOT (Eq (Cst ja, hp))) = [] + | bset (NOT (Eq (Var jb, hp))) = [] + | bset (NOT (Eq (Neg jc, hp))) = [] + | bset (NOT (Eq (Add (Cst js, je), hp))) = [] + | bset (NOT (Eq (Add (Var jt, je), hp))) = [] + | bset (NOT (Eq (Add (Neg ju, je), hp))) = [] + | bset (NOT (Eq (Add (Add (jv, jw), je), hp))) = [] + | bset (NOT (Eq (Add (Sub (jx, jy), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Cst lc), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Var lu), je), hp))) = + (if (lu = 0) then (if (kk < 0) then [je] else [lin_neg je]) else []) + | bset (NOT (Eq (Add (Mult (Cst kk, Neg le), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Add (lf, lg)), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Sub (lh, li)), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Cst kk, Mult (lj, lk)), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Var kl, ka), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Neg km, ka), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Add (kn, ko), ka), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Sub (kp, kq), ka), je), hp))) = [] + | bset (NOT (Eq (Add (Mult (Mult (kr, ks), ka), je), hp))) = [] + | bset (NOT (Eq (Sub (jf, jg), hp))) = [] + | bset (NOT (Eq (Mult (jh, ji), hp))) = [] + | bset (NOT (Divides (hq, hr))) = [] + | bset (NOT T) = [] + | bset (NOT F) = [] + | bset (NOT (NOT hs)) = [] + | bset (NOT (And (ht, hu))) = [] + | bset (NOT (Or (hv, hw))) = [] + | bset (NOT (Imp (hx, hy))) = [] + | bset (NOT (Equ (hz, ia))) = [] + | bset (NOT (QAll ib)) = [] + | bset (NOT (QEx ic)) = [] + | bset (Imp (al, am)) = [] + | bset (Equ (an, ao)) = [] + | bset (QAll ap) = [] + | bset (QEx aq) = []; + +fun adjustcoeff (l, Le (Add (Mult (Cst c, Var 0), r), Cst i)) = + (if (c <= 0) + then Le (Add (Mult (Cst ~1, Var 0), lin_mul (~ (op_div_def1 l c), r)), + Cst 0) + else Le (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0)) + | adjustcoeff (l, Eq (Add (Mult (Cst c, Var 0), r), Cst i)) = + Eq (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst c, Var 0), r), Cst i))) = + NOT (Eq (Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l c, r)), Cst 0)) + | adjustcoeff (l, And (p, q)) = And (adjustcoeff (l, p), adjustcoeff (l, q)) + | adjustcoeff (l, Or (p, q)) = Or (adjustcoeff (l, p), adjustcoeff (l, q)) + | adjustcoeff (l, Lt (w, x)) = Lt (w, x) + | adjustcoeff (l, Gt (y, z)) = Gt (y, z) + | adjustcoeff (l, Le (Cst bq, ab)) = Le (Cst bq, ab) + | adjustcoeff (l, Le (Var br, ab)) = Le (Var br, ab) + | adjustcoeff (l, Le (Neg bs, ab)) = Le (Neg bs, ab) + | adjustcoeff (l, Le (Add (Cst ci, bu), ab)) = Le (Add (Cst ci, bu), ab) + | adjustcoeff (l, Le (Add (Var cj, bu), ab)) = Le (Add (Var cj, bu), ab) + | adjustcoeff (l, Le (Add (Neg ck, bu), ab)) = Le (Add (Neg ck, bu), ab) + | adjustcoeff (l, Le (Add (Add (cl, cm), bu), ab)) = + Le (Add (Add (cl, cm), bu), ab) + | adjustcoeff (l, Le (Add (Sub (cn, co), bu), ab)) = + Le (Add (Sub (cn, co), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Cst ds), bu), ab)) = + Le (Add (Mult (Cst da, Cst ds), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Var en)) = + Le (Add (Mult (Cst da, Var 0), bu), Var en) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Neg eo)) = + Le (Add (Mult (Cst da, Var 0), bu), Neg eo) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Add (ep, eq))) = + Le (Add (Mult (Cst da, Var 0), bu), Add (ep, eq)) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Sub (er, es))) = + Le (Add (Mult (Cst da, Var 0), bu), Sub (er, es)) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var 0), bu), Mult (et, eu))) = + Le (Add (Mult (Cst da, Var 0), bu), Mult (et, eu)) + | adjustcoeff (l, Le (Add (Mult (Cst da, Var ek), bu), ab)) = + Le (Add (Mult (Cst da, Var ek), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Neg du), bu), ab)) = + Le (Add (Mult (Cst da, Neg du), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Add (dv, dw)), bu), ab)) = + Le (Add (Mult (Cst da, Add (dv, dw)), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Sub (dx, dy)), bu), ab)) = + Le (Add (Mult (Cst da, Sub (dx, dy)), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Cst da, Mult (dz, ea)), bu), ab)) = + Le (Add (Mult (Cst da, Mult (dz, ea)), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Var db, cq), bu), ab)) = + Le (Add (Mult (Var db, cq), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Neg dc, cq), bu), ab)) = + Le (Add (Mult (Neg dc, cq), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Add (dd, de), cq), bu), ab)) = + Le (Add (Mult (Add (dd, de), cq), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Sub (df, dg), cq), bu), ab)) = + Le (Add (Mult (Sub (df, dg), cq), bu), ab) + | adjustcoeff (l, Le (Add (Mult (Mult (dh, di), cq), bu), ab)) = + Le (Add (Mult (Mult (dh, di), cq), bu), ab) + | adjustcoeff (l, Le (Sub (bv, bw), ab)) = Le (Sub (bv, bw), ab) + | adjustcoeff (l, Le (Mult (bx, by), ab)) = Le (Mult (bx, by), ab) + | adjustcoeff (l, Ge (ac, ad)) = Ge (ac, ad) + | adjustcoeff (l, Eq (Cst fe, af)) = Eq (Cst fe, af) + | adjustcoeff (l, Eq (Var ff, af)) = Eq (Var ff, af) + | adjustcoeff (l, Eq (Neg fg, af)) = Eq (Neg fg, af) + | adjustcoeff (l, Eq (Add (Cst fw, fi), af)) = Eq (Add (Cst fw, fi), af) + | adjustcoeff (l, Eq (Add (Var fx, fi), af)) = Eq (Add (Var fx, fi), af) + | adjustcoeff (l, Eq (Add (Neg fy, fi), af)) = Eq (Add (Neg fy, fi), af) + | adjustcoeff (l, Eq (Add (Add (fz, ga), fi), af)) = + Eq (Add (Add (fz, ga), fi), af) + | adjustcoeff (l, Eq (Add (Sub (gb, gc), fi), af)) = + Eq (Add (Sub (gb, gc), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Cst hg), fi), af)) = + Eq (Add (Mult (Cst go, Cst hg), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Var ib)) = + Eq (Add (Mult (Cst go, Var 0), fi), Var ib) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Neg ic)) = + Eq (Add (Mult (Cst go, Var 0), fi), Neg ic) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Add (id, ie))) = + Eq (Add (Mult (Cst go, Var 0), fi), Add (id, ie)) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Sub (if', ig))) = + Eq (Add (Mult (Cst go, Var 0), fi), Sub (if', ig)) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var 0), fi), Mult (ih, ii))) = + Eq (Add (Mult (Cst go, Var 0), fi), Mult (ih, ii)) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Var hy), fi), af)) = + Eq (Add (Mult (Cst go, Var hy), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Neg hi), fi), af)) = + Eq (Add (Mult (Cst go, Neg hi), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Add (hj, hk)), fi), af)) = + Eq (Add (Mult (Cst go, Add (hj, hk)), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Sub (hl, hm)), fi), af)) = + Eq (Add (Mult (Cst go, Sub (hl, hm)), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Cst go, Mult (hn, ho)), fi), af)) = + Eq (Add (Mult (Cst go, Mult (hn, ho)), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Var gp, ge), fi), af)) = + Eq (Add (Mult (Var gp, ge), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Neg gq, ge), fi), af)) = + Eq (Add (Mult (Neg gq, ge), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Add (gr, gs), ge), fi), af)) = + Eq (Add (Mult (Add (gr, gs), ge), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Sub (gt, gu), ge), fi), af)) = + Eq (Add (Mult (Sub (gt, gu), ge), fi), af) + | adjustcoeff (l, Eq (Add (Mult (Mult (gv, gw), ge), fi), af)) = + Eq (Add (Mult (Mult (gv, gw), ge), fi), af) + | adjustcoeff (l, Eq (Sub (fj, fk), af)) = Eq (Sub (fj, fk), af) + | adjustcoeff (l, Eq (Mult (fl, fm), af)) = Eq (Mult (fl, fm), af) + | adjustcoeff (l, Divides (Cst is, Cst jk)) = Divides (Cst is, Cst jk) + | adjustcoeff (l, Divides (Cst is, Var jl)) = Divides (Cst is, Var jl) + | adjustcoeff (l, Divides (Cst is, Neg jm)) = Divides (Cst is, Neg jm) + | adjustcoeff (l, Divides (Cst is, Add (Cst kc, jo))) = + Divides (Cst is, Add (Cst kc, jo)) + | adjustcoeff (l, Divides (Cst is, Add (Var kd, jo))) = + Divides (Cst is, Add (Var kd, jo)) + | adjustcoeff (l, Divides (Cst is, Add (Neg ke, jo))) = + Divides (Cst is, Add (Neg ke, jo)) + | adjustcoeff (l, Divides (Cst is, Add (Add (kf, kg), jo))) = + Divides (Cst is, Add (Add (kf, kg), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Sub (kh, ki), jo))) = + Divides (Cst is, Add (Sub (kh, ki), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Cst lm), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Cst lm), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Var me), jo))) = + (if (me = 0) + then Divides + (Cst (op_div_def1 l ku * is), + Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l ku, jo))) + else Divides + (Cst is, + Add (Mult (Cst ku, Var (op_45_def0 me id_1_def0 + 1)), jo))) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Neg lo), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Neg lo), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Add (lp, lq)), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Add (lp, lq)), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Sub (lr, ls)), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Sub (lr, ls)), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Cst ku, Mult (lt, lu)), jo))) = + Divides (Cst is, Add (Mult (Cst ku, Mult (lt, lu)), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Var kv, kk), jo))) = + Divides (Cst is, Add (Mult (Var kv, kk), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Neg kw, kk), jo))) = + Divides (Cst is, Add (Mult (Neg kw, kk), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Add (kx, ky), kk), jo))) = + Divides (Cst is, Add (Mult (Add (kx, ky), kk), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Sub (kz, la), kk), jo))) = + Divides (Cst is, Add (Mult (Sub (kz, la), kk), jo)) + | adjustcoeff (l, Divides (Cst is, Add (Mult (Mult (lb, lc), kk), jo))) = + Divides (Cst is, Add (Mult (Mult (lb, lc), kk), jo)) + | adjustcoeff (l, Divides (Cst is, Sub (jp, jq))) = + Divides (Cst is, Sub (jp, jq)) + | adjustcoeff (l, Divides (Cst is, Mult (jr, js))) = + Divides (Cst is, Mult (jr, js)) + | adjustcoeff (l, Divides (Var it, ah)) = Divides (Var it, ah) + | adjustcoeff (l, Divides (Neg iu, ah)) = Divides (Neg iu, ah) + | adjustcoeff (l, Divides (Add (iv, iw), ah)) = Divides (Add (iv, iw), ah) + | adjustcoeff (l, Divides (Sub (ix, iy), ah)) = Divides (Sub (ix, iy), ah) + | adjustcoeff (l, Divides (Mult (iz, ja), ah)) = Divides (Mult (iz, ja), ah) + | adjustcoeff (l, T) = T + | adjustcoeff (l, F) = F + | adjustcoeff (l, NOT (Lt (mg, mh))) = NOT (Lt (mg, mh)) + | adjustcoeff (l, NOT (Gt (mi, mj))) = NOT (Gt (mi, mj)) + | adjustcoeff (l, NOT (Le (mk, ml))) = NOT (Le (mk, ml)) + | adjustcoeff (l, NOT (Ge (mm, mn))) = NOT (Ge (mm, mn)) + | adjustcoeff (l, NOT (Eq (Cst oa, mp))) = NOT (Eq (Cst oa, mp)) + | adjustcoeff (l, NOT (Eq (Var ob, mp))) = NOT (Eq (Var ob, mp)) + | adjustcoeff (l, NOT (Eq (Neg oc, mp))) = NOT (Eq (Neg oc, mp)) + | adjustcoeff (l, NOT (Eq (Add (Cst os, oe), mp))) = + NOT (Eq (Add (Cst os, oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Var ot, oe), mp))) = + NOT (Eq (Add (Var ot, oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Neg ou, oe), mp))) = + NOT (Eq (Add (Neg ou, oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Add (ov, ow), oe), mp))) = + NOT (Eq (Add (Add (ov, ow), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Sub (ox, oy), oe), mp))) = + NOT (Eq (Add (Sub (ox, oy), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Cst qc), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Cst qc), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Var qx))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Var qx)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Neg qy))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Neg qy)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Add (qz, ra)))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Add (qz, ra))) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Sub (rb, rc)))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Sub (rb, rc))) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Mult (rd, re)))) = + NOT (Eq (Add (Mult (Cst pk, Var 0), oe), Mult (rd, re))) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Var qu), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Var qu), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Neg qe), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Neg qe), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Add (qf, qg)), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Add (qf, qg)), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Sub (qh, qi)), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Sub (qh, qi)), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Cst pk, Mult (qj, qk)), oe), mp))) = + NOT (Eq (Add (Mult (Cst pk, Mult (qj, qk)), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Var pl, pa), oe), mp))) = + NOT (Eq (Add (Mult (Var pl, pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Neg pm, pa), oe), mp))) = + NOT (Eq (Add (Mult (Neg pm, pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Add (pn, po), pa), oe), mp))) = + NOT (Eq (Add (Mult (Add (pn, po), pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Sub (pp, pq), pa), oe), mp))) = + NOT (Eq (Add (Mult (Sub (pp, pq), pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Add (Mult (Mult (pr, ps), pa), oe), mp))) = + NOT (Eq (Add (Mult (Mult (pr, ps), pa), oe), mp)) + | adjustcoeff (l, NOT (Eq (Sub (of', og), mp))) = NOT (Eq (Sub (of', og), mp)) + | adjustcoeff (l, NOT (Eq (Mult (oh, oi), mp))) = NOT (Eq (Mult (oh, oi), mp)) + | adjustcoeff (l, NOT (Divides (Cst ro, Cst sg))) = + NOT (Divides (Cst ro, Cst sg)) + | adjustcoeff (l, NOT (Divides (Cst ro, Var sh))) = + NOT (Divides (Cst ro, Var sh)) + | adjustcoeff (l, NOT (Divides (Cst ro, Neg si))) = + NOT (Divides (Cst ro, Neg si)) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Cst sy, sk)))) = + NOT (Divides (Cst ro, Add (Cst sy, sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Var sz, sk)))) = + NOT (Divides (Cst ro, Add (Var sz, sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Neg ta, sk)))) = + NOT (Divides (Cst ro, Add (Neg ta, sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Add (tb, tc), sk)))) = + NOT (Divides (Cst ro, Add (Add (tb, tc), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Sub (td, te), sk)))) = + NOT (Divides (Cst ro, Add (Sub (td, te), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Cst ui), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Cst ui), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Var va), sk)))) = + (if (va = 0) + then NOT (Divides + (Cst (op_div_def1 l tq * ro), + Add (Mult (Cst 1, Var 0), lin_mul (op_div_def1 l tq, sk)))) + else NOT (Divides + (Cst ro, + Add (Mult (Cst tq, Var (op_45_def0 va id_1_def0 + 1)), + sk)))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Neg uk), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Neg uk), sk))) + | adjustcoeff + (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Add (ul, um)), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Add (ul, um)), sk))) + | adjustcoeff + (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Sub (un, uo)), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Sub (un, uo)), sk))) + | adjustcoeff + (l, NOT (Divides (Cst ro, Add (Mult (Cst tq, Mult (up, uq)), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Cst tq, Mult (up, uq)), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Var tr, tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Var tr, tg), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Neg ts, tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Neg ts, tg), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Add (tt, tu), tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Add (tt, tu), tg), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Add (Mult (Sub (tv, tw), tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Sub (tv, tw), tg), sk))) + | adjustcoeff + (l, NOT (Divides (Cst ro, Add (Mult (Mult (tx, ty), tg), sk)))) = + NOT (Divides (Cst ro, Add (Mult (Mult (tx, ty), tg), sk))) + | adjustcoeff (l, NOT (Divides (Cst ro, Sub (sl, sm)))) = + NOT (Divides (Cst ro, Sub (sl, sm))) + | adjustcoeff (l, NOT (Divides (Cst ro, Mult (sn, so)))) = + NOT (Divides (Cst ro, Mult (sn, so))) + | adjustcoeff (l, NOT (Divides (Var rp, mr))) = NOT (Divides (Var rp, mr)) + | adjustcoeff (l, NOT (Divides (Neg rq, mr))) = NOT (Divides (Neg rq, mr)) + | adjustcoeff (l, NOT (Divides (Add (rr, rs), mr))) = + NOT (Divides (Add (rr, rs), mr)) + | adjustcoeff (l, NOT (Divides (Sub (rt, ru), mr))) = + NOT (Divides (Sub (rt, ru), mr)) + | adjustcoeff (l, NOT (Divides (Mult (rv, rw), mr))) = + NOT (Divides (Mult (rv, rw), mr)) + | adjustcoeff (l, NOT T) = NOT T + | adjustcoeff (l, NOT F) = NOT F + | adjustcoeff (l, NOT (NOT ms)) = NOT (NOT ms) + | adjustcoeff (l, NOT (And (mt, mu))) = NOT (And (mt, mu)) + | adjustcoeff (l, NOT (Or (mv, mw))) = NOT (Or (mv, mw)) + | adjustcoeff (l, NOT (Imp (mx, my))) = NOT (Imp (mx, my)) + | adjustcoeff (l, NOT (Equ (mz, na))) = NOT (Equ (mz, na)) + | adjustcoeff (l, NOT (QAll nb)) = NOT (QAll nb) + | adjustcoeff (l, NOT (QEx nc)) = NOT (QEx nc) + | adjustcoeff (l, Imp (an, ao)) = Imp (an, ao) + | adjustcoeff (l, Equ (ap, aq)) = Equ (ap, aq) + | adjustcoeff (l, QAll ar) = QAll ar + | adjustcoeff (l, QEx as') = QEx as'; + +fun formlcm (Le (Add (Mult (Cst c, Var 0), r), Cst i)) = abs c + | formlcm (Eq (Add (Mult (Cst c, Var 0), r), Cst i)) = abs c + | formlcm (NOT p) = formlcm p + | formlcm (And (p, q)) = ilcm (formlcm p) (formlcm q) + | formlcm (Or (p, q)) = ilcm (formlcm p) (formlcm q) + | formlcm (Lt (u, v)) = 1 + | formlcm (Gt (w, x)) = 1 + | formlcm (Le (Cst bo, z)) = 1 + | formlcm (Le (Var bp, z)) = 1 + | formlcm (Le (Neg bq, z)) = 1 + | formlcm (Le (Add (Cst cg, bs), z)) = 1 + | formlcm (Le (Add (Var ch, bs), z)) = 1 + | formlcm (Le (Add (Neg ci, bs), z)) = 1 + | formlcm (Le (Add (Add (cj, ck), bs), z)) = 1 + | formlcm (Le (Add (Sub (cl, cm), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Cst dq), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Var el)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Neg em)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Add (en, eo))) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Sub (ep, eq))) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var 0), bs), Mult (er, es))) = 1 + | formlcm (Le (Add (Mult (Cst cy, Var ei ), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Neg ds), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Add (dt, du)), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Sub (dv, dw)), bs), z)) = 1 + | formlcm (Le (Add (Mult (Cst cy, Mult (dx, dy)), bs), z)) = 1 + | formlcm (Le (Add (Mult (Var cz, co), bs), z)) = 1 + | formlcm (Le (Add (Mult (Neg da, co), bs), z)) = 1 + | formlcm (Le (Add (Mult (Add (db, dc), co), bs), z)) = 1 + | formlcm (Le (Add (Mult (Sub (dd, de), co), bs), z)) = 1 + | formlcm (Le (Add (Mult (Mult (df, dg), co), bs), z)) = 1 + | formlcm (Le (Sub (bt, bu), z)) = 1 + | formlcm (Le (Mult (bv, bw), z)) = 1 + | formlcm (Ge (aa, ab)) = 1 + | formlcm (Eq (Cst fc, ad)) = 1 + | formlcm (Eq (Var fd, ad)) = 1 + | formlcm (Eq (Neg fe, ad)) = 1 + | formlcm (Eq (Add (Cst fu, fg), ad)) = 1 + | formlcm (Eq (Add (Var fv, fg), ad)) = 1 + | formlcm (Eq (Add (Neg fw, fg), ad)) = 1 + | formlcm (Eq (Add (Add (fx, fy), fg), ad)) = 1 + | formlcm (Eq (Add (Sub (fz, ga), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Cst he), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Var hz)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Neg ia)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Add (ib, ic))) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Sub (id, ie))) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var 0), fg), Mult (if', ig))) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Var hw), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Neg hg), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Add (hh, hi)), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Sub (hj, hk)), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Cst gm, Mult (hl, hm)), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Var gn, gc), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Neg go, gc), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Add (gp, gq), gc), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Sub (gr, gs), gc), fg), ad)) = 1 + | formlcm (Eq (Add (Mult (Mult (gt, gu), gc), fg), ad)) = 1 + | formlcm (Eq (Sub (fh, fi), ad)) = 1 + | formlcm (Eq (Mult (fj, fk), ad)) = 1 + | formlcm (Divides (Cst iq, Cst ji)) = 1 + | formlcm (Divides (Cst iq, Var jj)) = 1 + | formlcm (Divides (Cst iq, Neg jk)) = 1 + | formlcm (Divides (Cst iq, Add (Cst ka, jm))) = 1 + | formlcm (Divides (Cst iq, Add (Var kb, jm))) = 1 + | formlcm (Divides (Cst iq, Add (Neg kc, jm))) = 1 + | formlcm (Divides (Cst iq, Add (Add (kd, ke), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Sub (kf, kg), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Cst lk), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Var mc), jm))) = + (if (mc = 0) then abs ks else 1) + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Neg lm), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Add (ln, lo)), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Sub (lp, lq)), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Cst ks, Mult (lr, ls)), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Var kt, ki), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Neg ku, ki), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Add (kv, kw), ki), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Sub (kx, ky), ki), jm))) = 1 + | formlcm (Divides (Cst iq, Add (Mult (Mult (kz, la), ki), jm))) = 1 + | formlcm (Divides (Cst iq, Sub (jn, jo))) = 1 + | formlcm (Divides (Cst iq, Mult (jp, jq))) = 1 + | formlcm (Divides (Var ir, af)) = 1 + | formlcm (Divides (Neg is, af)) = 1 + | formlcm (Divides (Add (it, iu), af)) = 1 + | formlcm (Divides (Sub (iv, iw), af)) = 1 + | formlcm (Divides (Mult (ix, iy), af)) = 1 + | formlcm T = 1 + | formlcm F = 1 + | formlcm (Imp (al, am)) = 1 + | formlcm (Equ (an, ao)) = 1 + | formlcm (QAll ap) = 1 + | formlcm (QEx aq) = 1; + +fun unitycoeff p = + let val l = formlcm p; val p' = adjustcoeff (l, p) + in (if (l = 1) then p' + else And (Divides (Cst l, Add (Mult (Cst 1, Var 0), Cst 0)), p')) + end; + +fun unify p = + let val q = unitycoeff p; val B = list_set (bset q); val A = list_set (aset q) + in (if op_60_61_def0 (size_def1 B) (size_def1 A) then (q, B) + else (mirror q, map lin_neg A)) + end; + +fun cooper p = + lift_un (fn q => decrvars (explode_minf (unify q))) (linform (nnf p)); + +fun pa p = lift_un psimpl (qelim (cooper, p)); + +val test = pa; + +end; diff -r afaa031ed4da -r 105519771c67 src/HOL/Tools/comm_ring.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/comm_ring.ML Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,134 @@ +(* ID: $Id$ + Author: Amine Chaieb + + Tactic for solving equalities over commutative rings *) +signature CRING = +sig + val cring_tac : int -> tactic + val cring_method: int -> Proof.method + val setup : (theory -> theory) list +end +structure CommRing :CRING = +struct + +(* The Cring exception for erronous uses of cring_tac *) +exception CRing of string; + +(* Zero and One of the commutative ring *) +fun cring_zero T = Const("0",T); +fun cring_one T = Const("1",T); + +(* reification functions *) +(* add two polynom expressions *) +fun polT t = Type ("Commutative_Ring.pol",[t]); +fun polexT t = Type("Commutative_Ring.polex",[t]); +val nT = HOLogic.natT; +fun listT T = Type ("List.list",[T]); + +(* Reification of the constructors *) +(* Nat*) +val succ = Const("Suc",nT --> nT); +val zero = Const("0",nT); +val one = Const("1",nT); + +(* Lists *) +fun reif_list T [] = Const("List.list.Nil",listT T) + | reif_list T (x::xs) = Const("List.list.Cons",[T,listT T] ---> listT T) + $x$(reif_list T xs); + +(* pol*) +fun pol_Pc t = Const("Commutative_Ring.pol.Pc",t --> polT t); +fun pol_Pinj t = Const("Commutative_Ring.pol.Pinj",[nT,polT t] ---> polT t); +fun pol_PX t = Const("Commutative_Ring.pol.PX",[polT t, nT, polT t] ---> polT t); + +(* polex *) +fun polex_add t = Const("Commutative_Ring.polex.Add",[polexT t,polexT t] ---> polexT t); +fun polex_sub t = Const("Commutative_Ring.polex.Sub",[polexT t,polexT t] ---> polexT t); +fun polex_mul t = Const("Commutative_Ring.polex.Mul",[polexT t,polexT t] ---> polexT t); +fun polex_neg t = Const("Commutative_Ring.polex.Neg",polexT t --> polexT t); +fun polex_pol t = Const("Commutative_Ring.polex.Pol",polT t --> polexT t); +fun polex_pow t = Const("Commutative_Ring.polex.Pow",[polexT t, nT] ---> polexT t); +(* reification of natural numbers *) +fun reif_nat n = + if n>0 then succ$(reif_nat (n-1)) + else if n=0 then zero + else raise CRing "ring_tac: reif_nat negative n"; + +(* reification of polynoms : primitive cring expressions *) +fun reif_pol T vs t = + case t of + Free(_,_) => + let val i = find_index_eq t vs + in if i = 0 + then (pol_PX T)$((pol_Pc T)$ (cring_one T)) + $one$((pol_Pc T)$(cring_zero T)) + else (pol_Pinj T)$(reif_nat i)$ + ((pol_PX T)$((pol_Pc T)$ (cring_one T)) + $one$ + ((pol_Pc T)$(cring_zero T))) + end + | _ => (pol_Pc T)$ t; + + +(* reification of polynom expressions *) +fun reif_polex T vs t = + case t of + Const("op +",_)$a$b => (polex_add T) + $ (reif_polex T vs a)$(reif_polex T vs b) + | Const("op -",_)$a$b => (polex_sub T) + $ (reif_polex T vs a)$(reif_polex T vs b) + | Const("op *",_)$a$b => (polex_mul T) + $ (reif_polex T vs a)$ (reif_polex T vs b) + | Const("uminus",_)$a => (polex_neg T) + $ (reif_polex T vs a) + | (Const("Nat.power",_)$a$n) => (polex_pow T) $ (reif_polex T vs a) $ n + + | _ => (polex_pol T) $ (reif_pol T vs t); + +(* reification of the equation *) +val cr_sort = Sign.read_sort (the_context ()) "{comm_ring,recpower}"; +fun reif_eq sg (eq as Const("op =",Type("fun",a::_))$lhs$rhs) = + if Sign.of_sort (the_context()) (a,cr_sort) + then + let val fs = term_frees eq + val cvs = cterm_of sg (reif_list a fs) + val clhs = cterm_of sg (reif_polex a fs lhs) + val crhs = cterm_of sg (reif_polex a fs rhs) + val ca = ctyp_of sg a + in (ca,cvs,clhs, crhs) + end + else raise CRing "reif_eq: not an equation over comm_ring + recpower" + | reif_eq sg _ = raise CRing "reif_eq: not an equation"; + +(*The cring tactic *) +(* Attention: You have to make sure that no t^0 is in the goal!! *) +(* Use simply rewriting t^0 = 1 *) +val cring_ss = simpset_of (the_context()) + addsimps + (map thm ["mkPX_def", "mkPinj_def","sub_def", + "power_add","even_def","pow_if"]) + addsimps [sym OF [thm "power_add"]]; + +val norm_eq = thm "norm_eq" +fun comm_ring_tac i =(fn st => + let + val g = BasisLibrary.List.nth (prems_of st, i - 1) + val sg = sign_of_thm st + val (ca,cvs,clhs,crhs) = reif_eq sg (HOLogic.dest_Trueprop g) + val norm_eq_th = simplify cring_ss + (instantiate' [SOME ca] [SOME clhs, SOME crhs, SOME cvs] + norm_eq) + in ((cut_rules_tac [norm_eq_th] i) + THEN (simp_tac cring_ss i) + THEN (simp_tac cring_ss i)) st + end); + +fun cring_method i = Method.METHOD (fn facts => + Method.insert_tac facts 1 THEN comm_ring_tac i); + +val setup = + [Method.add_method ("comm_ring", + Method.no_args (cring_method 1), + "reflective decision procedure for equalities over commutative rings")]; + +end; \ No newline at end of file diff -r afaa031ed4da -r 105519771c67 src/HOL/ex/Commutative_RingEx.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Commutative_RingEx.thy Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,47 @@ + (* $Id$ *) +theory Commutative_RingEx +imports Main +begin + + (* Some examples demonstrating the comm_ring method *) + +lemma "4*(x::int)^5*y^3*x^2*3 + x*z + 3^5 = 12*x^7*y^3 + z*x + 243" +by comm_ring + +lemma "((x::int) + y)^2 = x^2 + y^2 + 2*x*y" +by comm_ring + +lemma "((x::int) + y)^3 = x^3 + y^3 + 3*x^2*y + 3*y^2*x" +by comm_ring + +lemma "((x::int) - y)^3 = x^3 + 3*x*y^2 + (-3)*y*x^2 - y^3" +by comm_ring + +lemma "((x::int) - y)^2 = x^2 + y^2 - 2*x*y" +by comm_ring + +lemma " ((a::int) + b + c)^2 = a^2 + b^2 + c^2 + 2*a*b + 2*b*c + 2*a*c" +by comm_ring + +lemma "((a::int) - b - c)^2 = a^2 + b^2 + c^2 - 2*a*b + 2*b*c - 2*a*c" +by comm_ring + +lemma "(a::int)*b + a*c = a*(b+c)" +by comm_ring + +lemma "(a::int)^2 - b^2 = (a - b) * (a + b)" +by comm_ring + +lemma "(a::int)^3 - b^3 = (a - b) * (a^2 + a*b + b^2)" +by comm_ring + +lemma "(a::int)^3 + b^3 = (a + b) * (a^2 - a*b + b^2)" +by comm_ring + +lemma "(a::int)^4 - b^4 = (a - b) * (a + b)*(a^2 + b^2)" +by comm_ring + +lemma "(a::int)^10 - b^10 = (a - b) * (a^9 + a^8*b + a^7*b^2 + a^6*b^3 + a^5*b^4 + a^4*b^5 + a^3*b^6 + a^2*b^7 + a*b^8 + b^9 )" +by comm_ring + +end \ No newline at end of file diff -r afaa031ed4da -r 105519771c67 src/HOL/ex/Commutative_Ring_Complete.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Commutative_Ring_Complete.thy Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,384 @@ +(* ID: $Id$ + Author: Bernhard Haeupler + + This theory is about the relative completeness of the tactic + As long as the reified atomic polynomials of type 'a pol + are in normal form, the cring method is complete *) + +theory Commutative_Ring_Complete +imports Main +begin + + (* Fromalization of normal form *) +consts isnorm :: "('a::{comm_ring,recpower}) pol \ bool" +recdef isnorm "measure size" + "isnorm (Pc c) = True" + "isnorm (Pinj i (Pc c)) = False" + "isnorm (Pinj i (Pinj j Q)) = False" + "isnorm (Pinj 0 P) = False" + "isnorm (Pinj i (PX Q1 j Q2)) = isnorm (PX Q1 j Q2)" + "isnorm (PX P 0 Q) = False" + "isnorm (PX (Pc c) i Q) = (c \ 0 & isnorm Q)" + "isnorm (PX (PX P1 j (Pc c)) i Q) = (c\0 \ isnorm(PX P1 j (Pc c))\isnorm Q)" + "isnorm (PX P i Q) = (isnorm P \ isnorm Q)" + +(* Some helpful lemmas *) +lemma norm_Pinj_0_False:"isnorm (Pinj 0 P) = False" +by(cases P, auto) + +lemma norm_PX_0_False:"isnorm (PX (Pc 0) i Q) = False" +by(cases i, auto) + +lemma norm_Pinj:"isnorm (Pinj i Q) \ isnorm Q" +by(cases i,simp add: norm_Pinj_0_False norm_PX_0_False,cases Q) auto + +lemma norm_PX2:"isnorm (PX P i Q) \ isnorm Q" +by(cases i, auto, cases P, auto, case_tac pol2, auto) + +lemma norm_PX1:"isnorm (PX P i Q) \ isnorm P" +by(cases i, auto, cases P, auto, case_tac pol2, auto) + +lemma mkPinj_cn:"\y~=0; isnorm Q\ \ isnorm (mkPinj y Q)" +apply(auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split) +apply(case_tac nat, auto simp add: norm_Pinj_0_False) +by(case_tac pol, auto) (case_tac y, auto) + +lemma norm_PXtrans: + assumes A:"isnorm (PX P x Q)" and "isnorm Q2" + shows "isnorm (PX P x Q2)" +proof(cases P) + case (PX p1 y p2) from prems show ?thesis by(cases x, auto, cases p2, auto) +next + case Pc from prems show ?thesis by(cases x, auto) +next + case Pinj from prems show ?thesis by(cases x, auto) +qed + + +lemma norm_PXtrans2: assumes A:"isnorm (PX P x Q)" and "isnorm Q2" shows "isnorm (PX P (Suc (n+x)) Q2)" +proof(cases P) + case (PX p1 y p2) + from prems show ?thesis by(cases x, auto, cases p2, auto) +next + case Pc + from prems show ?thesis by(cases x, auto) +next + case Pinj + from prems show ?thesis by(cases x, auto) +qed + + (* mkPX conserves normalizedness (_cn)*) +lemma mkPX_cn: + assumes "x \ 0" and "isnorm P" and "isnorm Q" + shows "isnorm (mkPX P x Q)" +proof(cases P) + case (Pc c) + from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) +next + case (Pinj i Q) + from prems show ?thesis by (cases x) (auto simp add: mkPinj_cn mkPX_def) +next + case (PX P1 y P2) + from prems have Y0:"y>0" by(cases y, auto) + from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2]) + with prems Y0 show ?thesis by (cases x, auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _], cases P2, auto) +qed + + (* add conserves normalizedness *) +lemma add_cn:"\isnorm P; (isnorm Q)\ \ isnorm (add (P, Q))" +proof(induct P Q rule: add.induct) + case (2 c i P2) thus ?case by (cases P2, simp_all, cases i, simp_all) +next + case (3 i P2 c) thus ?case by (cases P2, simp_all, cases i, simp_all) +next + case (4 c P2 i Q2) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) +next + case (5 P2 i Q2 c) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case by(cases i, simp, cases P2, auto, case_tac pol2, auto) +next + case (6 x P2 y Q2) + from prems have Y0:"y>0" by (cases y, auto simp add: norm_Pinj_0_False) + from prems have X0:"x>0" by (cases x, auto simp add: norm_Pinj_0_False) + have "x < y \ x = y \ x > y" by arith + moreover + { assume "xy" hence "EX d. x=d+y" by arith + then obtain d where "x=d+y".. + moreover + note prems Y0 + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) + ultimately have ?case by (simp add: mkPinj_cn)} + ultimately show ?case by blast +next + case (7 x P2 Q2 y R) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (add (R, P2))" by simp + with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + with prems NR have "isnorm( add (R, Pinj (x - 1) P2))" "isnorm(PX Q2 y R)" by simp + with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + ultimately show ?case by blast +next + case (8 Q2 y R x P2) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (add (R, P2))" by simp + with prems have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + with prems NR have "isnorm( add (R, Pinj (x - 1) P2))" "isnorm(PX Q2 y R)" by simp + with X have ?case by (simp add: norm_PXtrans[of Q2 y _]) } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + from prems have Y0:"y>0" by(cases y, auto) + from prems have X0:"x>0" by(cases x, auto) + from prems have NP1:"isnorm P1" and NP2:"isnorm P2" by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2]) + from prems have NQ1:"isnorm Q1" and NQ2:"isnorm Q2" by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2]) + have "y < x \ x = y \ x < y" by arith + moreover + {assume sm1:"y < x" hence "EX d. x=d+y" by arith + then obtain d where sm2:"x=d+y".. + note prems NQ1 NP1 NP2 NQ2 sm1 sm2 + moreover + have "isnorm (PX P1 d (Pc 0))" + proof(cases P1) + case (PX p1 y p2) + with prems show ?thesis by(cases d, simp,cases p2, auto) + next case Pc from prems show ?thesis by(cases d, auto) + next case Pinj from prems show ?thesis by(cases d, auto) + qed + ultimately have "isnorm (add (P2, Q2))" "isnorm (add (PX P1 (x - y) (Pc 0), Q1))" by auto + with Y0 sm1 sm2 have ?case by (simp add: mkPX_cn)} + moreover + {assume "x=y" + from prems NP1 NP2 NQ1 NQ2 have "isnorm (add (P2, Q2))" "isnorm (add (P1, Q1))" by auto + with Y0 prems have ?case by (simp add: mkPX_cn) } + moreover + {assume sm1:"xisnorm P; (isnorm Q)\ \ isnorm (mul (P, Q))" +proof(induct P Q rule: mul.induct) + case (2 c i P2) thus ?case + by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) +next + case (3 i P2 c) thus ?case + by (cases P2, simp_all) (cases "i",simp_all add: mkPinj_cn) +next + case (4 c P2 i Q2) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case + by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) +next + case (5 P2 i Q2 c) + from prems have "isnorm P2" "isnorm Q2" by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) + with prems show ?case + by - (case_tac "c=0",simp_all,case_tac "i=0",simp_all add: mkPX_cn) +next + case (6 x P2 y Q2) + have "x < y \ x = y \ x > y" by arith + moreover + { assume "x0" by (cases x, auto simp add: norm_Pinj_0_False) + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d Q2)" by (cases d, simp, cases Q2, auto) + ultimately have ?case by (simp add: mkPinj_cn)} + moreover + { assume "x=y" + moreover + from prems have "isnorm P2" "isnorm Q2" by(auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) + moreover + note prems + moreover + ultimately have ?case by (simp add: mkPinj_cn) } + moreover + { assume "x>y" hence "EX d. x=d+y" by arith + then obtain d where "x=d+y".. + moreover + note prems + moreover + from prems have "y>0" by (cases y, auto simp add: norm_Pinj_0_False) + moreover + from prems have "isnorm P2" "isnorm Q2" by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) + moreover + with prems have "isnorm (Pinj d P2)" by (cases d, simp, cases P2, auto) + ultimately have ?case by (simp add: mkPinj_cn) } + ultimately show ?case by blast +next + case (7 x P2 Q2 y R) + from prems have Y0:"y>0" by(cases y, auto) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (mul (R, P2))" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) + with Y0 prems have ?case by (simp add: mkPX_cn)} + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) + moreover + from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + moreover + from prems have "isnorm (Pinj x P2)" by(cases P2, auto) + moreover + note prems + ultimately have "isnorm (mul (R, Pinj (x - 1) P2))" "isnorm (mul (Pinj x P2, Q2))" by auto + with Y0 X have ?case by (simp add: mkPX_cn)} + ultimately show ?case by blast +next + case (8 Q2 y R x P2) + from prems have Y0:"y>0" by(cases y, auto) + have "x=0 \ (x = 1) \ (x > 1)" by arith + moreover + { assume "x=0" with prems have ?case by (auto simp add: norm_Pinj_0_False)} + moreover + { assume "x=1" + from prems have "isnorm R" "isnorm P2" by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) + with prems have "isnorm (mul (R, P2))" "isnorm Q2" by (auto simp add: norm_PX1[of Q2 y R]) + with Y0 prems have ?case by (simp add: mkPX_cn) } + moreover + { assume "x > 1" hence "EX d. x=Suc (Suc d)" by arith + then obtain d where X:"x=Suc (Suc d)" .. + from prems have NR:"isnorm R" "isnorm Q2" by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) + moreover + from prems have "isnorm (Pinj (x - 1) P2)" by(cases P2, auto) + moreover + from prems have "isnorm (Pinj x P2)" by(cases P2, auto) + moreover + note prems + ultimately have "isnorm (mul (R, Pinj (x - 1) P2))" "isnorm (mul (Pinj x P2, Q2))" by auto + with Y0 X have ?case by (simp add: mkPX_cn) } + ultimately show ?case by blast +next + case (9 P1 x P2 Q1 y Q2) + from prems have X0:"x>0" by(cases x, auto) + from prems have Y0:"y>0" by(cases y, auto) + note prems + moreover + from prems have "isnorm P1" "isnorm P2" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + moreover + from prems have "isnorm Q1" "isnorm Q2" by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2]) + ultimately have "isnorm (mul (P1, Q1))" "isnorm (mul (P2, Q2))" "isnorm (mul (P1, mkPinj 1 Q2))" "isnorm (mul (Q1, mkPinj 1 P2))" + by (auto simp add: mkPinj_cn) + with prems X0 Y0 have "isnorm (mkPX (mul (P1, Q1)) (x + y) (mul (P2, Q2)))" "isnorm (mkPX (mul (P1, mkPinj (Suc 0) Q2)) x (Pc 0))" + "isnorm (mkPX (mul (Q1, mkPinj (Suc 0) P2)) y (Pc 0))" + by (auto simp add: mkPX_cn) + thus ?case by (simp add: add_cn) +qed(simp) + + (* neg conserves normalizedness *) +lemma neg_cn: "isnorm P \ isnorm (neg P)" +proof(induct P rule: neg.induct) + case (Pinj i P2) + from prems have "isnorm P2" by (simp add: norm_Pinj[of i P2]) + with prems show ?case by(cases P2, auto, cases i, auto) +next + case (PX P1 x P2) + from prems have "isnorm P2" "isnorm P1" by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + with prems show ?case + proof(cases P1) + case (PX p1 y p2) + with prems show ?thesis by(cases x, auto, cases p2, auto) + next + case Pinj + with prems show ?thesis by(cases x, auto) + qed(cases x, auto) +qed(simp) + + (* sub conserves normalizedness *) +lemma sub_cn:"\isnorm p; isnorm q\ \ isnorm (sub p q)" +by (simp add: sub_def add_cn neg_cn) + + (* sqr conserves normalizizedness *) +lemma sqr_cn:"isnorm P \ isnorm (sqr P)" +proof(induct P) + case (Pinj i Q) + from prems show ?case by(cases Q, auto simp add: mkPX_cn mkPinj_cn, cases i, auto simp add: mkPX_cn mkPinj_cn) +next + case (PX P1 x P2) + from prems have "x+x~=0" "isnorm P2" "isnorm P1" by (cases x, auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) + with prems have "isnorm (mkPX (mul (mul (Pc ((1\'a) + (1\'a)), P1), mkPinj (Suc 0) P2)) x (Pc (0\'a)))" + and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))" by( auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) + thus ?case by( auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) +qed(simp) + + + (* pow conserves normalizedness *) +lemma pow_cn:"!! P. \isnorm P\ \ isnorm (pow (P, n))" +proof(induct n rule: nat_less_induct) + case (1 k) + show ?case + proof(cases "k=0") + case False + hence K2:"k div 2 < k" by (cases k, auto) + from prems have "isnorm (sqr P)" by (simp add: sqr_cn) + with prems K2 show ?thesis by(simp add: allE[of _ "(k div 2)" _] allE[of _ "(sqr P)" _], cases k, auto simp add: mul_cn) + qed(simp) +qed + +end \ No newline at end of file diff -r afaa031ed4da -r 105519771c67 src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Wed Sep 14 10:24:39 2005 +0200 +++ b/src/HOL/ex/ROOT.ML Wed Sep 14 17:25:52 2005 +0200 @@ -25,6 +25,7 @@ time_use_thy "CTL"; time_use_thy "mesontest2"; time_use_thy "PresburgerEx"; +time_use_thy "Reflected_Presburger"; time_use_thy "BT"; time_use_thy "InSort"; time_use_thy "Qsort"; @@ -33,6 +34,8 @@ time_use_thy "Lagrange"; +time_use_thy "Commutative_RingEx"; +time_use_thy "Commutative_Ring_Complete"; time_use_thy "set"; time_use_thy "MT"; diff -r afaa031ed4da -r 105519771c67 src/HOL/ex/Reflected_Presburger.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Reflected_Presburger.thy Wed Sep 14 17:25:52 2005 +0200 @@ -0,0 +1,5726 @@ +(* A formalization of quantifier elimination for Presburger arithmetic*) +(* based on a generic quantifier elimination algorithm and + cooper's methos to eliminate on \ *) + +theory Reflected_Presburger +imports Main +uses ("rcooper.ML") ("rpresbtac.ML") +begin + +(* Shadow syntax for integer terms *) +datatype intterm = + Cst int + | Var nat + | Neg intterm + | Add intterm intterm + | Sub intterm intterm + | Mult intterm intterm + +(* interpretatation of intterms , takes bound variables and free variables *) +consts I_intterm :: "int list \ intterm \ int" +primrec +"I_intterm ats (Cst b) = b" +"I_intterm ats (Var n) = (ats!n)" +"I_intterm ats (Neg it) = -(I_intterm ats it)" +"I_intterm ats (Add it1 it2) = (I_intterm ats it1) + (I_intterm ats it2)" +"I_intterm ats (Sub it1 it2) = (I_intterm ats it1) - (I_intterm ats it2)" +"I_intterm ats (Mult it1 it2) = (I_intterm ats it1) * (I_intterm ats it2)" + +(*Shadow syntax for Presburger formulae *) +datatype QF = + Lt intterm intterm + |Gt intterm intterm + |Le intterm intterm + |Ge intterm intterm + |Eq intterm intterm + |Divides intterm intterm + |T + |F + |NOT QF + |And QF QF + |Or QF QF + |Imp QF QF + |Equ QF QF + |QAll QF + |QEx QF + +(* Interpretation of Presburger formulae *) +consts qinterp :: "int list \ QF \ bool" +primrec +"qinterp ats (Lt it1 it2) = (I_intterm ats it1 < I_intterm ats it2)" +"qinterp ats (Gt it1 it2) = (I_intterm ats it1 > I_intterm ats it2)" +"qinterp ats (Le it1 it2) = (I_intterm ats it1 \ I_intterm ats it2)" +"qinterp ats (Ge it1 it2) = (I_intterm ats it1 \ I_intterm ats it2)" +"qinterp ats (Divides it1 it2) = (I_intterm ats it1 dvd I_intterm ats it2)" +"qinterp ats (Eq it1 it2) = (I_intterm ats it1 = I_intterm ats it2)" +"qinterp ats T = True" +"qinterp ats F = False" +"qinterp ats (NOT p) = (\(qinterp ats p))" +"qinterp ats (And p q) = (qinterp ats p \ qinterp ats q)" +"qinterp ats (Or p q) = (qinterp ats p \ qinterp ats q)" +"qinterp ats (Imp p q) = (qinterp ats p \ qinterp ats q)" +"qinterp ats (Equ p q) = (qinterp ats p = qinterp ats q)" +"qinterp ats (QAll p) = (\x. qinterp (x#ats) p)" +"qinterp ats (QEx p) = (\x. qinterp (x#ats) p)" + +(* quantifier elimination based on qe, quantifier elimination for an + existential formula, with quantifier free body + Since quantifier elimination for one formula is allowed to fail, + the reult is of type QF option *) + +consts lift_bin:: "('a \ 'a \ 'b) \ 'a option \ 'a option \ 'b option" +recdef lift_bin "measure (\(c,a,b). size a)" +"lift_bin (c,Some a,Some b) = Some (c a b)" +"lift_bin (c,x, y) = None" + +lemma lift_bin_Some: + assumes ls: "lift_bin (c,x,y) = Some t" + shows "(\a. x = Some a) \ (\b. y = Some b)" + using ls + by (cases "x", auto) (cases "y", auto)+ + +consts lift_un:: "('a \ 'b) \ 'a option \ 'b option" +primrec +"lift_un c None = None" +"lift_un c (Some p) = Some (c p)" + +consts lift_qe:: "('a \ 'b option) \ 'a option \ 'b option" +primrec +"lift_qe qe None = None" +"lift_qe qe (Some p) = qe p" + +(* generic quantifier elimination *) +consts qelim :: "(QF \ QF option) \ QF \ QF option" +recdef qelim "measure (\(qe,p). size p)" +"qelim (qe, (QAll p)) = lift_un NOT (lift_qe qe (lift_un NOT (qelim (qe ,p))))" +"qelim (qe, (QEx p)) = lift_qe qe (qelim (qe,p))" +"qelim (qe, (And p q)) = lift_bin (And, (qelim (qe, p)), (qelim (qe, q)))" +"qelim (qe, (Or p q)) = lift_bin (Or, (qelim (qe, p)), (qelim (qe, q)))" +"qelim (qe, (Imp p q)) = lift_bin (Imp, (qelim (qe, p)), (qelim (qe, q)))" +"qelim (qe, (Equ p q)) = lift_bin (Equ, (qelim (qe, p)), (qelim (qe, q)))" +"qelim (qe,NOT p) = lift_un NOT (qelim (qe,p))" +"qelim (qe, p) = Some p" + +(* quantifier free-ness *) +consts isqfree :: "QF \ bool" +recdef isqfree "measure size" +"isqfree (QAll p) = False" +"isqfree (QEx p) = False" +"isqfree (And p q) = (isqfree p \ isqfree q)" +"isqfree (Or p q) = (isqfree p \ isqfree q)" +"isqfree (Imp p q) = (isqfree p \ isqfree q)" +"isqfree (Equ p q) = (isqfree p \ isqfree q)" +"isqfree (NOT p) = isqfree p" +"isqfree p = True" + +(* qelim lifts quantifierfreeness*) +lemma qelim_qfree: + assumes qeqf: "(\ q q'. \isqfree q ; qe q = Some q'\ \ isqfree q')" + shows qff:"\ p'. qelim (qe, p) = Some p' \ isqfree p'" + using qeqf +proof (induct p) + case (Lt a b) + have "qelim (qe, Lt a b) = Some (Lt a b)" by simp + moreover have "qelim (qe,Lt a b) = Some p'" . + ultimately have "p' = Lt a b" by simp + moreover have "isqfree (Lt a b)" by simp + ultimately + show ?case by simp +next + case (Gt a b) + have "qelim (qe, Gt a b) = Some (Gt a b)" by simp + moreover have "qelim (qe,Gt a b) = Some p'" . + ultimately have "p' = Gt a b" by simp + moreover have "isqfree (Gt a b)" by simp + ultimately + show ?case by simp +next + case (Le a b) + have "qelim (qe, Le a b) = Some (Le a b)" by simp + moreover have "qelim (qe,Le a b) = Some p'" . + ultimately have "p' = Le a b" by simp + moreover have "isqfree (Le a b)" by simp + ultimately + show ?case by simp +next + case (Ge a b) + have "qelim (qe, Ge a b) = Some (Ge a b)" by simp + moreover have "qelim (qe,Ge a b) = Some p'" . + ultimately have "p' = Ge a b" by simp + moreover have "isqfree (Ge a b)" by simp + ultimately + show ?case by simp +next + case (Eq a b) + have "qelim (qe, Eq a b) = Some (Eq a b)" by simp + moreover have "qelim (qe,Eq a b) = Some p'" . + ultimately have "p' = Eq a b" by simp + moreover have "isqfree (Eq a b)" by simp + ultimately + show ?case by simp +next + case (Divides a b) + have "qelim (qe, Divides a b) = Some (Divides a b)" by simp + moreover have "qelim (qe,Divides a b) = Some p'" . + ultimately have "p' = Divides a b" by simp + moreover have "isqfree (Divides a b)" by simp + ultimately + show ?case by simp +next + case T + have "qelim(qe,T) = Some T" by simp + moreover have "qelim(qe,T) = Some p'" . + ultimately have "p' = T" by simp + moreover have "isqfree T" by simp + ultimately show ?case by simp +next + case F + have "qelim(qe,F) = Some F" by simp + moreover have "qelim(qe,F) = Some p'" . + ultimately have "p' = F" by simp + moreover have "isqfree F" by simp + ultimately show ?case by simp +next + case (NOT p) + from "NOT.prems" have "\ p1. qelim(qe,p) = Some p1" + by (cases "qelim(qe,p)") simp_all + then obtain "p1" where p1_def: "qelim(qe,p) = Some p1" by blast + from "NOT.prems" have "\q q'. \isqfree q; qe q = Some q'\ \ isqfree q'" + by blast + with "NOT.hyps" p1_def have p1qf: "isqfree p1" by blast + then have "p' = NOT p1" using "NOT.prems" p1_def + by (cases "qelim(qe,NOT p)") simp_all + then show ?case using p1qf by simp +next + case (And p q) + from "And.prems" have p1q1: "(\p1. qelim(qe,p) = Some p1) \ + (\q1. qelim(qe,q) = Some q1)" using lift_bin_Some[where c="And"] by simp + from p1q1 obtain "p1" and "q1" + where p1_def: "qelim(qe,p) = Some p1" + and q1_def: "qelim(qe,q) = Some q1" by blast + from prems have qf1:"isqfree p1" + using p1_def by blast + from prems have qf2:"isqfree q1" + using q1_def by blast + from "And.prems" have "qelim(qe,And p q) = Some p'" by blast + then have "p' = And p1 q1" using p1_def q1_def by simp + then + show ?case using qf1 qf2 by simp +next + case (Or p q) + from "Or.prems" have p1q1: "(\p1. qelim(qe,p) = Some p1) \ + (\q1. qelim(qe,q) = Some q1)" using lift_bin_Some[where c="Or"] by simp + from p1q1 obtain "p1" and "q1" + where p1_def: "qelim(qe,p) = Some p1" + and q1_def: "qelim(qe,q) = Some q1" by blast + from prems have qf1:"isqfree p1" + using p1_def by blast + from prems have qf2:"isqfree q1" + using q1_def by blast + from "Or.prems" have "qelim(qe,Or p q) = Some p'" by blast + then have "p' = Or p1 q1" using p1_def q1_def by simp + then + show ?case using qf1 qf2 by simp +next + case (Imp p q) + from "Imp.prems" have p1q1: "(\p1. qelim(qe,p) = Some p1) \ + (\q1. qelim(qe,q) = Some q1)" using lift_bin_Some[where c="Imp"] by simp + from p1q1 obtain "p1" and "q1" + where p1_def: "qelim(qe,p) = Some p1" + and q1_def: "qelim(qe,q) = Some q1" by blast + from prems have qf1:"isqfree p1" + using p1_def by blast + from prems have qf2:"isqfree q1" + using q1_def by blast + from "Imp.prems" have "qelim(qe,Imp p q) = Some p'" by blast + then have "p' = Imp p1 q1" using p1_def q1_def by simp + then + show ?case using qf1 qf2 by simp +next + case (Equ p q) + from "Equ.prems" have p1q1: "(\p1. qelim(qe,p) = Some p1) \ + (\q1. qelim(qe,q) = Some q1)" using lift_bin_Some[where c="Equ"] by simp + from p1q1 obtain "p1" and "q1" + where p1_def: "qelim(qe,p) = Some p1" + and q1_def: "qelim(qe,q) = Some q1" by blast + from prems have qf1:"isqfree p1" + using p1_def by blast + from prems have qf2:"isqfree q1" + using q1_def by blast + from "Equ.prems" have "qelim(qe,Equ p q) = Some p'" by blast + then have "p' = Equ p1 q1" using p1_def q1_def by simp + then + show ?case using qf1 qf2 by simp +next + case (QEx p) + from "QEx.prems" have "\ p1. qelim(qe,p) = Some p1" + by (cases "qelim(qe,p)") simp_all + then obtain "p1" where p1_def: "qelim(qe,p) = Some p1" by blast + from "QEx.prems" have "\q q'. \isqfree q; qe q = Some q'\ \ isqfree q'" + by blast + with "QEx.hyps" p1_def have p1qf: "isqfree p1" by blast + from "QEx.prems" have "qe p1 = Some p'" using p1_def by simp + with "QEx.prems" show ?case using p1qf + by simp +next + case (QAll p) + from "QAll.prems" + have "\ p1. lift_qe qe (lift_un NOT (qelim (qe ,p))) = Some p1" + by (cases "lift_qe qe (lift_un NOT (qelim (qe ,p)))") simp_all + then obtain "p1" where + p1_def:"lift_qe qe (lift_un NOT (qelim (qe ,p))) = Some p1" by blast + then have "\ p2. lift_un NOT (qelim (qe ,p)) = Some p2" + by (cases "qelim (qe ,p)") simp_all + then obtain "p2" + where p2_def:"lift_un NOT (qelim (qe ,p)) = Some p2" by blast + then have "\ p3. qelim(qe,p) = Some p3" by (cases "qelim(qe,p)") simp_all + then obtain "p3" where p3_def: "qelim(qe,p) = Some p3" by blast + with prems have qf3: "isqfree p3" by blast + have p2_def2: "p2 = NOT p3" using p2_def p3_def by simp + then have qf2: "isqfree p2" using qf3 by simp + have p1_edf2: "qe p2 = Some p1" using p1_def p2_def by simp + with "QAll.prems" have qf1: "isqfree p1" using qf2 by blast + from "QAll.prems" have "p' = NOT p1" using p1_def by simp + with qf1 show ?case by simp +qed + +(* qelim lifts semantical equivalence *) +lemma qelim_corr: + assumes qecorr: "(\ q q' ats. \isqfree q ; qe q = Some q'\ \ (qinterp ats (QEx q)) = (qinterp ats q'))" + and qeqf: "(\ q q'. \isqfree q ; qe q = Some q'\ \ isqfree q')" + shows qff:"\ p' ats. qelim (qe, p) = Some p' \ (qinterp ats p = qinterp ats p')" (is "\ p' ats. ?Qe p p' \ (?F ats p = ?F ats p')") + using qeqf qecorr +proof (induct p) + case (NOT f) + from "NOT.prems" have "\f'. ?Qe f f' " by (cases "qelim(qe,f)") simp_all + then obtain "f'" where df': "?Qe f f'" by blast + with prems have feqf': "?F ats f = ?F ats f'" by blast + from "NOT.prems" df' have "p' = NOT f'" by simp + with feqf' show ?case by simp + +next + case (And f g) + from "And.prems" have f1g1: "(\f1. qelim(qe,f) = Some f1) \ + (\g1. qelim(qe,g) = Some g1)" using lift_bin_Some[where c="And"] by simp + from f1g1 obtain "f1" and "g1" + where f1_def: "qelim(qe, f) = Some f1" + and g1_def: "qelim(qe,g) = Some g1" by blast + from prems f1_def have feqf1: "?F ats f = ?F ats f1" by blast + from prems g1_def have geqg1: "?F ats g = ?F ats g1" by blast + from "And.prems" f1_def g1_def have "p' = And f1 g1" by simp + with feqf1 geqg1 show ?case by simp + +next + case (Or f g) + from "Or.prems" have f1g1: "(\f1. qelim(qe,f) = Some f1) \ + (\g1. qelim(qe,g) = Some g1)" using lift_bin_Some[where c="Or"] by simp + from f1g1 obtain "f1" and "g1" + where f1_def: "qelim(qe, f) = Some f1" + and g1_def: "qelim(qe,g) = Some g1" by blast + from prems f1_def have feqf1: "?F ats f = ?F ats f1" by blast + from prems g1_def have geqg1: "?F ats g = ?F ats g1" by blast + from "Or.prems" f1_def g1_def have "p' = Or f1 g1" by simp + with feqf1 geqg1 show ?case by simp +next + case (Imp f g) + from "Imp.prems" have f1g1: "(\f1. qelim(qe,f) = Some f1) \ + (\g1. qelim(qe,g) = Some g1)" using lift_bin_Some[where c="Imp"] by simp + from f1g1 obtain "f1" and "g1" + where f1_def: "qelim(qe, f) = Some f1" + and g1_def: "qelim(qe,g) = Some g1" by blast + from prems f1_def have feqf1: "?F ats f = ?F ats f1" by blast + from prems g1_def have geqg1: "?F ats g = ?F ats g1" by blast + from "Imp.prems" f1_def g1_def have "p' = Imp f1 g1" by simp + with feqf1 geqg1 show ?case by simp +next + case (Equ f g) + from "Equ.prems" have f1g1: "(\f1. qelim(qe,f) = Some f1) \ + (\g1. qelim(qe,g) = Some g1)" using lift_bin_Some[where c="Equ"] by simp + from f1g1 obtain "f1" and "g1" + where f1_def: "qelim(qe, f) = Some f1" + and g1_def: "qelim(qe,g) = Some g1" by blast + from prems f1_def have feqf1: "?F ats f = ?F ats f1" by blast + from prems g1_def have geqg1: "?F ats g = ?F ats g1" by blast + from "Equ.prems" f1_def g1_def have "p' = Equ f1 g1" by simp + with feqf1 geqg1 show ?case by simp +next + case (QEx f) + from "QEx.prems" have "\ f1. ?Qe f f1" + by (cases "qelim(qe,f)") simp_all + then obtain "f1" where f1_def: "qelim(qe,f) = Some f1" by blast + from prems have qf1:"isqfree f1" using qelim_qfree by blast + from prems have feqf1: "\ ats. qinterp ats f = qinterp ats f1" + using f1_def qf1 by blast + then have "?F ats (QEx f) = ?F ats (QEx f1)" + by simp + from prems have "qelim (qe,QEx f) = Some p'" by blast + then have "\ f'. qe f1 = Some f'" using f1_def by simp + then obtain "f'" where fdef': "qe f1 = Some f'" by blast + with prems have exf1: "?F ats (QEx f1) = ?F ats f'" using qf1 by blast + have fp: "?Qe (QEx f) f'" using f1_def fdef' by simp + from prems have "?Qe (QEx f) p'" by blast + then have "p' = f'" using fp by simp + then show ?case using feqf1 exf1 by simp +next + case (QAll f) + from "QAll.prems" + have "\ f0. lift_un NOT (lift_qe qe (lift_un NOT (qelim (qe ,f)))) = + Some f0" + by (cases "lift_un NOT (lift_qe qe (lift_un NOT (qelim (qe ,f))))") + simp_all + then obtain "f0" + where f0_def: "lift_un NOT (lift_qe qe (lift_un NOT (qelim (qe ,f)))) = + Some f0" by blast + then have "\ f1. lift_qe qe (lift_un NOT (qelim (qe ,f))) = Some f1" + by (cases "lift_qe qe (lift_un NOT (qelim (qe ,f)))") simp_all + then obtain "f1" where + f1_def:"lift_qe qe (lift_un NOT (qelim (qe ,f))) = Some f1" by blast + then have "\ f2. lift_un NOT (qelim (qe ,f)) = Some f2" + by (cases "qelim (qe ,f)") simp_all + then obtain "f2" + where f2_def:"lift_un NOT (qelim (qe ,f)) = Some f2" by blast + then have "\ f3. qelim(qe,f) = Some f3" by (cases "qelim(qe,f)") simp_all + then obtain "f3" where f3_def: "qelim(qe,f) = Some f3" by blast + from prems have qf3:"isqfree f3" using qelim_qfree by blast + from prems have feqf3: "\ ats. qinterp ats f = qinterp ats f3" + using f3_def qf3 by blast + have f23:"f2 = NOT f3" using f2_def f3_def by simp + then have feqf2: "\ ats. qinterp ats f = qinterp ats (NOT f2)" + using feqf3 by simp + have qf2: "isqfree f2" using f23 qf3 by simp + have "qe f2 = Some f1" using f1_def f2_def f23 by simp + with prems have exf2eqf1: "?F ats (QEx f2) = ?F ats f1" using qf2 by blast + have "f0 = NOT f1" using f0_def f1_def by simp + then have f0eqf1: "?F ats f0 = ?F ats (NOT f1)" by simp + from prems have "qelim (qe, QAll f) = Some p'" by blast + then have f0eqp': "p' = f0" using f0_def by simp + have "?F ats (QAll f) = (\x. ?F (x#ats) f)" by simp + also have "\ = (\ (\ x. ?F (x#ats) (NOT f)))" by simp + also have "\ = (\ (\ x. ?F (x#ats) (NOT (NOT f2))))" using feqf2 + by auto + also have "\ = (\ (\ x. ?F (x#ats) f2))" by simp + also have "\ = (\ (?F ats f1))" using exf2eqf1 by simp + finally show ?case using f0eqp' f0eqf1 by simp +qed simp_all + +(* Cooper's algorithm *) + + +(* Transform an intform into NNF *) +consts lgth :: "QF \ nat" + nnf :: "QF \ QF" +primrec +"lgth (Lt it1 it2) = 1" +"lgth (Gt it1 it2) = 1" +"lgth (Le it1 it2) = 1" +"lgth (Ge it1 it2) = 1" +"lgth (Eq it1 it2) = 1" +"lgth (Divides it1 it2) = 1" +"lgth T = 1" +"lgth F = 1" +"lgth (NOT p) = 1 + lgth p" +"lgth (And p q) = 1 + lgth p + lgth q" +"lgth (Or p q) = 1 + lgth p + lgth q" +"lgth (Imp p q) = 1 + lgth p + lgth q" +"lgth (Equ p q) = 1 + lgth p + lgth q" +"lgth (QAll p) = 1 + lgth p" +"lgth (QEx p) = 1 + lgth p" + +lemma [simp] :"0 < lgth q" +apply (induct_tac q) +apply(auto) +done + +(* NNF *) +recdef nnf "measure (\p. lgth p)" + "nnf (Lt it1 it2) = Le (Sub it1 it2) (Cst (- 1))" + "nnf (Gt it1 it2) = Le (Sub it2 it1) (Cst (- 1))" + "nnf (Le it1 it2) = Le it1 it2 " + "nnf (Ge it1 it2) = Le it2 it1" + "nnf (Eq it1 it2) = Eq it2 it1" + "nnf (Divides d t) = Divides d t" + "nnf T = T" + "nnf F = F" + "nnf (And p q) = And (nnf p) (nnf q)" + "nnf (Or p q) = Or (nnf p) (nnf q)" + "nnf (Imp p q) = Or (nnf (NOT p)) (nnf q)" + "nnf (Equ p q) = Or (And (nnf p) (nnf q)) + (And (nnf (NOT p)) (nnf (NOT q)))" + "nnf (NOT (Lt it1 it2)) = (Le it2 it1)" + "nnf (NOT (Gt it1 it2)) = (Le it1 it2)" + "nnf (NOT (Le it1 it2)) = (Le (Sub it2 it1) (Cst (- 1)))" + "nnf (NOT (Ge it1 it2)) = (Le (Sub it1 it2) (Cst (- 1)))" + "nnf (NOT (Eq it1 it2)) = (NOT (Eq it1 it2))" + "nnf (NOT (Divides d t)) = (NOT (Divides d t))" + "nnf (NOT T) = F" + "nnf (NOT F) = T" + "nnf (NOT (NOT p)) = (nnf p)" + "nnf (NOT (And p q)) = (Or (nnf (NOT p)) (nnf (NOT q)))" + "nnf (NOT (Or p q)) = (And (nnf (NOT p)) (nnf (NOT q)))" + "nnf (NOT (Imp p q)) = (And (nnf p) (nnf (NOT q)))" + "nnf (NOT (Equ p q)) = (Or (And (nnf p) (nnf (NOT q))) (And (nnf (NOT p)) (nnf q)))" + +consts isnnf :: "QF \ bool" +recdef isnnf "measure (\p. lgth p)" + "isnnf (Le it1 it2) = True" + "isnnf (Eq it1 it2) = True" + "isnnf (Divides d t) = True" + "isnnf T = True" + "isnnf F = True" + "isnnf (And p q) = (isnnf p \ isnnf q)" + "isnnf (Or p q) = (isnnf p \ isnnf q)" + "isnnf (NOT (Divides d t)) = True" + "isnnf (NOT (Eq it1 it2)) = True" + "isnnf p = False" + +(* nnf preserves semantics *) +lemma nnf_corr: "isqfree p \ qinterp ats p = qinterp ats (nnf p)" +by (induct p rule: nnf.induct,simp_all) +(arith, arith, arith, arith, arith, arith, arith, arith, arith, blast) + + +(* the result of nnf is in NNF *) +lemma nnf_isnnf : "isqfree p \ isnnf (nnf p)" +by (induct p rule: nnf.induct, auto) + +lemma nnf_isqfree: "isnnf p \ isqfree p" +by (induct p rule: isnnf.induct) auto + +(* nnf preserves quantifier freeness *) +lemma nnf_qfree: "isqfree p \ isqfree(nnf p)" + using nnf_isqfree nnf_isnnf by simp + +(* Linearization and normalization of formulae *) +(* Definition of linearity of an intterm *) + +consts islinintterm :: "intterm \ bool" +recdef islinintterm "measure size" +"islinintterm (Cst i) = True" +"islinintterm (Add (Mult (Cst i) (Var n)) (Cst i')) = (i \ 0)" +"islinintterm (Add (Mult (Cst i) (Var n)) (Add (Mult (Cst i') (Var n')) r)) = ( i \ 0 \ i' \ 0 \ n < n' \ islinintterm (Add (Mult (Cst i') (Var n')) r))" +"islinintterm i = False" + +(* subterms of linear terms are linear *) +lemma islinintterm_subt: + assumes lr: "islinintterm (Add (Mult (Cst i) (Var n)) r)" + shows "islinintterm r" +using lr +by (induct r rule: islinintterm.induct) auto + +(* c \ 0 for linear term c.n + r*) +lemma islinintterm_cnz: + assumes lr: "islinintterm (Add (Mult (Cst i) (Var n)) r)" + shows "i \ 0" +using lr +by (induct r rule: islinintterm.induct) auto + +lemma islininttermc0r: "islinintterm (Add (Mult (Cst c) (Var n)) r) \ (c \ 0 \ islinintterm r)" +by (induct r rule: islinintterm.induct, simp_all) + +(* An alternative linearity definition *) +consts islintn :: "(nat \ intterm) \ bool" +recdef islintn "measure (\ (n,t). (size t))" +"islintn (n0, Cst i) = True" +"islintn (n0, Add (Mult (Cst i) (Var n)) r) = (i \ 0 \ n0 \ n \ islintn (n+1,r))" +"islintn (n0, t) = False" + +constdefs islint :: "intterm \ bool" + "islint t \ islintn(0,t)" + +(* And the equivalence to the first definition *) +lemma islinintterm_eq_islint: "islinintterm t = islint t" + using islint_def +by (induct t rule: islinintterm.induct) auto + +(* monotony *) +lemma islintn_mon: + assumes lin: "islintn (n,t)" + and mgen: "m \ n" + shows "islintn(m,t)" + using lin mgen +by (induct t rule: islintn.induct) auto + +lemma islintn_subt: + assumes lint: "islintn(n,Add (Mult (Cst i) (Var m)) r)" + shows "islintn (m+1,r)" +using lint +by auto + +(* List indexin for n > 0 *) +lemma nth_pos: "0 < n \ (x#xs) ! n = (y#xs) ! n" +using Nat.gr0_conv_Suc +by clarsimp + +lemma nth_pos2: "0 < n \ (x#xs) ! n = xs ! (n - 1)" +using Nat.gr0_conv_Suc +by clarsimp + +lemma intterm_novar0: + assumes lin: "islinintterm (Add (Mult (Cst i) (Var n)) r)" + shows "I_intterm (x#ats) r = I_intterm (y#ats) r" +using lin +by (induct r rule: islinintterm.induct) (simp_all add: nth_pos2) +(* a simple version of a general theorem: Interpretation doese not depend + on the first variable if it doese not occur in the term *) + +lemma linterm_novar0: + assumes lin: "islintn (n,t)" + and npos: "0 < n" + shows "I_intterm (x#ats) t = I_intterm (y#ats) t" +using lin npos +by (induct n t rule: islintn.induct) (simp_all add: nth_pos2) + +(* Periodicity of dvd *) +lemma dvd_period: + assumes advdd: "(a::int) dvd d" + shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))" +using advdd +proof- + from advdd have "\x.\k. (((a::int) dvd (x + t)) = (a dvd + (x+k*d + t)))" by (rule dvd_modd_pinf) + then show ?thesis by simp +qed + +(* lin_ad adds two linear terms*) +consts lin_add :: "intterm \ intterm \ intterm" +recdef lin_add "measure (\(x,y). ((size x) + (size y)))" +"lin_add (Add (Mult (Cst c1) (Var n1)) (r1),Add (Mult (Cst c2) (Var n2)) (r2)) = + (if n1=n2 then + (let c = Cst (c1 + c2) + in (if c1+c2=0 then lin_add(r1,r2) else Add (Mult c (Var n1)) (lin_add (r1,r2)))) + else if n1 \ n2 then (Add (Mult (Cst c1) (Var n1)) (lin_add (r1,Add (Mult (Cst c2) (Var n2)) (r2)))) + else (Add (Mult (Cst c2) (Var n2)) (lin_add (Add (Mult (Cst c1) (Var n1)) r1,r2))))" +"lin_add (Add (Mult (Cst c1) (Var n1)) (r1),Cst b) = + (Add (Mult (Cst c1) (Var n1)) (lin_add (r1, Cst b)))" +"lin_add (Cst x,Add (Mult (Cst c2) (Var n2)) (r2)) = + Add (Mult (Cst c2) (Var n2)) (lin_add (Cst x,r2))" +"lin_add (Cst b1, Cst b2) = Cst (b1+b2)" + +lemma lin_add_cst_corr: + assumes blin : "islintn(n0,b)" + shows "I_intterm ats (lin_add (Cst a,b)) = (I_intterm ats (Add (Cst a) b))" +using blin +by (induct n0 b rule: islintn.induct) auto + +lemma lin_add_cst_corr2: + assumes blin : "islintn(n0,b)" + shows "I_intterm ats (lin_add (b,Cst a)) = (I_intterm ats (Add b (Cst a)))" +using blin +by (induct n0 b rule: islintn.induct) auto + +lemma lin_add_corrh: "\ n01 n02. \ islintn (n01,a) ; islintn (n02,b)\ + \ I_intterm ats (lin_add(a,b)) = I_intterm ats (Add a b)" +proof(induct a b rule: lin_add.induct) + case (58 i n r j m s) + have "(n = m \ i+j = 0) \ (n = m \ i+j \ 0) \ n < m \ m < n " by arith + moreover + {assume "n=m\i+j=0" hence ?case using prems by (auto simp add: sym[OF zadd_zmult_distrib]) } + moreover + {assume "n=m\i+j\0" hence ?case using prems by (auto simp add: Let_def zadd_zmult_distrib)} + moreover + {assume "n < m" hence ?case using prems by auto } + moreover + {assume "n > m" hence ?case using prems by auto } + ultimately show ?case by blast +qed (auto simp add: lin_add_cst_corr lin_add_cst_corr2 Let_def) + +(* lin_add has the semantics of Add*) +lemma lin_add_corr: + assumes lina: "islinintterm a" + and linb: "islinintterm b" + shows "I_intterm ats (lin_add (a,b)) = (I_intterm ats (Add a b))" +using lina linb islinintterm_eq_islint islint_def lin_add_corrh +by blast + +lemma lin_add_cst_lint: + assumes lin: "islintn (n0,b)" + shows "islintn (n0, lin_add (Cst i, b))" +using lin +by (induct n0 b rule: islintn.induct) auto + +lemma lin_add_cst_lint2: + assumes lin: "islintn (n0,b)" + shows "islintn (n0, lin_add (b,Cst i))" +using lin +by (induct n0 b rule: islintn.induct) auto + +(* lin_add preserves linearity..*) +lemma lin_add_lint: "\ n0 n01 n02. \ islintn (n01,a) ; islintn (n02,b); n0 \ min n01 n02 \ + \ islintn (n0, lin_add (a,b))" +proof (induct a b rule: lin_add.induct) + case (58 i n r j m s) + have "(n =m \ i + j = 0) \ (n = m \ i+j \ 0) \ n m < n" by arith + moreover + { assume "n = m" + and "i+j = 0" + hence ?case using "58" islintn_mon[where m = "n01" and n = "Suc m"] + islintn_mon[where m = "n02" and n = "Suc m"] by auto } + moreover + { assume "n = m" + and "i+j \ 0" + hence ?case using "58" islintn_mon[where m = "n01" and n = "Suc m"] + islintn_mon[where m = "n02" and n = "Suc m"] by (auto simp add: Let_def) } + moreover + { assume "n < m" hence ?case using 58 by force } +moreover + { assume "m < n" + hence ?case using 58 + apply (auto simp add: Let_def) + apply (erule allE[where x = "Suc m" ] ) + by (erule allE[where x = "Suc m" ] ) simp } + ultimately show ?case by blast +qed(simp_all add: Let_def lin_add_cst_lint lin_add_cst_lint2) + +lemma lin_add_lin: + assumes lina: "islinintterm a" + and linb: "islinintterm b" + shows "islinintterm (lin_add (a,b))" +using islinintterm_eq_islint islint_def lin_add_lint lina linb by auto + +(* lin_mul multiplies a linear term by a constant *) +consts lin_mul :: "int \ intterm \ intterm" +recdef lin_mul "measure (\(c,t). size t)" +"lin_mul (c,Cst i) = (Cst (c*i))" +"lin_mul (c,Add (Mult (Cst c') (Var n)) r) = + (if c = 0 then (Cst 0) else + (Add (Mult (Cst (c*c')) (Var n)) (lin_mul (c,r))))" + +lemma zmult_zadd_distrib[simp]: "(a::int) * (b+c) = a*b + a*c" +proof- + have "a*(b+c) = (b+c)*a" by simp + moreover have "(b+c)*a = b*a + c*a" by (simp add: zadd_zmult_distrib) + ultimately show ?thesis by simp +qed + +(* lin_mul has the semantics of Mult *) +lemma lin_mul_corr: + assumes lint: "islinintterm t" + shows "I_intterm ats (lin_mul (c,t)) = I_intterm ats (Mult (Cst c) t)" +using lint +proof (induct c t rule: lin_mul.induct) + case (21 c c' n r) + have "islinintterm (Add (Mult (Cst c') (Var n)) r)" . + then have "islinintterm r" + by (rule islinintterm_subt[of "c'" "n" "r"]) + then show ?case using "21.hyps" "21.prems" by simp +qed(auto) + +(* lin_mul preserves linearity *) +lemma lin_mul_lin: + assumes lint: "islinintterm t" + shows "islinintterm (lin_mul(c,t))" +using lint +by (induct t rule: islinintterm.induct) auto + +lemma lin_mul0: + assumes lint: "islinintterm t" + shows "lin_mul(0,t) = Cst 0" + using lint + by (induct t rule: islinintterm.induct) auto + +lemma lin_mul_lintn: + "\ m. islintn(m,t) \ islintn(m,lin_mul(l,t))" + by (induct l t rule: lin_mul.induct) simp_all + +(* lin_neg neagtes a linear term *) +constdefs lin_neg :: "intterm \ intterm" +"lin_neg i == lin_mul ((-1::int),i)" + +(* lin_neg has the semantics of Neg *) +lemma lin_neg_corr: + assumes lint: "islinintterm t" + shows "I_intterm ats (lin_neg t) = I_intterm ats (Neg t)" + using lint lin_mul_corr + by (simp add: lin_neg_def lin_mul_corr) + +(* lin_neg preserves linearity *) +lemma lin_neg_lin: + assumes lint: "islinintterm t" + shows "islinintterm (lin_neg t)" +using lint +by (simp add: lin_mul_lin lin_neg_def) + +(* Some properties about lin_add and lin-neg should be moved above *) + +lemma lin_neg_idemp: + assumes lini: "islinintterm i" + shows "lin_neg (lin_neg i) = i" +using lini +by (induct i rule: islinintterm.induct) (auto simp add: lin_neg_def) + +lemma lin_neg_lin_add_distrib: + assumes lina : "islinintterm a" + and linb :"islinintterm b" + shows "lin_neg (lin_add(a,b)) = lin_add (lin_neg a, lin_neg b)" +using lina linb +proof (induct a b rule: lin_add.induct) + case (58 c1 n1 r1 c2 n2 r2) + from prems have lincnr1:"islinintterm (Add (Mult (Cst c1) (Var n1)) r1)" by simp + have linr1: "islinintterm r1" by (rule islinintterm_subt[OF lincnr1]) + from prems have lincnr2: "islinintterm (Add (Mult (Cst c2) (Var n2)) r2)" by simp + have linr2: "islinintterm r2" by (rule islinintterm_subt[OF lincnr2]) + have "n1 = n2 \ n1 < n2 \ n1 > n2" by arith + show ?case using prems linr1 linr2 by (simp_all add: lin_neg_def Let_def) +next + case (59 c n r b) + from prems have lincnr: "islinintterm (Add (Mult (Cst c) (Var n)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + show ?case using prems linr by (simp add: lin_neg_def Let_def) +next + case (60 b c n r) + from prems have lincnr: "islinintterm (Add (Mult (Cst c) (Var n)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + show ?case using prems linr by (simp add: lin_neg_def Let_def) +qed (simp_all add: lin_neg_def) + +(* linearize tries to linearize a term *) +consts linearize :: "intterm \ intterm option" +recdef linearize "measure (\t. size t)" +"linearize (Cst b) = Some (Cst b)" +"linearize (Var n) = Some (Add (Mult (Cst 1) (Var n)) (Cst 0))" +"linearize (Neg i) = lift_un lin_neg (linearize i)" + "linearize (Add i j) = lift_bin(\ x. \ y. lin_add(x,y), linearize i, linearize j)" +"linearize (Sub i j) = + lift_bin(\ x. \ y. lin_add(x,lin_neg y), linearize i, linearize j)" +"linearize (Mult i j) = + (case linearize i of + None \ None + | Some li \ (case li of + Cst b \ (case linearize j of + None \ None + | (Some lj) \ Some (lin_mul(b,lj))) + | _ \ (case linearize j of + None \ None + | (Some lj) \ (case lj of + Cst b \ Some (lin_mul (b,li)) + | _ \ None))))" + +lemma linearize_linear1: + assumes lin: "linearize t \ None" + shows "islinintterm (the (linearize t))" +using lin +proof (induct t rule: linearize.induct) + case (1 b) show ?case by simp +next + case (2 n) show ?case by simp +next + case (3 i) show ?case + proof- + have "(linearize i = None) \ (\li. linearize i = Some li)" by auto + moreover + { assume "linearize i = None" with prems have ?thesis by auto} + moreover + { assume lini: "\li. linearize i = Some li" + from lini obtain "li" where "linearize i = Some li" by blast + have linli: "islinintterm li" by (simp!) + moreover have "linearize (Neg i) = Some (lin_neg li)" using prems by simp + moreover from linli have "islinintterm(lin_neg li)" by (simp add: lin_neg_lin) + ultimately have ?thesis by simp + } + ultimately show ?thesis by blast + qed +next + case (4 i j) show ?case + proof- + have "(linearize i = None) \ ((\ li. linearize i = Some li) \ linearize j = None) \ ((\ li. linearize i = Some li) \ (\ lj. linearize j = Some lj))" by auto + moreover + { + assume nlini: "linearize i = None" + from nlini have "linearize (Add i j) = None" + by (simp add: Let_def measure_def inv_image_def) then have ?thesis using prems by auto} + moreover + { assume nlinj: "linearize j = None" + and lini: "\ li. linearize i = Some li" + from nlinj lini have "linearize (Add i j) = None" + by (simp add: Let_def measure_def inv_image_def, auto) with prems have ?thesis by auto} + moreover + { assume lini: "\li. linearize i = Some li" + and linj: "\lj. linearize j = Some lj" + from lini obtain "li" where "linearize i = Some li" by blast + have linli: "islinintterm li" by (simp!) + from linj obtain "lj" where "linearize j = Some lj" by blast + have linlj: "islinintterm lj" by (simp!) + moreover from lini linj have "linearize (Add i j) = Some (lin_add (li,lj))" + by (simp add: measure_def inv_image_def, auto!) + moreover from linli linlj have "islinintterm(lin_add (li,lj))" by (simp add: lin_add_lin) + ultimately have ?thesis by simp } + ultimately show ?thesis by blast + qed +next + case (5 i j)show ?case + proof- + have "(linearize i = None) \ ((\ li. linearize i = Some li) \ linearize j = None) \ ((\ li. linearize i = Some li) \ (\ lj. linearize j = Some lj))" by auto + moreover + { + assume nlini: "linearize i = None" + from nlini have "linearize (Sub i j) = None" by (simp add: Let_def measure_def inv_image_def) then have ?thesis by (auto!) + } + moreover + { + assume lini: "\ li. linearize i = Some li" + and nlinj: "linearize j = None" + from nlinj lini have "linearize (Sub i j) = None" + by (simp add: Let_def measure_def inv_image_def, auto) then have ?thesis by (auto!) + } + moreover + { + assume lini: "\li. linearize i = Some li" + and linj: "\lj. linearize j = Some lj" + from lini obtain "li" where "linearize i = Some li" by blast + have linli: "islinintterm li" by (simp!) + from linj obtain "lj" where "linearize j = Some lj" by blast + have linlj: "islinintterm lj" by (simp!) + moreover from lini linj have "linearize (Sub i j) = Some (lin_add (li,lin_neg lj))" + by (simp add: measure_def inv_image_def, auto!) + moreover from linli linlj have "islinintterm(lin_add (li,lin_neg lj))" by (simp add: lin_add_lin lin_neg_lin) + ultimately have ?thesis by simp + } + ultimately show ?thesis by blast + qed +next + case (6 i j)show ?case + proof- + have cses: "(linearize i = None) \ + ((\ li. linearize i = Some li) \ linearize j = None) \ + ((\ li. linearize i = Some li) \ (\ bj. linearize j = Some (Cst bj))) + \ ((\ bi. linearize i = Some (Cst bi)) \ (\ lj. linearize j = Some lj)) + \ ((\ li. linearize i = Some li \ \ (\ bi. li = Cst bi)) \ (\ lj. linearize j = Some lj \ \ (\ bj. lj = Cst bj)))" by auto + moreover + { + assume nlini: "linearize i = None" + from nlini have "linearize (Mult i j) = None" + by (simp add: Let_def measure_def inv_image_def) + with prems have ?thesis by auto } + moreover + { assume lini: "\ li. linearize i = Some li" + and nlinj: "linearize j = None" + from lini obtain "li" where "linearize i = Some li" by blast + moreover from nlinj lini have "linearize (Mult i j) = None" + using prems + by (cases li ) (auto simp add: Let_def measure_def inv_image_def) + with prems have ?thesis by auto} + moreover + { assume lini: "\li. linearize i = Some li" + and linj: "\bj. linearize j = Some (Cst bj)" + from lini obtain "li" where li_def: "linearize i = Some li" by blast + from prems have linli: "islinintterm li" by simp + moreover + from linj obtain "bj" where bj_def: "linearize j = Some (Cst bj)" by blast + have linlj: "islinintterm (Cst bj)" by simp + moreover from lini linj prems + have "linearize (Mult i j) = Some (lin_mul (bj,li))" + by (cases li) (auto simp add: measure_def inv_image_def) + moreover from linli linlj have "islinintterm(lin_mul (bj,li))" by (simp add: lin_mul_lin) + ultimately have ?thesis by simp } + moreover + { assume lini: "\bi. linearize i = Some (Cst bi)" + and linj: "\lj. linearize j = Some lj" + from lini obtain "bi" where "linearize i = Some (Cst bi)" by blast + from prems have linli: "islinintterm (Cst bi)" by simp + moreover + from linj obtain "lj" where "linearize j = Some lj" by blast + from prems have linlj: "islinintterm lj" by simp + moreover from lini linj prems have "linearize (Mult i j) = Some (lin_mul (bi,lj))" + by (simp add: measure_def inv_image_def) + moreover from linli linlj have "islinintterm(lin_mul (bi,lj))" by (simp add: lin_mul_lin) + ultimately have ?thesis by simp } + moreover + { assume linc: "\ li. linearize i = Some li \ \ (\ bi. li = Cst bi)" + and ljnc: "\ lj. linearize j = Some lj \ \ (\ bj. lj = Cst bj)" + from linc obtain "li" where "linearize i = Some li \ \ (\ bi. li = Cst bi)" by blast + moreover + from ljnc obtain "lj" where "linearize j = Some lj \ \ (\ bj. lj = Cst bj)" by blast + ultimately have "linearize (Mult i j) = None" + by (cases li, auto simp add: measure_def inv_image_def) (cases lj, auto)+ + with prems have ?thesis by simp } + ultimately show ?thesis by blast + qed +qed + +(* the result of linearize, when successful, is a linear term*) +lemma linearize_linear: "\ t'. linearize t = Some t' \ islinintterm t'" +proof- + fix t' + assume lint: "linearize t = Some t'" + from lint have lt: "linearize t \ None" by auto + then have "islinintterm (the (linearize t))" by (rule_tac linearize_linear1[OF lt]) + with lint show "islinintterm t'" by simp +qed + +lemma linearize_corr1: + assumes lin: "linearize t \ None" + shows "I_intterm ats t = I_intterm ats (the (linearize t))" +using lin +proof (induct t rule: linearize.induct) + case (3 i) show ?case + proof- + have "(linearize i = None) \ (\li. linearize i = Some li)" by auto + moreover + { + assume "linearize i = None" + have ?thesis using prems by simp + } + moreover + { + assume lini: "\li. linearize i = Some li" + from lini have lini2: "linearize i \ None" by simp + from lini obtain "li" where "linearize i = Some li" by blast + from lini2 lini have "islinintterm (the (linearize i))" + by (simp add: linearize_linear1[OF lini2]) + then have linli: "islinintterm li" using prems by simp + have ieqli: "I_intterm ats i = I_intterm ats li" using prems by simp + moreover have "linearize (Neg i) = Some (lin_neg li)" using prems by simp + moreover from ieqli linli have "I_intterm ats (Neg i) = I_intterm ats (lin_neg li)" by (simp add: lin_neg_corr[OF linli]) + ultimately have ?thesis using prems by (simp add: lin_neg_corr) + } + ultimately show ?thesis by blast + qed +next + case (4 i j) show ?case + proof- + have "(linearize i = None) \ ((\ li. linearize i = Some li) \ linearize j = None) \ ((\ li. linearize i = Some li) \ (\ lj. linearize j = Some lj))" by auto + moreover + { + assume nlini: "linearize i = None" + from nlini have "linearize (Add i j) = None" by (simp add: Let_def measure_def inv_image_def) then have ?thesis using prems by auto + } + moreover + { + assume nlinj: "linearize j = None" + and lini: "\ li. linearize i = Some li" + from nlinj lini have "linearize (Add i j) = None" + by (simp add: Let_def measure_def inv_image_def, auto) + then have ?thesis using prems by auto + } + moreover + { + assume lini: "\li. linearize i = Some li" + and linj: "\lj. linearize j = Some lj" + from lini have lini2: "linearize i \ None" by simp + from linj have linj2: "linearize j \ None" by simp + from lini obtain "li" where "linearize i = Some li" by blast + from lini2 have "islinintterm (the (linearize i))" by (simp add: linearize_linear1) + then have linli: "islinintterm li" using prems by simp + from linj obtain "lj" where "linearize j = Some lj" by blast + from linj2 have "islinintterm (the (linearize j))" by (simp add: linearize_linear1) + then have linlj: "islinintterm lj" using prems by simp + moreover from lini linj have "linearize (Add i j) = Some (lin_add (li,lj))" + using prems by (simp add: measure_def inv_image_def) + moreover from linli linlj have "I_intterm ats (lin_add (li,lj)) = I_intterm ats (Add li lj)" by (simp add: lin_add_corr) + ultimately have ?thesis using prems by simp + } + ultimately show ?thesis by blast + qed +next + case (5 i j)show ?case + proof- + have "(linearize i = None) \ ((\ li. linearize i = Some li) \ linearize j = None) \ ((\ li. linearize i = Some li) \ (\ lj. linearize j = Some lj))" by auto + moreover + { + assume nlini: "linearize i = None" + from nlini have "linearize (Sub i j) = None" by (simp add: Let_def measure_def inv_image_def) then have ?thesis using prems by auto + } + moreover + { + assume lini: "\ li. linearize i = Some li" + and nlinj: "linearize j = None" + from nlinj lini have "linearize (Sub i j) = None" + by (simp add: Let_def measure_def inv_image_def, auto) with prems have ?thesis by auto + } + moreover + { + assume lini: "\li. linearize i = Some li" + and linj: "\lj. linearize j = Some lj" + from lini have lini2: "linearize i \ None" by simp + from linj have linj2: "linearize j \ None" by simp + from lini obtain "li" where "linearize i = Some li" by blast + from lini2 have "islinintterm (the (linearize i))" by (simp add: linearize_linear1) + with prems have linli: "islinintterm li" by simp + from linj obtain "lj" where "linearize j = Some lj" by blast + from linj2 have "islinintterm (the (linearize j))" by (simp add: linearize_linear1) + with prems have linlj: "islinintterm lj" by simp + moreover from prems have "linearize (Sub i j) = Some (lin_add (li,lin_neg lj))" + by (simp add: measure_def inv_image_def) + moreover from linlj have linnlj:"islinintterm (lin_neg lj)" by (simp add: lin_neg_lin) + moreover from linli linnlj have "I_intterm ats (lin_add (li,lin_neg lj)) = I_intterm ats (Add li (lin_neg lj))" by (simp only: lin_add_corr[OF linli linnlj]) + moreover from linli linlj linnlj have "I_intterm ats (Add li (lin_neg lj)) = I_intterm ats (Sub li lj)" + by (simp add: lin_neg_corr) + ultimately have ?thesis using prems by simp + } + ultimately show ?thesis by blast + qed +next + case (6 i j)show ?case + proof- + have cses: "(linearize i = None) \ + ((\ li. linearize i = Some li) \ linearize j = None) \ + ((\ li. linearize i = Some li) \ (\ bj. linearize j = Some (Cst bj))) + \ ((\ bi. linearize i = Some (Cst bi)) \ (\ lj. linearize j = Some lj)) + \ ((\ li. linearize i = Some li \ \ (\ bi. li = Cst bi)) \ (\ lj. linearize j = Some lj \ \ (\ bj. lj = Cst bj)))" by auto + moreover + { + assume nlini: "linearize i = None" + from nlini have "linearize (Mult i j) = None" by (simp add: Let_def measure_def inv_image_def) with prems have ?thesis by auto + } + moreover + { + assume lini: "\ li. linearize i = Some li" + and nlinj: "linearize j = None" + + from lini obtain "li" where "linearize i = Some li" by blast + moreover from prems have "linearize (Mult i j) = None" + by (cases li) (simp_all add: Let_def measure_def inv_image_def) + with prems have ?thesis by auto + } + moreover + { + assume lini: "\li. linearize i = Some li" + and linj: "\bj. linearize j = Some (Cst bj)" + from lini have lini2: "linearize i \ None" by simp + from linj have linj2: "linearize j \ None" by auto + from lini obtain "li" where "linearize i = Some li" by blast + from lini2 have "islinintterm (the (linearize i))" by (simp add: linearize_linear1) + with prems have linli: "islinintterm li" by simp + moreover + from linj obtain "bj" where "linearize j = Some (Cst bj)" by blast + have linlj: "islinintterm (Cst bj)" by simp + moreover from prems have "linearize (Mult i j) = Some (lin_mul (bj,li))" + by (cases li) (auto simp add: measure_def inv_image_def) + then have lm1: "I_intterm ats (the(linearize (Mult i j))) = I_intterm ats (lin_mul (bj,li))" by simp + moreover from linli linlj have "I_intterm ats (lin_mul(bj,li)) = I_intterm ats (Mult li (Cst bj))" by (simp add: lin_mul_corr) + with prems + have "I_intterm ats (lin_mul(bj,li)) = I_intterm ats (Mult li (the (linearize j)))" + by auto + moreover have "I_intterm ats (Mult li (the (linearize j))) = I_intterm ats (Mult i (the (linearize j)))" using prems by simp + moreover have "I_intterm ats i = I_intterm ats (the (linearize i))" + using lini2 lini "6.hyps" by simp + moreover have "I_intterm ats j = I_intterm ats (the (linearize j))" + using prems by (cases li) (auto simp add: measure_def inv_image_def) + ultimately have ?thesis by auto } + moreover + { assume lini: "\bi. linearize i = Some (Cst bi)" + and linj: "\lj. linearize j = Some lj" + from lini have lini2 : "linearize i \ None" by auto + from linj have linj2 : "linearize j \ None" by auto + from lini obtain "bi" where "linearize i = Some (Cst bi)" by blast + have linli: "islinintterm (Cst bi)" using prems by simp + moreover + from linj obtain "lj" where "linearize j = Some lj" by blast + from linj2 have "islinintterm (the (linearize j))" by (simp add: linearize_linear1) + then have linlj: "islinintterm lj" by (simp!) + moreover from linli lini linj have "linearize (Mult i j) = Some (lin_mul (bi,lj))" apply (simp add: measure_def inv_image_def) + apply auto by (case_tac "li::intterm",auto!) + then have lm1: "I_intterm ats (the(linearize (Mult i j))) = I_intterm ats (lin_mul (bi,lj))" by simp + moreover from linli linlj have "I_intterm ats (lin_mul(bi,lj)) = I_intterm ats (Mult (Cst bi) lj)" by (simp add: lin_mul_corr) + then have "I_intterm ats (lin_mul(bi,lj)) = I_intterm ats (Mult (the (linearize i)) lj)" by (auto!) + moreover have "I_intterm ats (Mult (the (linearize i)) lj) = I_intterm ats (Mult (the (linearize i)) j)" using lini lini2 by (simp!) + moreover have "I_intterm ats i = I_intterm ats (the (linearize i))" + using lini2 lini "6.hyps" by simp + moreover have "I_intterm ats j = I_intterm ats (the (linearize j))" + using linj linj2 lini lini2 linli linlj "6.hyps" by (auto!) + + ultimately have ?thesis by auto } + moreover + { assume linc: "\ li. linearize i = Some li \ \ (\ bi. li = Cst bi)" + and ljnc: "\ lj. linearize j = Some lj \ \ (\ bj. lj = Cst bj)" + from linc obtain "li" where "\ li. linearize i = Some li \ \ (\ bi. li = Cst bi)" by blast + moreover + from ljnc obtain "lj" where "\ lj. linearize j = Some lj \ \ (\ bj. lj = Cst bj)" by blast + ultimately have "linearize (Mult i j) = None" + apply (simp add: measure_def inv_image_def) + apply (case_tac "linearize i", auto) + apply (case_tac a) + apply (auto!) + by (case_tac "lj",auto)+ + then have ?thesis by (simp!) } + ultimately show ?thesis by blast + qed +qed simp_all + + +(* linearize, when successfull, preserves semantics *) +lemma linearize_corr: "\ t'. linearize t = Some t' \ I_intterm ats t = I_intterm ats t' " +proof- + fix t' + assume lint: "linearize t = Some t'" + show "I_intterm ats t = I_intterm ats t'" + proof- + from lint have lt: "linearize t \ None" by simp + then have "I_intterm ats t = I_intterm ats (the (linearize t))" + by (rule_tac linearize_corr1[OF lt]) + with lint show ?thesis by simp + qed +qed + +(* tries to linearize a formula *) +consts linform :: "QF \ QF option" +primrec +"linform (Le it1 it2) = + lift_bin(\x. \y. Le (lin_add(x,lin_neg y)) (Cst 0),linearize it1, linearize it2)" +"linform (Eq it1 it2) = + lift_bin(\x. \y. Eq (lin_add(x,lin_neg y)) (Cst 0),linearize it1, linearize it2)" +"linform (Divides d t) = + (case linearize d of + None \ None + | Some ld \ (case ld of + Cst b \ + (if (b=0) then None + else + (case linearize t of + None \ None + | Some lt \ Some (Divides ld lt))) + | _ \ None))" +"linform T = Some T" +"linform F = Some F" +"linform (NOT p) = lift_un NOT (linform p)" +"linform (And p q)= lift_bin(\f. \g. And f g, linform p, linform q)" +"linform (Or p q) = lift_bin(\f. \g. Or f g, linform p, linform q)" + +(* linearity of formulae *) +consts islinform :: "QF \ bool" +recdef islinform "measure size" +"islinform (Le it (Cst i)) = (i=0 \ islinintterm it )" +"islinform (Eq it (Cst i)) = (i=0 \ islinintterm it)" +"islinform (Divides (Cst d) t) = (d \ 0 \ islinintterm t)" +"islinform T = True" +"islinform F = True" +"islinform (NOT (Divides (Cst d) t)) = (d \ 0 \ islinintterm t)" +"islinform (NOT (Eq it (Cst i))) = (i=0 \ islinintterm it)" +"islinform (And p q)= ((islinform p) \ (islinform q))" +"islinform (Or p q) = ((islinform p) \ (islinform q))" +"islinform p = False" + +(* linform preserves nnf, if successful *) +lemma linform_nnf: + assumes nnfp: "isnnf p" + shows "\ p'. \linform p = Some p'\ \ isnnf p'" +using nnfp +proof (induct p rule: isnnf.induct, simp_all) + case (goal1 a b p') + show ?case + using prems + by (cases "linearize a", auto) (cases "linearize b", auto) +next + case (goal2 a b p') + show ?case + using prems + by (cases "linearize a", auto) (cases "linearize b", auto) +next + case (goal3 d t p') + show ?case + using prems + apply (cases "linearize d", auto) + apply (case_tac "a",auto) + apply (case_tac "int=0",auto) + by (cases "linearize t",auto) +next + case (goal4 f g p') show ?case + using prems + by (cases "linform f", auto) (cases "linform g", auto) +next + case (goal5 f g p') show ?case + using prems + by (cases "linform f", auto) (cases "linform g", auto) +next + case (goal6 d t p') show ?case + using prems + apply (cases "linearize d", auto) + apply (case_tac "a", auto) + apply (case_tac "int = 0",auto) + by (cases "linearize t", auto) +next + case (goal7 a b p') + show ?case + using prems + by (cases "linearize a", auto) (cases "linearize b", auto) + + +qed + + +lemma linform_corr: "\ lp. \ isnnf p ; linform p = Some lp \ \ + (qinterp ats p = qinterp ats lp)" +proof (induct p rule: linform.induct) + case (Le x y) + show ?case + using "Le.prems" + proof- + have "(\ lx ly. linearize x = Some lx \ linearize y = Some ly) \ + (linearize x = None) \ (linearize y = None)"by auto + moreover + { + assume linxy: "\ lx ly. linearize x = Some lx \ linearize y = Some ly" + from linxy obtain "lx" "ly" + where lxly:"linearize x = Some lx \ linearize y = Some ly" by blast + then + have lxeqx: "I_intterm ats x = I_intterm ats lx" + by (simp add: linearize_corr) + from lxly have lxlin: "islinintterm lx" + by (auto simp add: linearize_linear) + from lxly have lyeqy: "I_intterm ats y = I_intterm ats ly" + by (simp add: linearize_corr) + from lxly have lylin: "islinintterm ly" + by (auto simp add: linearize_linear) + from "prems" + have lpeqle: "lp = (Le (lin_add(lx,lin_neg ly)) (Cst 0))" + by auto + moreover + have lin1: "islinintterm (Cst 1)" by simp + then + have ?thesis + using lxlin lylin lin1 lin_add_lin lin_neg_lin "prems" lxly lpeqle + by (simp add: lin_add_corr lin_neg_corr lxeqx lyeqy) + + } + + moreover + { + assume "linearize x = None" + have ?thesis using "prems" by simp + } + + moreover + { + assume "linearize y = None" + then have ?thesis using "prems" + by (case_tac "linearize x", auto) + } + ultimately show ?thesis by blast + qed + +next + case (Eq x y) + show ?case + using "Eq.prems" + proof- + have "(\ lx ly. linearize x = Some lx \ linearize y = Some ly) \ + (linearize x = None) \ (linearize y = None)"by auto + moreover + { + assume linxy: "\ lx ly. linearize x = Some lx \ linearize y = Some ly" + from linxy obtain "lx" "ly" + where lxly:"linearize x = Some lx \ linearize y = Some ly" by blast + then + have lxeqx: "I_intterm ats x = I_intterm ats lx" + by (simp add: linearize_corr) + from lxly have lxlin: "islinintterm lx" + by (auto simp add: linearize_linear) + from lxly have lyeqy: "I_intterm ats y = I_intterm ats ly" + by (simp add: linearize_corr) + from lxly have lylin: "islinintterm ly" + by (auto simp add: linearize_linear) + from "prems" + have lpeqle: "lp = (Eq (lin_add(lx,lin_neg ly)) (Cst 0))" + by auto + moreover + have lin1: "islinintterm (Cst 1)" by simp + then + have ?thesis + using lxlin lylin lin1 lin_add_lin lin_neg_lin "prems" lxly lpeqle + by (simp add: lin_add_corr lin_neg_corr lxeqx lyeqy) + + } + + moreover + { + assume "linearize x = None" + have ?thesis using "prems" by simp + } + + moreover + { + assume "linearize y = None" + then have ?thesis using "prems" + by (case_tac "linearize x", auto) + } + ultimately show ?thesis by blast + qed + +next + case (Divides d t) + show ?case + using "Divides.prems" + apply (case_tac "linearize d",auto) + apply (case_tac a, auto) + apply (case_tac "int = 0", auto) + apply (case_tac "linearize t", auto) + apply (simp add: linearize_corr) + apply (case_tac a, auto) + apply (case_tac "int = 0", auto) + by (case_tac "linearize t", auto simp add: linearize_corr) +next + case (NOT f) show ?case + using "prems" + proof- + have "(\ lf. linform f = Some lf) \ (linform f = None)" by auto + moreover + { + assume linf: "\ lf. linform f = Some lf" + from prems have "isnnf (NOT f)" by simp + then have fnnf: "isnnf f" by (cases f) auto + from linf obtain "lf" where lf: "linform f = Some lf" by blast + then have "lp = NOT lf" using "prems" by auto + with "NOT.prems" "NOT.hyps" lf fnnf + have ?case by simp + } + moreover + { + assume "linform f = None" + then + have "linform (NOT f) = None" by simp + then + have ?thesis using "NOT.prems" by simp + } + ultimately show ?thesis by blast + qed +next + case (Or f g) + show ?case using "Or.hyps" + proof - + have "((\ lf. linform f = Some lf ) \ (\ lg. linform g = Some lg)) \ + (linform f = None) \ (linform g = None)" by auto + moreover + { + assume linf: "\ lf. linform f = Some lf" + and ling: "\ lg. linform g = Some lg" + from linf obtain "lf" where lf: "linform f = Some lf" by blast + from ling obtain "lg" where lg: "linform g = Some lg" by blast + from lf lg have "linform (Or f g) = Some (Or lf lg)" by simp + then have "lp = Or lf lg" using lf lg "prems" by simp + with lf lg "prems" have ?thesis by simp + } + moreover + { + assume "linform f = None" + then have ?thesis using "Or.prems" by auto + } + moreover + { + assume "linform g = None" + then have ?thesis using "Or.prems" by (case_tac "linform f", auto) + + } + ultimately show ?thesis by blast + qed +next + case (And f g) + show ?case using "And.hyps" + proof - + have "((\ lf. linform f = Some lf ) \ (\ lg. linform g = Some lg)) \ + (linform f = None) \ (linform g = None)" by auto + moreover + { + assume linf: "\ lf. linform f = Some lf" + and ling: "\ lg. linform g = Some lg" + from linf obtain "lf" where lf: "linform f = Some lf" by blast + from ling obtain "lg" where lg: "linform g = Some lg" by blast + from lf lg have "linform (And f g) = Some (And lf lg)" by simp + then have "lp = And lf lg" using lf lg "prems" by simp + with lf lg "prems" have ?thesis by simp + } + moreover + { + assume "linform f = None" + then have ?thesis using "And.prems" by auto + } + moreover + { + assume "linform g = None" + then have ?thesis using "And.prems" by (case_tac "linform f", auto) + + } + ultimately show ?thesis by blast + qed + +qed simp_all + + +(* the result of linform is a linear formula *) +lemma linform_lin: "\ lp. \ isnnf p ; linform p = Some lp\ \ islinform lp" +proof (induct p rule: linform.induct) + case (Le x y) + have "((\ lx. linearize x = Some lx) \ (\ ly. linearize y = Some ly)) \ + (linearize x = None) \ (linearize y = None) " by clarsimp + moreover + { + assume linx: "\ lx. linearize x = Some lx" + and liny: "\ ly. linearize y = Some ly" + from linx obtain "lx" where lx: "linearize x = Some lx" by blast + from liny obtain "ly" where ly: "linearize y = Some ly" by blast + from lx have lxlin: "islinintterm lx" by (simp add: linearize_linear) + from ly have lylin: "islinintterm ly" by (simp add: linearize_linear) + have lin1:"islinintterm (Cst 1)" by simp + have lin0: "islinintterm (Cst 0)" by simp + from "prems" have "lp = Le (lin_add(lx,lin_neg ly)) (Cst 0)" + by auto + with lin0 lin1 lxlin lylin "prems" + have ?case by (simp add: lin_add_lin lin_neg_lin) + + } + + moreover + { + assume "linearize x = None" + then have ?case using "prems" by simp + } + moreover + { + assume "linearize y = None" + then have ?case using "prems" by (case_tac "linearize x",simp_all) + } + ultimately show ?case by blast +next + case (Eq x y) + have "((\ lx. linearize x = Some lx) \ (\ ly. linearize y = Some ly)) \ + (linearize x = None) \ (linearize y = None) " by clarsimp + moreover + { + assume linx: "\ lx. linearize x = Some lx" + and liny: "\ ly. linearize y = Some ly" + from linx obtain "lx" where lx: "linearize x = Some lx" by blast + from liny obtain "ly" where ly: "linearize y = Some ly" by blast + from lx have lxlin: "islinintterm lx" by (simp add: linearize_linear) + from ly have lylin: "islinintterm ly" by (simp add: linearize_linear) + have lin1:"islinintterm (Cst 1)" by simp + have lin0: "islinintterm (Cst 0)" by simp + from "prems" have "lp = Eq (lin_add(lx,lin_neg ly)) (Cst 0)" + by auto + with lin0 lin1 lxlin lylin "prems" + have ?case by (simp add: lin_add_lin lin_neg_lin) + + } + + moreover + { + assume "linearize x = None" + then have ?case using "prems" by simp + } + moreover + { + assume "linearize y = None" + then have ?case using "prems" by (case_tac "linearize x",simp_all) + } + ultimately show ?case by blast +next + case (Divides d t) + show ?case + using prems + apply (case_tac "linearize d", auto) + apply (case_tac a, auto) + apply (case_tac "int = 0", auto) + + by (case_tac "linearize t",auto simp add: linearize_linear) +next + case (Or f g) + show ?case using "Or.hyps" + proof - + have "((\ lf. linform f = Some lf ) \ (\ lg. linform g = Some lg)) \ + (linform f = None) \ (linform g = None)" by auto + moreover + { + assume linf: "\ lf. linform f = Some lf" + and ling: "\ lg. linform g = Some lg" + from linf obtain "lf" where lf: "linform f = Some lf" by blast + from ling obtain "lg" where lg: "linform g = Some lg" by blast + from lf lg have "linform (Or f g) = Some (Or lf lg)" by simp + then have "lp = Or lf lg" using lf lg "prems" by simp + with lf lg "prems" have ?thesis by simp + } + moreover + { + assume "linform f = None" + then have ?thesis using "Or.prems" by auto + } + moreover + { + assume "linform g = None" + then have ?thesis using "Or.prems" by (case_tac "linform f", auto) + + } + ultimately show ?thesis by blast + qed +next + case (And f g) + show ?case using "And.hyps" + proof - + have "((\ lf. linform f = Some lf ) \ (\ lg. linform g = Some lg)) \ + (linform f = None) \ (linform g = None)" by auto + moreover + { + assume linf: "\ lf. linform f = Some lf" + and ling: "\ lg. linform g = Some lg" + from linf obtain "lf" where lf: "linform f = Some lf" by blast + from ling obtain "lg" where lg: "linform g = Some lg" by blast + from lf lg have "linform (And f g) = Some (And lf lg)" by simp + then have "lp = And lf lg" using lf lg "prems" by simp + with lf lg "prems" have ?thesis by simp + } + moreover + { + assume "linform f = None" + then have ?thesis using "And.prems" by auto + } + moreover + { + assume "linform g = None" + then have ?thesis using "And.prems" by (case_tac "linform f", auto) + + } + ultimately show ?thesis by blast + qed +next + case (NOT f) show ?case + using "prems" + proof- + have "(\ lf. linform f = Some lf) \ (linform f = None)" by auto + moreover + { + assume linf: "\ lf. linform f = Some lf" + from prems have "isnnf (NOT f)" by simp + then have fnnf: "isnnf f" by (cases f) auto + from linf obtain "lf" where lf: "linform f = Some lf" by blast + then have "lp = NOT lf" using "prems" by auto + with "NOT.prems" "NOT.hyps" lf fnnf + have ?thesis + using fnnf + apply (cases f, auto) + prefer 2 + apply (case_tac "linearize intterm1",auto) + apply (case_tac a, auto) + apply (case_tac "int = 0", auto) + apply (case_tac "linearize intterm2") + apply (auto simp add: linearize_linear) + apply (case_tac "linearize intterm1",auto) + by (case_tac "linearize intterm2") + (auto simp add: linearize_linear lin_add_lin lin_neg_lin) + } + moreover + { + assume "linform f = None" + then + have "linform (NOT f) = None" by simp + then + have ?thesis using "NOT.prems" by simp + } + ultimately show ?thesis by blast + qed +qed (simp_all) + + +(* linform, if successfull, preserves quantifier freeness *) +lemma linform_isnnf: "islinform p \ isnnf p" +by (induct p rule: islinform.induct) auto + +lemma linform_isqfree: "islinform p \ isqfree p" +using linform_isnnf nnf_isqfree by simp + +lemma linform_qfree: "\ p'. \ isnnf p ; linform p = Some p'\ \ isqfree p'" +using linform_isqfree linform_lin +by simp + +(* Definitions and lemmas about gcd and lcm *) +constdefs lcm :: "nat \ nat \ nat" + "lcm \ (\(m,n). m*n div gcd(m,n))" + +constdefs ilcm :: "int \ int \ int" + "ilcm \ \i.\j. int (lcm(nat(abs i),nat(abs j)))" + +(* ilcm_dvd12 are needed later *) +lemma lcm_dvd1: + assumes mpos: " m >0" + and npos: "n>0" + shows "m dvd (lcm(m,n))" +proof- + have "gcd(m,n) dvd n" by simp + then obtain "k" where "n = gcd(m,n) * k" using dvd_def by auto + then have "m*n div gcd(m,n) = m*(gcd(m,n)*k) div gcd(m,n)" by (simp add: mult_ac) + also have "\ = m*k" using mpos npos gcd_zero by simp + finally show ?thesis by (simp add: lcm_def) +qed + +lemma lcm_dvd2: + assumes mpos: " m >0" + and npos: "n>0" + shows "n dvd (lcm(m,n))" +proof- + have "gcd(m,n) dvd m" by simp + then obtain "k" where "m = gcd(m,n) * k" using dvd_def by auto + then have "m*n div gcd(m,n) = (gcd(m,n)*k)*n div gcd(m,n)" by (simp add: mult_ac) + also have "\ = n*k" using mpos npos gcd_zero by simp + finally show ?thesis by (simp add: lcm_def) +qed + +lemma ilcm_dvd1: +assumes anz: "a \ 0" + and bnz: "b \ 0" + shows "a dvd (ilcm a b)" +proof- + let ?na = "nat (abs a)" + let ?nb = "nat (abs b)" + have nap: "?na >0" using anz by simp + have nbp: "?nb >0" using bnz by simp + from nap nbp have "?na dvd lcm(?na,?nb)" using lcm_dvd1 by simp + thus ?thesis by (simp add: ilcm_def dvd_int_iff) +qed + + +lemma ilcm_dvd2: +assumes anz: "a \ 0" + and bnz: "b \ 0" + shows "b dvd (ilcm a b)" +proof- + let ?na = "nat (abs a)" + let ?nb = "nat (abs b)" + have nap: "?na >0" using anz by simp + have nbp: "?nb >0" using bnz by simp + from nap nbp have "?nb dvd lcm(?na,?nb)" using lcm_dvd2 by simp + thus ?thesis by (simp add: ilcm_def dvd_int_iff) +qed + +lemma zdvd_self_abs1: "(d::int) dvd (abs d)" +by (case_tac "d <0", simp_all) + +lemma zdvd_self_abs2: "(abs (d::int)) dvd d" +by (case_tac "d<0", simp_all) + +(* lcm a b is positive for positive a and b *) + +lemma lcm_pos: + assumes mpos: "m > 0" + and npos: "n>0" + shows "lcm (m,n) > 0" + +proof(rule ccontr, simp add: lcm_def gcd_zero) +assume h:"m*n div gcd(m,n) = 0" +from mpos npos have "gcd (m,n) \ 0" using gcd_zero by simp +hence gcdp: "gcd(m,n) > 0" by simp +with h +have "m*n < gcd(m,n)" + by (cases "m * n < gcd (m, n)") (auto simp add: div_if[OF gcdp, where m="m*n"]) +moreover +have "gcd(m,n) dvd m" by simp + with mpos dvd_imp_le have t1:"gcd(m,n) \ m" by simp + with npos have t1:"gcd(m,n)*n \ m*n" by simp + have "gcd(m,n) \ gcd(m,n)*n" using npos by simp + with t1 have "gcd(m,n) \ m*n" by arith +ultimately show "False" by simp +qed + +lemma ilcm_pos: + assumes apos: " 0 < a" + and bpos: "0 < b" + shows "0 < ilcm a b" +proof- + let ?na = "nat (abs a)" + let ?nb = "nat (abs b)" + have nap: "?na >0" using apos by simp + have nbp: "?nb >0" using bpos by simp + have "0 < lcm (?na,?nb)" by (rule lcm_pos[OF nap nbp]) + thus ?thesis by (simp add: ilcm_def) +qed + +(* fomlcm computes the lcm of all c, where c is the coeffitient of Var 0 *) +consts formlcm :: "QF \ int" +recdef formlcm "measure size" +"formlcm (Le (Add (Mult (Cst c) (Var 0)) r) (Cst i)) = abs c " +"formlcm (Eq (Add (Mult (Cst c) (Var 0)) r) (Cst i)) = abs c " +"formlcm (Divides (Cst d) (Add (Mult (Cst c) (Var 0)) r)) = abs c" +"formlcm (NOT p) = formlcm p" +"formlcm (And p q)= ilcm (formlcm p) (formlcm q)" +"formlcm (Or p q) = ilcm (formlcm p) (formlcm q)" +"formlcm p = 1" + +(* the property that formlcm should fullfill *) +consts divideallc:: "int \ QF \ bool" +recdef divideallc "measure (\(i,p). size p)" +"divideallc (l,Le (Add (Mult (Cst c) (Var 0)) r) (Cst i)) = (c dvd l)" +"divideallc (l,Eq (Add (Mult (Cst c) (Var 0)) r) (Cst i)) = (c dvd l)" +"divideallc(l,Divides (Cst d) (Add (Mult (Cst c) (Var 0)) r)) = (c dvd l)" +"divideallc (l,NOT p) = divideallc(l,p)" +"divideallc (l,And p q) = (divideallc (l,p) \ divideallc (l,q))" +"divideallc (l,Or p q) = (divideallc (l,p) \ divideallc (l,q))" +"divideallc p = True" + +(* formlcm retuns a positive integer *) +lemma formlcm_pos: + assumes linp: "islinform p" + shows "0 < formlcm p" +using linp +proof (induct p rule: formlcm.induct, simp_all add: ilcm_pos) + case (goal1 c r i) + have "i=0 \ i \ 0" by simp + moreover + { + assume "i \ 0" then have ?case using prems by simp + } + moreover + { + assume iz: "i = 0" + then have "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + then have "c\0" + using prems + by (simp add: islininttermc0r[where c="c" and n="0" and r="r"]) + then have ?case by simp + } + ultimately + show ?case by blast +next + case (goal2 c r i) + have "i=0 \ i \ 0" by simp + moreover + { + assume "i \ 0" then have ?case using prems by simp + } + moreover + { + assume iz: "i = 0" + then have "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + then have "c\0" + using prems + by (simp add: islininttermc0r[where c="c" and n="0" and r="r"]) + then have ?case by simp + } + ultimately + show ?case by blast + +next + case (goal3 d c r) + show ?case using prems by (simp add: islininttermc0r[where c="c" and n="0" and r="r"]) +next + case (goal4 f) + show ?case using prems + by (cases f,auto) (case_tac "intterm2", auto,case_tac "intterm1", auto) +qed + +lemma divideallc_mono: "\ c. \ divideallc(c,p) ; c dvd d\ \ divideallc (d,p)" +proof (induct d p rule: divideallc.induct, simp_all) + case (goal1 l a b) show ?case by ( rule zdvd_trans [where m="a" and n="b" and k="l"]) +next + case (goal2 l a b) show ?case by ( rule zdvd_trans [where m="a" and n="b" and k="l"]) +next + case (goal3 l a b) show ?case by ( rule zdvd_trans [where m="a" and n="b" and k="l"]) +next + case (goal4 l f g k) + have "divideallc (l,g)" using prems by clarsimp + moreover have "divideallc (l,f)" using prems by clarsimp + ultimately + show ?case by simp +next + case (goal5 l f g k) + have "divideallc (l,g)" using prems by clarsimp + moreover have "divideallc (l,f)" using prems by clarsimp + ultimately + show ?case by simp + +qed + +(* fomlcm retuns a number all coeffitients of Var 0 divide *) + +lemma formlcm_divideallc: + assumes linp: "islinform p" + shows "divideallc(formlcm p, p)" +using linp +proof (induct p rule: formlcm.induct, simp_all add: zdvd_self_abs1) + case (goal1 f) + show ?case using prems + by (cases f,auto) (case_tac "intterm2", auto, case_tac "intterm1",auto) +next + case (goal2 f g) + have "formlcm f >0" using formlcm_pos prems by simp + hence "formlcm f \ 0" by simp + moreover have "formlcm g > 0" using formlcm_pos prems by simp + hence "formlcm g \ 0" by simp + ultimately + show ?case using prems formlcm_pos + by (simp add: ilcm_dvd1 ilcm_dvd2 + divideallc_mono[where c="formlcm f" and d="ilcm (formlcm f) (formlcm g)"] + divideallc_mono[where c="formlcm g" and d="ilcm (formlcm f) (formlcm g)"]) +next + case (goal3 f g) + have "formlcm f >0" using formlcm_pos prems by simp + hence "formlcm f \ 0" by simp + moreover have "formlcm g > 0" using formlcm_pos prems by simp + hence "formlcm g \ 0" by simp + ultimately + show ?case using prems + by (simp add: ilcm_dvd1 ilcm_dvd2 + divideallc_mono[where c="formlcm f" and d="ilcm (formlcm f) (formlcm g)"] + divideallc_mono[where c="formlcm g" and d="ilcm (formlcm f) (formlcm g)"]) +qed + +(* adjustcoeff transforms the formula given an l , look at correctness thm*) +consts adjustcoeff :: "int \ QF \ QF" +recdef adjustcoeff "measure (\(l,p). size p)" +"adjustcoeff (l,(Le (Add (Mult (Cst c) (Var 0)) r) (Cst i))) = + (if c\0 then + Le (Add (Mult (Cst -1) (Var 0)) (lin_mul (- (l div c), r))) (Cst (0::int)) + else + Le (Add (Mult (Cst 1) (Var 0)) (lin_mul (l div c, r))) (Cst (0::int)))" +"adjustcoeff (l,(Eq (Add (Mult (Cst c) (Var 0)) r) (Cst i))) = + (Eq (Add (Mult (Cst 1) (Var 0)) (lin_mul (l div c, r))) (Cst (0::int)))" +"adjustcoeff (l,Divides (Cst d) (Add (Mult (Cst c) (Var 0)) r)) = + Divides (Cst ((l div c) * d)) + (Add (Mult (Cst 1) (Var 0)) (lin_mul (l div c, r)))" +"adjustcoeff (l,NOT (Divides (Cst d) (Add (Mult (Cst c) (Var 0)) r))) = NOT (Divides (Cst ((l div c) * d)) + (Add (Mult (Cst 1) (Var 0)) (lin_mul (l div c, r))))" +"adjustcoeff (l,(NOT(Eq (Add (Mult (Cst c) (Var 0)) r) (Cst i)))) = + (NOT(Eq (Add (Mult (Cst 1) (Var 0)) (lin_mul (l div c, r))) (Cst (0::int))))" +"adjustcoeff (l,And p q) = And (adjustcoeff (l,p)) (adjustcoeff(l,q))" +"adjustcoeff (l,Or p q) = Or (adjustcoeff (l,p)) (adjustcoeff(l,q))" +"adjustcoeff (l,p) = p" + + +(* unitycoeff expects a quantifier free formula an transforms it to an equivalent formula where the bound variable occurs only with coeffitient 1 or -1 *) +constdefs unitycoeff :: "QF \ QF" + "unitycoeff p == + (let l = formlcm p; + p' = adjustcoeff (l,p) + in (if l=1 then p' else + (And (Divides (Cst l) (Add (Mult (Cst 1) (Var 0)) (Cst 0))) p')))" + +(* what is a unified formula *) +consts isunified :: "QF \ bool" +recdef isunified "measure size" +"isunified (Le (Add (Mult (Cst i) (Var 0)) r) (Cst z)) = + ((abs i) = 1 \ (islinform(Le (Add (Mult (Cst i) (Var 0)) r) (Cst z))))" +"isunified (Eq (Add (Mult (Cst i) (Var 0)) r) (Cst z)) = + ((abs i) = 1 \ (islinform(Le (Add (Mult (Cst i) (Var 0)) r) (Cst z))))" +"isunified (NOT(Eq (Add (Mult (Cst i) (Var 0)) r) (Cst z))) = + ((abs i) = 1 \ (islinform(Le (Add (Mult (Cst i) (Var 0)) r) (Cst z))))" +"isunified (Divides (Cst d) (Add (Mult (Cst i) (Var 0)) r)) = + ((abs i) = 1 \ (islinform(Divides (Cst d) (Add (Mult (Cst i) (Var 0)) r))))" +"isunified (NOT(Divides (Cst d) (Add (Mult (Cst i) (Var 0)) r))) = + ((abs i) = 1 \ (islinform(NOT(Divides (Cst d) (Add (Mult (Cst i) (Var 0)) r)))))" +"isunified (And p q) = (isunified p \ isunified q)" +"isunified (Or p q) = (isunified p \ isunified q)" +"isunified p = islinform p" + +lemma unified_islinform: "isunified p \ islinform p" +by (induct p rule: isunified.induct) auto + +lemma adjustcoeff_lenpos: + "0 < n \ adjustcoeff (l, Le (Add (Mult (Cst i) (Var n)) r) (Cst c)) = + Le (Add (Mult (Cst i) (Var n)) r) (Cst c)" +by (cases n, auto) + +lemma adjustcoeff_eqnpos: + "0 < n \ adjustcoeff (l, Eq (Add (Mult (Cst i) (Var n)) r) (Cst c)) = + Eq (Add (Mult (Cst i) (Var n)) r) (Cst c)" +by (cases n, auto) + + +(* Properties of adjustcoeff and unitycoeff *) + +(* Some simple lemmas used afterwards *) +lemma zmult_zle_mono: "(i::int) \ j \ 0 \ k \ k * i \ k * j" + apply (erule order_le_less [THEN iffD1, THEN disjE, of "0::int"]) + apply (erule order_le_less [THEN iffD1, THEN disjE]) + apply (rule order_less_imp_le) + apply (rule zmult_zless_mono2) + apply simp_all + done + +lemma zmult_zle_mono_eq: + assumes kpos: "0 < k" + shows "((i::int) \ j) = (k*i \ k*j)" (is "?P = ?Q") +proof + assume P: ?P + from kpos have kge0: "0 \ k" by simp + show ?Q + by (rule zmult_zle_mono[OF P kge0]) +next + assume ?Q + then have "k*i - k*j \ 0" by simp + then have le1: "k*(i-j) \ k*0" + by (simp add: zdiff_zmult_distrib2) + have "i -j \ 0" + by (rule mult_left_le_imp_le[OF le1 kpos]) + then + show ?P by simp +qed + + +lemma adjustcoeff_le_corr: + assumes lpos: "0 < l" + and ipos: "0 < (i::int)" + and dvd: "i dvd l" + shows "(i*x + r \ 0) = (l*x + ((l div i)*r) \ 0)" +proof- + from lpos ipos have ilel: "i\l" by (simp add: zdvd_imp_le [OF dvd lpos]) + from ipos have inz: "i \ 0" by simp + have "i div i\ l div i" + by (simp add: zdiv_mono1[OF ilel ipos]) + then have ldivipos:"0 < l div i" + by (simp add: zdiv_self[OF inz]) + + from dvd have "\i'. i*i' = l" by (auto simp add: dvd_def) + then obtain "i'" where ii'eql: "i*i' = l" by blast + have "(i * x + r \ 0) = (l div i * (i * x + r) \ l div i * 0)" + by (rule zmult_zle_mono_eq[OF ldivipos, where i="i*x + r" and j="0"]) + also + have "(l div i * (i * x + r) \ l div i * 0) = ((l div i * i) * x + ((l div i)*r) \ 0)" + by (simp add: mult_ac) + also have "((l div i * i) * x + ((l div i)*r) \ 0) = (l*x + ((l div i)*r) \ 0)" + using sym[OF ii'eql] inz + by (simp add: zmult_ac) + finally + show ?thesis + by simp +qed + +lemma adjustcoeff_le_corr2: + assumes lpos: "0 < l" + and ineg: "(i::int) < 0" + and dvd: "i dvd l" + shows "(i*x + r \ 0) = ((-l)*x + ((-(l div i))*r) \ 0)" +proof- + from dvd have midvdl: "-i dvd l" by simp + from ineg have mipos: "0 < -i" by simp + from lpos ineg have milel: "-i\l" by (simp add: zdvd_imp_le [OF midvdl lpos]) + from ineg have inz: "i \ 0" by simp + have "l div i\ -i div i" + by (simp add: zdiv_mono1_neg[OF milel ineg]) + then have "l div i \ -1" + apply (simp add: zdiv_zminus1_eq_if[OF inz, where a="i"]) + by (simp add: zdiv_self[OF inz]) + then have ldivineg: "l div i < 0" by simp + then have mldivipos: "0 < - (l div i)" by simp + + from dvd have "\i'. i*i' = l" by (auto simp add: dvd_def) + then obtain "i'" where ii'eql: "i*i' = l" by blast + have "(i * x + r \ 0) = (- (l div i) * (i * x + r) \ - (l div i) * 0)" + by (rule zmult_zle_mono_eq[OF mldivipos, where i="i*x + r" and j="0"]) + also + have "(- (l div i) * (i * x + r) \ - (l div i) * 0) = (-((l div i) * i) * x \ (l div i)*r)" + by (simp add: mult_ac) + also have " (-((l div i) * i) * x \ (l div i)*r) = (- (l*x) \ (l div i)*r)" + using sym[OF ii'eql] inz + by (simp add: zmult_ac) + finally + show ?thesis + by simp +qed + +(* FIXME : Move this theorem above, it simplifies the 2 theorems above : adjustcoeff_le_corr1,2 *) +lemma dvd_div_pos: + assumes bpos: "0 < (b::int)" + and anz: "a\0" + and dvd: "a dvd b" + shows "(b div a)*a = b" +proof- + from anz have "0 < a \ a < 0" by arith + moreover + { + assume apos: "0 < a" + from bpos apos have aleb: "a\b" by (simp add: zdvd_imp_le [OF dvd bpos]) + have "a div a\ b div a" + by (simp add: zdiv_mono1[OF aleb apos]) + then have bdivapos:"0 < b div a" + by (simp add: zdiv_self[OF anz]) + + from dvd have "\a'. a*a' = b" by (auto simp add: dvd_def) + then obtain "a'" where aa'eqb: "a*a' = b" by blast + then have ?thesis using anz sym[OF aa'eqb] by simp + + } + moreover + { + assume aneg: "a < 0" + from dvd have midvdb: "-a dvd b" by simp + from aneg have mapos: "0 < -a" by simp + from bpos aneg have maleb: "-a\b" by (simp add: zdvd_imp_le [OF midvdb bpos]) + from aneg have anz: "a \ 0" by simp + have "b div a\ -a div a" + by (simp add: zdiv_mono1_neg[OF maleb aneg]) + then have "b div a \ -1" + apply (simp add: zdiv_zminus1_eq_if[OF anz, where a="a"]) + by (simp add: zdiv_self[OF anz]) + then have bdivaneg: "b div a < 0" by simp + then have mbdivapos: "0 < - (b div a)" by simp + + from dvd have "\a'. a*a' = b" by (auto simp add: dvd_def) + then obtain "a'" where aa'eqb: "a*a' = b" by blast + then have ?thesis using anz sym[OF aa'eqb] by (simp) + } + ultimately show ?thesis by blast +qed + +lemma adjustcoeff_eq_corr: + assumes lpos: "(0::int) < l" + and inz: "i \ 0" + and dvd: "i dvd l" + shows "(i*x + r = 0) = (l*x + ((l div i)*r) = 0)" +proof- + thm dvd_div_pos[OF lpos inz dvd] + have ldvdii: "(l div i)*i = l" by (rule dvd_div_pos[OF lpos inz dvd]) + have ldivinz: "l div i \ 0" using inz ldvdii lpos by auto + have "(i*x + r = 0) = ((l div i)*(i*x + r) = (l div i)*0)" + using ldivinz by arith + also have "\ = (((l div i)*i)*x + (l div i)*r = 0)" + by (simp add: zmult_ac) + finally show ?thesis using ldvdii by simp +qed + + + +(* Correctness theorem for adjustcoeff *) +lemma adjustcoeff_corr: + assumes linp: "islinform p" + and alldvd: "divideallc (l,p)" + and lpos: "0 < l" + shows "qinterp (a#ats) p = qinterp ((a*l)#ats) (adjustcoeff(l, p))" +using linp alldvd +proof (induct p rule: islinform.induct,simp_all) + case (goal1 t c) + from prems have cz: "c=0" by simp + then have ?case + using prems + proof(induct t rule: islinintterm.induct) + case (2 i n i') show ?case using prems + proof- + from prems have "i\0" by simp + then + have "(n=0 \ i < 0) \ (n=0 \ i > 0) \ n\0" by arith + moreover + { + assume "n\0" then have ?thesis + by (simp add: nth_pos2 adjustcoeff_lenpos) + } + moreover + { + assume nz: "n=0" + and ipos: "0 < i" + from prems nz have idvdl: "i dvd l" by simp + have "(i*a + i' \ 0) = (l*a+ ((l div i)*i') \ 0)" + by (rule adjustcoeff_le_corr[OF lpos ipos idvdl]) + then + have ?thesis using prems by (simp add: mult_ac) + } + moreover + { + assume nz: "n=0" + and ineg: "i < 0" + from prems nz have idvdl: "i dvd l" by simp + have "(i*a+i' \ 0) = (-l*a + (-(l div i) * i') \ 0)" + by (rule adjustcoeff_le_corr2[OF lpos ineg idvdl]) + then + have ?thesis using prems + by (simp add: zmult_ac) + } + ultimately show ?thesis by blast + qed + next + case (3 i n i' n' r) show ?case using prems + proof- + from prems + have lininrp: "islinintterm (Add (Mult (Cst i') (Var n')) r)" + by simp + then + have "islint (Add (Mult (Cst i') (Var n')) (r))" + by (simp add: islinintterm_eq_islint) + then have linr: "islintn(Suc n',r)" + by (simp add: islinintterm_subt[OF lininrp] islinintterm_eq_islint islint_def) + from lininrp have linr2: "islinintterm r" + by (simp add: islinintterm_subt[OF lininrp]) + from prems have "n < n'" by simp + then have nppos: "0 < n'" by simp + from prems have "i\0" by simp + then + have "(n=0 \ i < 0) \ (n=0 \ i > 0) \ n\0" by arith + moreover + { + assume nnz: "n\0" + from linr have ?thesis using nppos nnz intterm_novar0[OF lininrp] prems + apply (simp add: adjustcoeff_lenpos linterm_novar0[OF linr, where x="a" and y="a*l"]) + by (simp add: nth_pos2) + + } + moreover + { + assume nz: "n=0" + and ipos: "0 < i" + from prems nz have idvdl: "i dvd l" by simp + have "(i * a + (i' * (a # ats) ! n' + I_intterm (a # ats) r) \ 0) = + (l * a + l div i * (i' * (a # ats) ! n' + I_intterm (a # ats) r) \ 0)" + by (rule adjustcoeff_le_corr[OF lpos ipos idvdl]) + then + have ?thesis using prems linr linr2 + by (simp add: mult_ac nth_pos2 lin_mul_corr + linterm_novar0[OF linr, where x="a" and y="a*l"]) + } + moreover + { + assume nz: "n=0" + and ineg: "i < 0" + from prems nz have idvdl: "i dvd l" by simp + have "(i * a + (i' * (a # ats) ! n' + I_intterm (a # ats) r) \ 0) = + (- l * a + - (l div i) * (i' * (a # ats) ! n' + I_intterm (a # ats) r) \ 0)" + by (rule adjustcoeff_le_corr2[OF lpos ineg idvdl, where x="a" and r="(i'* (a#ats) ! n' + I_intterm (a#ats) r )"]) + then + have ?thesis using prems linr linr2 + by (simp add: zmult_ac nth_pos2 lin_mul_corr + linterm_novar0[OF linr, where x="a" and y="a*l"] ) + } + ultimately show ?thesis by blast + qed + qed simp_all + then show ?case by simp + +next + case (goal2 t c) + from prems have cz: "c=0" by simp + then have ?case + using prems + proof(induct t rule: islinintterm.induct) + case (2 i n i') show ?case using prems + proof- + from prems have inz: "i\0" by simp + then + have "n=0 \ n\0" by arith + moreover + { + assume "n\0" then have ?thesis + by (simp add: nth_pos2 adjustcoeff_eqnpos) + } + moreover + { + assume nz: "n=0" + from prems nz have idvdl: "i dvd l" by simp + have "(i*a + i' = 0) = (l*a+ ((l div i)*i') = 0)" + by (rule adjustcoeff_eq_corr[OF lpos inz idvdl]) + then + have ?thesis using prems by (simp add: mult_ac) + } + ultimately show ?thesis by blast + qed + next + case (3 i n i' n' r) show ?case using prems + proof- + from prems + have lininrp: "islinintterm (Add (Mult (Cst i') (Var n')) r)" + by simp + then + have "islint (Add (Mult (Cst i') (Var n')) (r))" + by (simp add: islinintterm_eq_islint) + then have linr: "islintn(Suc n',r)" + by (simp add: islinintterm_subt[OF lininrp] islinintterm_eq_islint islint_def) + from lininrp have linr2: "islinintterm r" + by (simp add: islinintterm_subt[OF lininrp]) + from prems have "n < n'" by simp + then have nppos: "0 < n'" by simp + from prems have "i\0" by simp + then + have "n=0 \ n\0" by arith + moreover + { + assume nnz: "n\0" + from linr have ?thesis using nppos nnz intterm_novar0[OF lininrp] prems + apply (simp add: adjustcoeff_eqnpos linterm_novar0[OF linr, where x="a" and y="a*l"]) + by (simp add: nth_pos2) + + } + moreover + { + assume nz: "n=0" + from prems have inz: "i \ 0" by auto + from prems nz have idvdl: "i dvd l" by simp + thm adjustcoeff_eq_corr[OF lpos inz idvdl] + have "(i * a + (i' * (a # ats) ! n' + I_intterm (a # ats) r) = 0) = + (l * a + l div i * (i' * (a # ats) ! n' + I_intterm (a # ats) r) = 0)" + by (rule adjustcoeff_eq_corr[OF lpos inz idvdl]) + then + have ?thesis using prems linr linr2 + by (simp add: mult_ac nth_pos2 lin_mul_corr + linterm_novar0[OF linr, where x="a" and y="a*l"]) + } + ultimately show ?thesis by blast + qed + qed simp_all + then show ?case by simp + +next + case (goal3 d t) show ?case + using prems + proof (induct t rule: islinintterm.induct) + case (2 i n i') + have "n=0 \ (\m. (n = Suc m))" by arith + moreover + { + assume "\m. n = Suc m" + then have ?case using prems by auto + } + moreover + { + assume nz: "n=0" + from prems have inz: "i\0" by simp + from prems have idvdl: "i dvd l" by simp + have ldiviieql: "l div i * i = l" by (rule dvd_div_pos[OF lpos inz idvdl]) + with lpos have ldivinz: "0 \ l div i" by auto + + then have ?case using prems + apply simp + apply (simp add: + ac_dvd_eq[OF ldivinz, where m="d" and c="i" and n="a" and t="i'"] + ldiviieql) + by (simp add: zmult_commute) + } + ultimately show ?case by blast + + next + case (3 i n i' n' r) + from prems + have lininrp: "islinintterm (Add (Mult (Cst i') (Var n')) r)" + by simp + then + have "islint (Add (Mult (Cst i') (Var n')) (r))" + by (simp add: islinintterm_eq_islint) + then have linr: "islintn(Suc n',r)" + by (simp add: islinintterm_subt[OF lininrp] islinintterm_eq_islint islint_def) + from lininrp have linr2: "islinintterm r" + by (simp add: islinintterm_subt[OF lininrp]) + from prems have "n < n'" by simp + then have nppos: "0 < n'" by simp + from prems have inz: "i\0" by simp + + have "n=0 \ (\m. (n = Suc m))" by arith + moreover + { + assume "\m. n = Suc m" + then have npos: "0 < n" by arith + have ?case using nppos intterm_novar0[OF lininrp] prems + apply (auto simp add: linterm_novar0[OF linr, where x="a" and y="a*l"]) + by (simp_all add: nth_pos2) + } + moreover + { + assume nz: "n=0" + from prems have idvdl: "i dvd l" by simp + have ldiviieql: "l div i * i = l" by (rule dvd_div_pos[OF lpos inz idvdl]) + with lpos have ldivinz: "0 \ l div i" by auto + + then have ?case using prems linr2 linr + apply (simp add: nth_pos2 lin_mul_corr linterm_novar0) + + apply (simp add: ac_dvd_eq[OF ldivinz, where m="d" and c="i" and n="a" and t="(i' * ats ! (n' - Suc 0) + I_intterm (a # ats) r)"] ldiviieql) + by (simp add: zmult_ac linterm_novar0[OF linr, where x="a" and y="a*l"]) + } + ultimately show ?case by blast + + qed simp_all +next + case (goal4 d t) show ?case + using prems + proof (induct t rule: islinintterm.induct) + case (2 i n i') + have "n=0 \ (\m. (n = Suc m))" by arith + moreover + { + assume "\m. n = Suc m" + then have ?case using prems by auto + } + moreover + { + assume nz: "n=0" + from prems have inz: "i\0" by simp + from prems have idvdl: "i dvd l" by simp + have ldiviieql: "l div i * i = l" by (rule dvd_div_pos[OF lpos inz idvdl]) + with lpos have ldivinz: "0 \ l div i" by auto + + then have ?case using prems + apply simp + apply (simp add: + ac_dvd_eq[OF ldivinz, where m="d" and c="i" and n="a" and t="i'"] + ldiviieql) + by (simp add: zmult_commute) + } + ultimately show ?case by blast + + next + case (3 i n i' n' r) + from prems + have lininrp: "islinintterm (Add (Mult (Cst i') (Var n')) r)" + by simp + then + have "islint (Add (Mult (Cst i') (Var n')) (r))" + by (simp add: islinintterm_eq_islint) + then have linr: "islintn(Suc n',r)" + by (simp add: islinintterm_subt[OF lininrp] islinintterm_eq_islint islint_def) + from lininrp have linr2: "islinintterm r" + by (simp add: islinintterm_subt[OF lininrp]) + from prems have "n < n'" by simp + then have nppos: "0 < n'" by simp + from prems have inz: "i\0" by simp + + have "n=0 \ (\m. (n = Suc m))" by arith + moreover + { + assume "\m. n = Suc m" + then have npos: "0 < n" by arith + have ?case using nppos intterm_novar0[OF lininrp] prems + apply (auto simp add: linterm_novar0[OF linr, where x="a" and y="a*l"]) + by (simp_all add: nth_pos2) + } + moreover + { + assume nz: "n=0" + from prems have idvdl: "i dvd l" by simp + have ldiviieql: "l div i * i = l" by (rule dvd_div_pos[OF lpos inz idvdl]) + with lpos have ldivinz: "0 \ l div i" by auto + + then have ?case using prems linr2 linr + apply (simp add: nth_pos2 lin_mul_corr linterm_novar0) + + apply (simp add: ac_dvd_eq[OF ldivinz, where m="d" and c="i" and n="a" and t="(i' * ats ! (n' - Suc 0) + I_intterm (a # ats) r)"] ldiviieql) + by (simp add: zmult_ac linterm_novar0[OF linr, where x="a" and y="a*l"]) + } + ultimately show ?case by blast + + qed simp_all +next + case (goal5 t c) + from prems have cz: "c=0" by simp + then have ?case + using prems + proof(induct t rule: islinintterm.induct) + case (2 i n i') show ?case using prems + proof- + from prems have inz: "i\0" by simp + then + have "n=0 \ n\0" by arith + moreover + { + assume "n\0" then have ?thesis + using prems + by (cases n, simp_all) + } + moreover + { + assume nz: "n=0" + from prems nz have idvdl: "i dvd l" by simp + have "(i*a + i' = 0) = (l*a+ ((l div i)*i') = 0)" + by (rule adjustcoeff_eq_corr[OF lpos inz idvdl]) + then + have ?thesis using prems by (simp add: mult_ac) + } + ultimately show ?thesis by blast + qed + next + case (3 i n i' n' r) show ?case using prems + proof- + from prems + have lininrp: "islinintterm (Add (Mult (Cst i') (Var n')) r)" + by simp + then + have "islint (Add (Mult (Cst i') (Var n')) (r))" + by (simp add: islinintterm_eq_islint) + then have linr: "islintn(Suc n',r)" + by (simp add: islinintterm_subt[OF lininrp] islinintterm_eq_islint islint_def) + from lininrp have linr2: "islinintterm r" + by (simp add: islinintterm_subt[OF lininrp]) + from prems have "n < n'" by simp + then have nppos: "0 < n'" by simp + from prems have "i\0" by simp + then + have "n=0 \ n\0" by arith + moreover + { + assume nnz: "n\0" + then have ?thesis using prems linr nppos nnz intterm_novar0[OF lininrp] + by (cases n, simp_all) + (simp add: nth_pos2 linterm_novar0[OF linr, where x="a" and y="a*l"]) + } + moreover + { + assume nz: "n=0" + from prems have inz: "i \ 0" by auto + from prems nz have idvdl: "i dvd l" by simp + have "(i * a + (i' * (a # ats) ! n' + I_intterm (a # ats) r) = 0) = + (l * a + l div i * (i' * (a # ats) ! n' + I_intterm (a # ats) r) = 0)" + by (rule adjustcoeff_eq_corr[OF lpos inz idvdl]) + then + have ?thesis using prems linr linr2 + by (simp add: mult_ac nth_pos2 lin_mul_corr + linterm_novar0[OF linr, where x="a" and y="a*l"]) + } + ultimately show ?thesis by blast + qed + qed simp_all + then show ?case by simp + +qed + +(* unitycoeff preserves semantics *) +lemma unitycoeff_corr: + assumes linp: "islinform p" + shows "qinterp ats (QEx p) = qinterp ats (QEx (unitycoeff p))" +proof- + + have lpos: "0 < formlcm p" by (rule formlcm_pos[OF linp]) + have dvd : "divideallc (formlcm p, p)" by (rule formlcm_divideallc[OF linp]) + show ?thesis using prems lpos dvd + proof (simp add: unitycoeff_def Let_def,case_tac "formlcm p = 1", + simp_all add: adjustcoeff_corr) + show "(\x. qinterp (x * formlcm p # ats) (adjustcoeff (formlcm p, p))) = + (\x. formlcm p dvd x \ + qinterp (x # ats) (adjustcoeff (formlcm p, p)))" + (is "(\x. ?P(x* (formlcm p))) = (\x. formlcm p dvd x \ ?P x)") + proof- + have "(\x. ?P(x* (formlcm p))) = (\x. ?P((formlcm p)*x))" + by (simp add: mult_commute) + also have "(\x. ?P((formlcm p)*x)) = (\x. (formlcm p dvd x) \ ?P x)" + by (simp add: unity_coeff_ex[where P="?P"]) + finally show ?thesis by simp + qed + qed +qed + +(* the resul of adjustcoeff is unified for all l with divideallc (l,p) *) +lemma adjustcoeff_unified: + assumes linp: "islinform p" + and dvdc: "divideallc(l,p)" + and lpos: "l > 0" + shows "isunified (adjustcoeff(l, p))" + using linp dvdc lpos + proof(induct l p rule: adjustcoeff.induct,simp_all add: lin_mul_lintn islinintterm_eq_islint islint_def) + case (goal1 l d c r) + from prems have "c >0 \ c < 0" by auto + moreover { + assume cpos: "c > 0 " + from prems have lp: "l > 0" by simp + from prems have cdvdl: "c dvd l" by simp + have clel: "c \ l" by (rule zdvd_imp_le[OF cdvdl lp]) + have "c div c \ l div c" by (rule zdiv_mono1[OF clel cpos]) + then have ?case using cpos by (simp add: zdiv_self) + } + moreover { + assume cneg: "c < 0" + + have mcpos: "-c > 0" by simp + then have mcnz: "-c \ 0" by simp + from prems have mcdvdl: "-c dvd l" + by simp + then have l1:"l mod -c = 0" by (simp add: zdvd_iff_zmod_eq_0) + from prems have lp: "l >0" by simp + have mclel: "-c \ l" by (rule zdvd_imp_le[OF mcdvdl lp]) + have "l div c = (-l div -c)" by simp + also have "\ = - (l div -c)" using l1 + by (simp only: zdiv_zminus1_eq_if[OF mcnz, where a="l"]) simp + finally have diveq: "l div c = - (l div -c)" by simp + + have "-c div -c \ l div -c" by (rule zdiv_mono1[OF mclel mcpos]) + then have "0 < l div -c" using cneg + by (simp add: zdiv_self) + then have ?case using diveq by simp + } + ultimately show ?case by blast + next + case (goal2 l p) from prems have "c >0 \ c < 0" by auto + moreover { + assume cpos: "c > 0 " + from prems have lp: "l > 0" by simp + from prems have cdvdl: "c dvd l" by simp + have clel: "c \ l" by (rule zdvd_imp_le[OF cdvdl lp]) + have "c div c \ l div c" by (rule zdiv_mono1[OF clel cpos]) + then have ?case using cpos by (simp add: zdiv_self) + } + moreover { + assume cneg: "c < 0" + + have mcpos: "-c > 0" by simp + then have mcnz: "-c \ 0" by simp + from prems have mcdvdl: "-c dvd l" + by simp + then have l1:"l mod -c = 0" by (simp add: zdvd_iff_zmod_eq_0) + from prems have lp: "l >0" by simp + have mclel: "-c \ l" by (rule zdvd_imp_le[OF mcdvdl lp]) + have "l div c = (-l div -c)" by simp + also have "\ = - (l div -c)" using l1 + by (simp only: zdiv_zminus1_eq_if[OF mcnz, where a="l"]) simp + finally have diveq: "l div c = - (l div -c)" by simp + + have "-c div -c \ l div -c" by (rule zdiv_mono1[OF mclel mcpos]) + then have "0 < l div -c" using cneg + by (simp add: zdiv_self) + then have ?case using diveq by simp + } + ultimately show ?case by blast + qed + +lemma adjustcoeff_lcm_unified: + assumes linp: "islinform p" + shows "isunified (adjustcoeff(formlcm p, p))" +using linp adjustcoeff_unified formlcm_pos formlcm_divideallc +by simp + +(* the result of unitycoeff is unified *) +lemma unitycoeff_unified: + assumes linp: "islinform p" + shows "isunified (unitycoeff p)" +using linp formlcm_pos[OF linp] +proof (auto simp add: unitycoeff_def Let_def adjustcoeff_lcm_unified) + assume f1: "formlcm p = 1" + have "isunified (adjustcoeff(formlcm p, p))" + by (rule adjustcoeff_lcm_unified[OF linp]) + with f1 + show "isunified (adjustcoeff(1,p))" by simp +qed + +lemma unified_isnnf: + assumes unifp: "isunified p" + shows "isnnf p" + using unified_islinform[OF unifp] linform_isnnf + by simp + +lemma unified_isqfree: "isunified p\ isqfree p" +using unified_islinform linform_isqfree +by auto + +(* Plus/Minus infinity , B and A set definitions *) + +consts minusinf :: "QF \ QF" + plusinf :: "QF \ QF" + aset :: "QF \ intterm list" + bset :: "QF \ intterm list" + +recdef minusinf "measure size" +"minusinf (Le (Add (Mult (Cst c) (Var 0)) r) z) = + (if c < 0 then F else T)" +"minusinf (Eq (Add (Mult (Cst c) (Var 0)) r) z) = F" +"minusinf (NOT(Eq (Add (Mult (Cst c) (Var 0)) r) z)) = T" +"minusinf (And p q) = And (minusinf p) (minusinf q)" +"minusinf (Or p q) = Or (minusinf p) (minusinf q)" +"minusinf p = p" + +recdef plusinf "measure size" +"plusinf (Le (Add (Mult (Cst c) (Var 0)) r) z) = + (if c < 0 then T else F)" +"plusinf (Eq (Add (Mult (Cst c) (Var 0)) r) z) = F" +"plusinf (NOT (Eq (Add (Mult (Cst c) (Var 0)) r) z)) = T" +"plusinf (And p q) = And (plusinf p) (plusinf q)" +"plusinf (Or p q) = Or (plusinf p) (plusinf q)" +"plusinf p = p" + +recdef bset "measure size" +"bset (Le (Add (Mult (Cst c) (Var 0)) r) z) = + (if c < 0 then [lin_add(r,(Cst -1)), r] + else [lin_add(lin_neg r,(Cst -1))])" +"bset (Eq (Add (Mult (Cst c) (Var 0)) r) z) = + (if c < 0 then [lin_add(r,(Cst -1))] + else [lin_add(lin_neg r,(Cst -1))])" +"bset (NOT(Eq (Add (Mult (Cst c) (Var 0)) r) z)) = + (if c < 0 then [r] + else [lin_neg r])" +"bset (And p q) = (bset p) @ (bset q)" +"bset (Or p q) = (bset p) @ (bset q)" +"bset p = []" + +recdef aset "measure size" +"aset (Le (Add (Mult (Cst c) (Var 0)) r) z) = + (if c < 0 then [lin_add (r, Cst 1)] + else [lin_add (lin_neg r, Cst 1), lin_neg r])" +"aset (Eq (Add (Mult (Cst c) (Var 0)) r) z) = + (if c < 0 then [lin_add(r,(Cst 1))] + else [lin_add(lin_neg r,(Cst 1))])" +"aset (NOT(Eq (Add (Mult (Cst c) (Var 0)) r) z)) = + (if c < 0 then [r] + else [lin_neg r])" +"aset (And p q) = (aset p) @ (aset q)" +"aset (Or p q) = (aset p) @ (aset q)" +"aset p = []" + +(* divlcm computes \ = lcm d , where d | x +t occurs in p *) +consts divlcm :: "QF \ int" +recdef divlcm "measure size" +"divlcm (Divides (Cst d) (Add (Mult (Cst c) (Var 0)) r)) = (abs d)" +"divlcm (NOT p) = divlcm p" +"divlcm (And p q)= ilcm (divlcm p) (divlcm q)" +"divlcm (Or p q) = ilcm (divlcm p) (divlcm q)" +"divlcm p = 1" + +(* the preoperty of \ *) +consts alldivide :: "int \ QF \ bool" +recdef alldivide "measure (%(d,p). size p)" +"alldivide (d,(Divides (Cst d') (Add (Mult (Cst c) (Var 0)) r))) = + (d' dvd d)" +"alldivide (d,(NOT p)) = alldivide (d,p)" +"alldivide (d,(And p q)) = (alldivide (d,p) \ alldivide (d,q))" +"alldivide (d,(Or p q)) = ((alldivide (d,p)) \ (alldivide (d,q)))" +"alldivide (d,p) = True" + +(* alldivide is monotone *) +lemma alldivide_mono: "\ d'. \ alldivide (d,p) ; d dvd d'\ \ alldivide (d',p)" +proof(induct d p rule: alldivide.induct, simp_all add: ilcm_dvd1 ilcm_dvd2) + fix "d1" "d2" "d3" + assume th1:"d2 dvd (d1::int)" + and th2: "d1 dvd d3" + show "d2 dvd d3" by (rule zdvd_trans[OF th1 th2]) +qed + +(* Some simple lemmas *) +lemma zdvd_eq_zdvd_abs: " (d::int) dvd d' = (d dvd (abs d')) " +proof- + have "d' < 0 \ d' \ 0" by arith + moreover + { + assume dn': "d' < 0" + then have "abs d' = - d'" by simp + then + have ?thesis by (simp) + } + moreover + { + assume dp': "d' \ 0" + then have "abs d' = d'" by simp + then have ?thesis by simp + } + ultimately show ?thesis by blast +qed + +lemma zdvd_refl_abs: "(d::int) dvd (abs d)" +proof- + have "d dvd d" by simp + then show ?thesis by (simp add: iffD1 [OF zdvd_eq_zdvd_abs [where d = "d" and d'="d"]]) +qed + +(* \ > 0*) +lemma divlcm_pos: + assumes + linp: "islinform p" + shows "0 < divlcm p" +using linp +proof (induct p rule: divlcm.induct,simp_all add: ilcm_pos) + case (goal1 f) show ?case + using prems + by (cases f, auto) (case_tac "intterm1", auto) +qed + +lemma nz_le: "(x::int) > 0 \ x \ 0" by auto +(* divlcm is correct *) +lemma divlcm_corr: + assumes + linp: "islinform p" + shows "alldivide (divlcm p,p)" +thm nz_le[OF divlcm_pos[OF linp]] + using linp divlcm_pos +proof (induct p rule: divlcm.induct,simp_all add: zdvd_refl_abs,clarsimp simp add: Nat.gr0_conv_Suc) + case (goal1 f) + have "islinform f" using prems + by (cases f, auto) (case_tac "intterm2", auto,case_tac "intterm1", auto) + then have "alldivide (divlcm f, f)" using prems by simp + moreover have "divlcm (NOT f) = divlcm f" by simp + moreover have "alldivide (x,f) = alldivide (x,NOT f)" by simp + ultimately show ?case by simp +next + case (goal2 f g) + have dvd1: "(divlcm f) dvd (ilcm (divlcm f) (divlcm g))" + using prems by(simp add: ilcm_dvd1 nz_le) + have dvd2: "(divlcm g) dvd (ilcm (divlcm f) (divlcm g))" + using prems by (simp add: ilcm_dvd2 nz_le) + from dvd1 prems + have "alldivide (ilcm (divlcm f) (divlcm g), f)" + by (simp add: alldivide_mono[where d= "divlcm f" and p="f" and d' ="ilcm (divlcm f) (divlcm g)"]) + moreover from dvd2 prems + have "alldivide (ilcm (divlcm f) (divlcm g), g)" + by (simp add: alldivide_mono[where d= "divlcm g" and p="g" and d' ="ilcm (divlcm f) (divlcm g)"]) + ultimately show ?case by simp +next + case (goal3 f g) + have dvd1: "(divlcm f) dvd (ilcm (divlcm f) (divlcm g))" + using prems by (simp add: nz_le ilcm_dvd1) + have dvd2: "(divlcm g) dvd (ilcm (divlcm f) (divlcm g))" + using prems by (simp add: nz_le ilcm_dvd2) + from dvd1 prems + have "alldivide (ilcm (divlcm f) (divlcm g), f)" + by (simp add: alldivide_mono[where d= "divlcm f" and p="f" and d' ="ilcm (divlcm f) (divlcm g)"]) + moreover from dvd2 prems + have "alldivide (ilcm (divlcm f) (divlcm g), g)" + by (simp add: alldivide_mono[where d= "divlcm g" and p="g" and d' ="ilcm (divlcm f) (divlcm g)"]) + ultimately show ?case by simp +qed + + +(* Properties of minusinf and plusinf*) + +(* minusinf p and p are the same for minusinfity \ *) +lemma minusinf_eq: + assumes unifp: "isunified p" + shows "\ z. \ x. x < z \ (qinterp (x#ats) p = qinterp (x#ats) (minusinf p))" +using unifp unified_islinform[OF unifp] +proof (induct p rule: minusinf.induct) + case (1 c r z) + have "c <0 \ 0 \ c" by arith + moreover + { + assume cneg: " c < 0" + from prems have z0: "z= Cst 0" + by (cases z,auto) + with prems have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" + by simp + + from prems z0 have ?case + proof- + show ?thesis + using prems z0 + apply auto + apply (rule exI[where x="I_intterm (a # ats) r"]) + apply (rule allI) + proof- + fix x + show "x < I_intterm (a # ats) r \ \ - x + I_intterm (x # ats) r \ 0" + by (simp add: intterm_novar0[OF lincnr, where x="a" and y="x"]) + qed + qed + } + moreover + { + assume cpos: "0 \ c" + from prems have z0: "z= Cst 0" + by (cases z) auto + with prems have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" + by simp + + from prems z0 have ?case + proof- + show ?thesis + using prems z0 + apply auto + apply (rule exI[where x="-(I_intterm (a # ats) r)"]) + apply (rule allI) + proof- + fix x + show "x < - I_intterm (a # ats) r \ x + I_intterm (x # ats) r \ 0" + by (simp add: intterm_novar0[OF lincnr, where x="a" and y="x"]) + qed + qed + } + + ultimately show ?case by blast +next + case (2 c r z) + from prems have z0: "z= Cst 0" + by (cases z,auto) + with prems have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" + by simp + have "c <0 \ 0 \ c" by arith + moreover + { + assume cneg: " c < 0" + from prems z0 have ?case + proof- + show ?thesis + using prems z0 + apply auto + apply (rule exI[where x="I_intterm (a # ats) r"]) + apply (rule allI) + proof- + fix x + show "x < I_intterm (a # ats) r \ \ - x + I_intterm (x # ats) r = 0" + by (simp add: intterm_novar0[OF lincnr, where x="a" and y="x"]) + qed + qed + } + moreover + { + assume cpos: "0 \ c" + from prems z0 have ?case + proof- + show ?thesis + using prems z0 + apply auto + apply (rule exI[where x="-(I_intterm (a # ats) r)"]) + apply (rule allI) + proof- + fix x + show "x < - I_intterm (a # ats) r \ x + I_intterm (x # ats) r \ 0" + by (simp add: intterm_novar0[OF lincnr, where x="a" and y="x"]) + qed + qed + } + + ultimately show ?case by blast +next + case (3 c r z) + from prems have z0: "z= Cst 0" + by (cases z,auto) + with prems have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" + by simp + have "c <0 \ 0 \ c" by arith + moreover + { + assume cneg: " c < 0" + from prems z0 have ?case + proof- + show ?thesis + using prems z0 + apply auto + apply (rule exI[where x="I_intterm (a # ats) r"]) + apply (rule allI) + proof- + fix x + show "x < I_intterm (a # ats) r \ \ - x + I_intterm (x # ats) r = 0" + by (simp add: intterm_novar0[OF lincnr, where x="a" and y="x"]) + qed + qed + } + moreover + { + assume cpos: "0 \ c" + from prems z0 have ?case + proof- + show ?thesis + using prems z0 + apply auto + apply (rule exI[where x="-(I_intterm (a # ats) r)"]) + apply (rule allI) + proof- + fix x + show "x < - I_intterm (a # ats) r \ x + I_intterm (x # ats) r \ 0" + by (simp add: intterm_novar0[OF lincnr, where x="a" and y="x"]) + qed + qed + } + + ultimately show ?case by blast +next + + case (4 f g) + from prems obtain "zf" where + zf:"\xxxx i \ 1" by auto + then have im1: "i=-1" by arith + then have "(a dvd i*x + i') = (a dvd x + (-i'))" + by (simp add: uminus_dvd_conv'[where d="a" and t="-x +i'"]) + moreover + from im1 have "(a dvd i*x + (i*(c * d)) + i') = (a dvd (x + c*d - i'))" + apply simp + apply (simp add: uminus_dvd_conv'[where d="a" and t="-x - c * d + i'"]) + by (simp add: zadd_ac) + ultimately + have eq1:"((a dvd i*x + i') = (a dvd i*x + (i*(c * d)) + i')) = + ((a dvd x + (-i')) = (a dvd (x + c*d - i')))" by simp + moreover + have dvd2: "(a dvd x + (-i')) = (a dvd x + c * d + (-i'))" + by (rule dvd_period[where a="a" and d="d" and x="x" and c="c"], assumption) + ultimately show ?case by simp + qed + next + case (goal2 i n i' n' r) + have "n = 0 \ 0 < n" by arith + moreover + { + assume npos: "0 < n" + from prems have "n < n'" by simp then have "0 < n'" by simp + moreover from prems + have linr: "islinintterm (Add (Mult (Cst i') (Var n')) r)" by simp + ultimately have ?case + using prems npos + by (simp add: nth_pos2 intterm_novar0[OF linr,where x="x" and y="x + c*d"]) + } + moreover + { + assume n0: "n=0" + from prems have lin2: "islinintterm (Add (Mult (Cst i') (Var n')) r)" by simp + from prems have "n < n'" by simp then have npos': "0 < n'" by simp + with prems have ?case + proof(simp add: intterm_novar0[OF lin2, where x="x" and y="x+c*d"] + nth_pos2 dvd_period,case_tac "i=1", + simp add: dvd_period[where a="a" and d="d" and x="x" and c="c"], simp) + case goal1 + from prems have "abs i = 1 \ i\1" by auto + then have mi: "i = -1" by arith + have "(a dvd -x + (i' * ats ! (n' - Suc 0) + I_intterm ((x + c * d) # ats) r)) = + (a dvd x + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r))" + by (simp add: + uminus_dvd_conv'[where d="a" and + t="-x + (i' * ats ! (n' - Suc 0) + I_intterm ((x + c * d) # ats) r)"]) + also + have "(a dvd x + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r)) = + (a dvd x +c*d + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r))" + by (rule dvd_period[where a="a" and d="d" and x="x" and c="c"], assumption) + also + have "(a dvd x +c*d + + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r)) = + (a dvd -(x +c*d + + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r)))" + by (rule uminus_dvd_conv'[where d="a" and + t="x +c*d + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r)"]) + also + have "(a dvd -(x +c*d + + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r))) + = (a dvd + - x - c * d + (i' * ats ! (n' - Suc 0) + I_intterm ((x + c * d) # ats) r))" + by (auto,simp_all add: zadd_ac) + finally show ?case using mi by auto + qed + } + ultimately show ?case by blast + qed +next + case (goal4 a t) + show ?case using prems + proof(induct t rule: islinintterm.induct, simp_all,case_tac "n=0", + simp_all add: nth_pos2) + case (goal1 i n i') + show ?case + using prems + proof(case_tac "i=1", simp, + simp add: dvd_period[where a="a" and d="d" and x="x" and c="c"]) + case goal1 + from prems have "abs i = 1 \ i\1" by auto + then have im1: "i=-1" by arith + then have "(a dvd i*x + i') = (a dvd x + (-i'))" + by (simp add: uminus_dvd_conv'[where d="a" and t="-x +i'"]) + moreover + from im1 have "(a dvd i*x + (i*(c * d)) + i') = (a dvd (x + c*d - i'))" + apply simp + apply (simp add: uminus_dvd_conv'[where d="a" and t="-x - c * d + i'"]) + by (simp add: zadd_ac) + ultimately + have eq1:"((a dvd i*x + i') = (a dvd i*x + (i*(c * d)) + i')) = + ((a dvd x + (-i')) = (a dvd (x + c*d - i')))" by simp + moreover + have dvd2: "(a dvd x + (-i')) = (a dvd x + c * d + (-i'))" + by (rule dvd_period[where a="a" and d="d" and x="x" and c="c"], assumption) + ultimately show ?thesis by simp + qed + next + case (goal2 i n i' n' r) + have "n = 0 \ 0 < n" by arith + moreover + { + assume npos: "0 < n" + from prems have "n < n'" by simp then have "0 < n'" by simp + moreover from prems + have linr: "islinintterm (Add (Mult (Cst i') (Var n')) r)" by simp + ultimately have ?case + using prems npos + by (simp add: nth_pos2 intterm_novar0[OF linr,where x="x" and y="x + c*d"]) + } + moreover + { + assume n0: "n=0" + from prems have lin2: "islinintterm (Add (Mult (Cst i') (Var n')) r)" by simp + from prems have "n < n'" by simp then have npos': "0 < n'" by simp + with prems have ?case + proof(simp add: intterm_novar0[OF lin2, where x="x" and y="x+c*d"] + nth_pos2 dvd_period,case_tac "i=1", + simp add: dvd_period[where a="a" and d="d" and x="x" and c="c"], simp) + case goal1 + from prems have "abs i = 1 \ i\1" by auto + then have mi: "i = -1" by arith + have "(a dvd -x + (i' * ats ! (n' - Suc 0) + I_intterm ((x + c * d) # ats) r)) = + (a dvd x + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r))" + by (simp add: + uminus_dvd_conv'[where d="a" and + t="-x + (i' * ats ! (n' - Suc 0) + I_intterm ((x + c * d) # ats) r)"]) + also + have "(a dvd x + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r)) = + (a dvd x +c*d + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r))" + by (rule dvd_period[where a="a" and d="d" and x="x" and c="c"], assumption) + also + have "(a dvd x +c*d + + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r)) = + (a dvd -(x +c*d + + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r)))" + by (rule uminus_dvd_conv'[where d="a" and + t="x +c*d + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r)"]) + also + have "(a dvd -(x +c*d + + (-i' * ats ! (n' - Suc 0) - I_intterm ((x + c * d) # ats) r))) + = (a dvd + - x - c * d + (i' * ats ! (n' - Suc 0) + I_intterm ((x + c * d) # ats) r))" + by (auto,simp_all add: zadd_ac) + finally show ?case using mi by auto + qed + } + ultimately show ?case by blast + qed +next + case (goal5 t a) + show ?case + using prems + apply (cases t, simp_all add: nth_pos2) + apply (case_tac "intterm1", simp_all) + apply (case_tac "intterm1a",simp_all) + by (case_tac "intterm2a",simp_all) + (case_tac "nat",simp_all add: nth_pos2 intterm_novar0[where x="x" and y="x+c*d"]) +qed + +lemma minusinf_repeats2: + assumes alldvd: "alldivide (d,p)" + and unity: "isunified p" + shows "\ x k. (qinterp (x#ats) (minusinf p) = qinterp ((x - k*d)#ats) (minusinf p))" + (is "\ x k. ?P x = ?P (x - k*d)") +proof(rule allI, rule allI) + fix x k + show "?P x = ?P (x - k*d)" + proof- + have "?P x = ?P (x + (-k)*d)" by (rule minusinf_repeats[OF alldvd unity]) + then have "?P x = ?P (x - (k*d))" by simp + then show ?thesis by blast + qed +qed + + +(* existence for minusinf p is existence for p *) +lemma minusinf_lemma: + assumes unifp: "isunified p" + and exminf: "\ j \ {1 ..d}. qinterp (j#ats) (minusinf p)" (is "\ j \ {1 .. d}. ?P1 j") + shows "\ x. qinterp (x#ats) p" (is "\ x. ?P x") +proof- + from exminf obtain "j" where P1j: "?P1 j" by blast + have ePeqP1: "\z. \ x. x < z \ (?P x = ?P1 x)" + by (rule minusinf_eq[OF unifp]) + then obtain "z" where P1eqP : "\ x. x < z \ (?P x = ?P1 x)" by blast + let ?d = "divlcm p" + have alldvd: "alldivide (?d,p)" using unified_islinform[OF unifp] divlcm_corr + by auto + have dpos: "0 < ?d" using unified_islinform[OF unifp] divlcm_pos + by simp + have P1eqP1 : "\ x k. ?P1 x = ?P1 (x - k*(?d))" + by (rule minusinf_repeats2[OF alldvd unifp]) + let ?w = "j - (abs (j-z) +1)* ?d" + show "\ x. ?P x" + proof + have w: "?w < z" + by (rule decr_lemma[OF dpos]) + + have "?P1 j = ?P1 ?w" using P1eqP1 by blast + also have "\ = ?P ?w" using w P1eqP by blast + finally show "?P ?w" using P1j by blast + qed +qed + +(* limited search for the withness for minusinf p, due to peridicity *) +lemma minusinf_disj: + assumes unifp: "isunified p" + shows "(\ x. qinterp (x#ats) (minusinf p)) = + (\ j \ { 1.. divlcm p}. qinterp (j#ats) (minusinf p))" + (is "(\ x. ?P x) = (\ j \ { 1.. ?d}. ?P j)") +proof + have linp: "islinform p" by (rule unified_islinform[OF unifp]) + have dpos: "0 < ?d" by (rule divlcm_pos[OF linp]) + have alldvd: "alldivide(?d,p)" by (rule divlcm_corr[OF linp]) + { + assume "\ j\ {1 .. ?d}. ?P j" + then show "\ x. ?P x" using dpos by auto + next + assume "\ x. ?P x" + then obtain "x" where P: "?P x" by blast + have modd: "\x k. ?P x = ?P (x - k*?d)" + by (rule minusinf_repeats2[OF alldvd unifp]) + + have "x mod ?d = x - (x div ?d)*?d" + by(simp add:zmod_zdiv_equality mult_ac eq_diff_eq) + hence Pmod: "?P x = ?P (x mod ?d)" using modd by simp + show "\ j\ {1 .. ?d}. ?P j" + proof (cases) + assume "x mod ?d = 0" + hence "?P 0" using P Pmod by simp + moreover have "?P 0 = ?P (0 - (-1)*?d)" using modd by blast + ultimately have "?P ?d" by simp + moreover have "?d \ {1 .. ?d}" using dpos + by (simp add:atLeastAtMost_iff) + ultimately show "\ j\ {1 .. ?d}. ?P j" .. + next + assume not0: "x mod ?d \ 0" + have "?P(x mod ?d)" using dpos P Pmod by(simp add:pos_mod_sign pos_mod_bound) + moreover have "x mod ?d : {1 .. ?d}" + proof - + have "0 \ x mod ?d" by(rule pos_mod_sign[OF dpos]) + moreover have "x mod ?d < ?d" by(rule pos_mod_bound[OF dpos]) + ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff) + qed + ultimately show "\ j\ {1 .. ?d}. ?P j" .. + qed + } +qed + +lemma minusinf_qfree: + assumes linp : "islinform p" + shows "isqfree (minusinf p)" + using linp + by (induct p rule: minusinf.induct) auto + +(* Properties of bset and a set *) + +(* The elements of a bset are linear *) +lemma bset_lin: + assumes unifp: "isunified p" + shows "\ b \ set (bset p). islinintterm b" +using unifp unified_islinform[OF unifp] +proof (induct p rule: bset.induct, auto) + case (goal1 c r z) + from prems have "z = Cst 0" by (cases z, simp_all) + then have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have "islinintterm (Cst -1)" by simp + then show ?case using linr lin_add_lin by simp +next + case (goal2 c r z) + from prems have "z = Cst 0" by (cases z, simp_all) + then have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + show ?case by (rule linr) +next + case (goal3 c r z) + from prems have "z = Cst 0" by (cases z, simp_all) + then have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have "islinintterm (Cst -1)" by simp + then show ?case using linr lin_add_lin lin_neg_lin by simp +next + case (goal4 c r z) + from prems have "z = Cst 0" by (cases z, simp_all) + then have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have "islinintterm (Cst -1)" by simp + then show ?case using linr lin_add_lin lin_neg_lin by simp +next + case (goal5 c r z) + from prems have "z = Cst 0" by (cases z, simp_all) + then have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have "islinintterm (Cst -1)" by simp + then show ?case using linr lin_add_lin lin_neg_lin by simp +next + case (goal6 c r z) + from prems have "z = Cst 0" by (cases z, simp_all) + then have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have "islinintterm (Cst -1)" by simp + then show ?case using linr lin_add_lin lin_neg_lin by simp +next + case (goal7 c r z) + from prems have "z = Cst 0" by (cases z, simp_all) + then have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" using prems by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have "islinintterm (Cst -1)" by simp + then show ?case using linr lin_add_lin lin_neg_lin by simp +qed + +(* The third lemma in Norrisch's Paper *) +lemma bset_disj_repeat: + assumes unifp: "isunified p" + and alldvd: "alldivide (d,p)" + and dpos: "0 < d" + and nob: "(qinterp (x#ats) q) \ \(\j\ {1 .. d}. \ b \ set (bset p). (qinterp (((I_intterm (a#ats) b) + j)#ats) q)) \(qinterp (x#ats) p)" + (is "?Q x \ \(\ j\ {1.. d}. \ b\ ?B. ?Q (?I a b + j)) \ ?P x") + shows "?P (x -d)" + using unifp nob alldvd unified_islinform[OF unifp] +proof (induct p rule: islinform.induct,auto) + case (goal1 t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have "n=0 \ n>0" by arith + moreover {assume "n>0" then have ?case + using prems + by (simp add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="x-d"]) } + moreover + {assume nz: "n = 0" + from prems have "abs i = 1" by auto + then have "i = -1 \ i =1" by arith + moreover + { + assume i1: "i=1" + have ?case using dpos prems + by (auto simp add: intterm_novar0[OF lininr, where x="x" and y="x - d"]) + } + moreover + { + assume im1: "i = -1" + have ?case + using prems + proof(auto simp add: intterm_novar0[OF lininr, where x="x - d" and y="x"], cases) + assume "- x + d + ?I x r \ 0" + then show "- x + d + ?I x r \ 0" . + next + assume np: "\ - x + d + ?I x r \ 0" + then have ltd:"x - ?I x r \ d - 1" by simp + from prems have "-x + ?I x r \ 0" by simp + then have ge0: "x - ?I x r \ 0" + by simp + from ltd ge0 have "x - ?I x r = 0 \ (1 \ x - ?I x r \ x - ?I x r \ d - 1) " by arith + moreover + { + assume "x - ?I x r = 0" + then have xeqr: "x = ?I x r" by simp + from prems have "?Q x" by simp + with xeqr have qr:"?Q (?I x r)" by simp + from prems have lininr: "islinintterm (Add (Mult (Cst i) (Var 0)) r)" by simp + have "islinintterm r" by (rule islinintterm_subt[OF lininr]) + from prems + have "\j\{1..d}. \ ?Q (?I a r + -1 + j)" + using linr by (auto simp add: lin_add_corr) + moreover from dpos have "1 \ {1..d}" by simp + ultimately have " \ ?Q (?I a r + -1 + 1)" by blast + with dpos linr have "\ ?Q (?I x r)" + by (simp add: intterm_novar0[OF lininr, where x="x" and y="a"] lin_add_corr) + with qr have "- x + d + ?I x r \ 0" by simp + } + moreover + { + assume gt0: "1 \ x - ?I x r \ x - ?I x r \ d - 1" + then have "\ j\ {1 .. d - 1}. x - ?I x r = j" by simp + then have "\ j\ {1 .. d}. x - ?I x r = j" by auto + then obtain "j" where con: "1\j \ j \ d \ x - ?I x r = j" by auto + then have xeqr: "x = ?I x r + j" by auto + with prems have "?Q (?I x r + j)" by simp + with con have qrpj: "\ j\ {1 .. d}. ?Q (?I x r + j)" by auto + from prems have "\j\{1..d}. \ ?Q (?I a r + j)" by auto + then have "\ (\ j\{1..d}. ?Q (?I x r + j))" + by (simp add: intterm_novar0[OF lininr, where x="x" and y="a"]) + with qrpj prems have "- x + d + ?I x r \ 0" by simp + + } + ultimately show "- x + d + ?I x r \ 0" by blast + qed + } + ultimately have ?case by blast + } + ultimately have ?case by blast + } + ultimately show ?case by blast +next + case (goal3 a t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have "n=0 \ n>0" by arith + moreover {assume "n>0" then have ?case using prems + by (simp add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="x-d"]) } + moreover { + assume nz: "n=0" + from prems have "abs i = 1" by auto + then have ipm: "i=1 \ i = -1" by arith + from nz prems have advdixr: "a dvd (i * x) + I_intterm (x # ats) r" + by simp + from prems have "a dvd d" by simp + then have advdid: "a dvd i*d" using ipm by auto + have ?case + using prems ipm + by (auto simp add: intterm_novar0[OF lininr, where x="x-d" and y="x"] dvd_period[OF advdid, where x="i*x" and c="-1"]) + } + ultimately have ?case by blast + } ultimately show ?case by blast +next + + case (goal4 a t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + + have "n=0 \ n>0" by arith + moreover {assume "n>0" then have ?case using prems + by (simp add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="x-d"]) } + moreover { + assume nz: "n=0" + from prems have "abs i = 1" by auto + then have ipm: "i =1 \ i = -1" by arith + from nz prems have advdixr: "\ (a dvd (i * x) + I_intterm (x # ats) r)" + by simp + from prems have "a dvd d" by simp + then have advdid: "a dvd i*d" using ipm by auto + have ?case + using prems ipm + by (auto simp add: intterm_novar0[OF lininr, where x="x-d" and y="x"] dvd_period[OF advdid, where x="i*x" and c="-1"]) + } + ultimately have ?case by blast + } ultimately show ?case by blast +next + case (goal2 t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have "n=0 \ n>0" by arith + moreover {assume "n>0" then have ?case + using prems + by (simp add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="x-d"]) } + moreover + {assume nz: "n = 0" + from prems have "abs i = 1" by auto + then have "i = -1 \ i =1" by arith + moreover + { + assume i1: "i=1" + with prems have px: "x + ?I x r = 0" by simp + then have "x = (- ?I x r - 1) + 1" by simp + hence q1: "?Q ((- ?I x r - 1) + 1)" by simp + from prems have "\ (?Q ((?I a (lin_add(lin_neg r, Cst -1))) + 1))" + by auto + hence "\ (?Q ((- ?I a r - 1) + 1))" + using lin_add_corr lin_neg_corr linr lin_neg_lin + by simp + hence "\ (?Q ((- ?I x r - 1) + 1))" + using intterm_novar0[OF lininr, where x="x" and y="a"] + by simp + with q1 have ?case by simp + } + moreover + { + assume im1: "i = -1" + with prems have px: "-x + ?I x r = 0" by simp + then have "x = ?I x r" by simp + hence q1: "?Q (?I x r)" by simp + from prems have "\ (?Q ((?I a (lin_add(r, Cst -1))) + 1))" + by auto + hence "\ (?Q (?I a r))" + using lin_add_corr lin_neg_corr linr lin_neg_lin + by simp + hence "\ (?Q (?I x r ))" + using intterm_novar0[OF lininr, where x="x" and y="a"] + by simp + with q1 have ?case by simp + } + ultimately have ?case by blast + } + ultimately have ?case by blast + } + ultimately show ?case by blast +next + case (goal5 t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have "n=0 \ n>0" by arith + moreover {assume "n>0" then have ?case + using prems + by (simp add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="x-d"]) } + moreover + {assume nz: "n = 0" + from prems have "abs i = 1" by auto + then have "i = -1 \ i =1" by arith + moreover + { + assume i1: "i=1" + with prems have px: "x -d + ?I (x-d) r = 0" by simp + hence "x = (- ?I x r) + d" + using intterm_novar0[OF lininr, where x="x" and y="x-d"] + by simp + hence q1: "?Q (- ?I x r + d)" by simp + from prems have "\ (?Q ((?I a (lin_neg r)) + d))" + by auto + hence "\ (?Q (- ?I a r + d))" + using lin_neg_corr linr by simp + hence "\ (?Q ((- ?I x r + d)))" + using intterm_novar0[OF lininr, where x="x" and y="a"] + by simp + with q1 have ?case by simp + } + moreover + { + assume im1: "i = -1" + with prems have px: "- (x -d) + ?I (x - d) r = 0" by simp + then have "x = ?I x r + d " + using intterm_novar0[OF lininr, where x="x" and y="x-d"] + by simp + hence q1: "?Q (?I x r + d)" by simp + from prems have "\ (?Q ((?I a r) + d))" + by auto + hence "\ (?Q (?I x r + d))" + using intterm_novar0[OF lininr, where x="x" and y="a"] + by simp + with q1 have ?case by simp + } + ultimately have ?case by blast + } + ultimately have ?case by blast + } + ultimately show ?case by blast + +qed + +lemma bset_disj_repeat2: + assumes unifp: "isunified p" + + shows "\ x. \(\j\ {1 .. (divlcm p)}. \ b \ set (bset p). + (qinterp (((I_intterm (a#ats) b) + j)#ats) p)) + \ (qinterp (x#ats) p) \ (qinterp ((x - (divlcm p))#ats) p)" + (is "\ x. \(\ j\ {1 .. ?d}. \ b\ ?B. ?P (?I a b + j)) \ ?P x \ ?P (x - ?d)") +proof + fix x + have linp: "islinform p" by (rule unified_islinform[OF unifp]) + have dpos: "?d > 0" by (rule divlcm_pos[OF linp]) + have alldvd: "alldivide(?d,p)" by (rule divlcm_corr[OF linp]) + show "\(\ j\ {1 .. ?d}. \ b\ ?B. ?P (?I a b + j)) \ ?P x \ ?P (x - ?d)" + using prems bset_disj_repeat[OF unifp alldvd dpos] + by blast +qed + +(* Cooper's theorem in the minusinfinity version *) +lemma cooper_mi_eq: + assumes unifp : "isunified p" + shows "(\ x. qinterp (x#ats) p) = + ((\ j \ {1 .. (divlcm p)}. qinterp (j#ats) (minusinf p)) \ + (\ j \ {1 .. (divlcm p)}. \ b \ set (bset p). + qinterp (((I_intterm (a#ats) b) + j)#ats) p))" + (is "(\ x. ?P x) = ((\ j\ {1 .. ?d}. ?MP j) \ (\ j \ ?D. \ b\ ?B. ?P (?I a b + j)))") +proof- + have linp :"islinform p" by (rule unified_islinform[OF unifp]) + have dpos: "?d > 0" by (rule divlcm_pos[OF linp]) + have alldvd: "alldivide(?d,p)" by (rule divlcm_corr[OF linp]) + have eMPimpeP: "(\j \ ?D. ?MP j) \ (\x. ?P x)" + by (simp add: minusinf_lemma[OF unifp, where d="?d" and ats="ats"]) + have ePimpeP: "(\ j \ ?D. \ b\ ?B. ?P (?I a b + j)) \ (\ x. ?P x)" + by blast + have bst_rep: "\ x. \ (\ j \ ?D. \ b \ ?B. ?P (?I a b + j)) \ ?P x \ ?P (x - ?d)" + by (rule bset_disj_repeat2[OF unifp]) + have MPrep: "\ x k. ?MP x = ?MP (x- k*?d)" + by (rule minusinf_repeats2[OF alldvd unifp]) + have MPeqP: "\ z. \ x < z. ?P x = ?MP x" + by (rule minusinf_eq[OF unifp]) + let ?B'= "{?I a b| b. b\ ?B}" + from bst_rep have bst_rep2: "\x. \ (\j\?D. \b\ ?B'. ?P (b+j)) \ ?P x \ ?P (x - ?d)" + by auto + show ?thesis + using cpmi_eq[OF dpos MPeqP bst_rep2 MPrep] + by auto +qed + +(* A formalized analogy between aset, bset, plusinfinity and minusinfinity *) + +consts mirror:: "QF \ QF" +recdef mirror "measure size" +"mirror (Le (Add (Mult (Cst c) (Var 0)) r) z) = + (Le (Add (Mult (Cst (- c)) (Var 0)) r) z)" +"mirror (Eq (Add (Mult (Cst c) (Var 0)) r) z) = + (Eq (Add (Mult (Cst (- c)) (Var 0)) r) z)" +"mirror (Divides (Cst d) (Add (Mult (Cst c) (Var 0)) r)) = + (Divides (Cst d) (Add (Mult (Cst (- c)) (Var 0)) r))" +"mirror (NOT(Divides (Cst d) (Add (Mult (Cst c) (Var 0)) r))) = + (NOT(Divides (Cst d) (Add (Mult (Cst (- c)) (Var 0)) r)))" +"mirror (NOT(Eq (Add (Mult (Cst c) (Var 0)) r) z)) = + (NOT(Eq (Add (Mult (Cst (- c)) (Var 0)) r) z))" +"mirror (And p q) = And (mirror p) (mirror q)" +"mirror (Or p q) = Or (mirror p) (mirror q)" +"mirror p = p" +(* mirror preserves unifiedness *) + +lemma[simp]: "(abs (i::int) = 1) = (i =1 \ i = -1)" by arith +lemma mirror_unified: + assumes unif: "isunified p" + shows "isunified (mirror p)" + using unif +proof (induct p rule: mirror.induct, simp_all) + case (goal1 c r z) + from prems have zz: "z = Cst 0" by (cases z, simp_all) + then show ?case using prems + by (auto simp add: islinintterm_eq_islint islint_def) +next + case (goal2 c r z) + from prems have zz: "z = Cst 0" by (cases z, simp_all) + then show ?case using prems + by (auto simp add: islinintterm_eq_islint islint_def) +next + case (goal3 d c r) show ?case using prems by (auto simp add: islinintterm_eq_islint islint_def) +next + case (goal4 d c r) show ?case using prems by (auto simp add: islinintterm_eq_islint islint_def) +next + case (goal5 c r z) + from prems have zz: "z = Cst 0" by (cases z, simp_all) + then show ?case using prems + by (auto simp add: islinintterm_eq_islint islint_def) +qed + +(* relationship between plusinf and minusinf *) +lemma plusinf_eq_minusinf_mirror: + assumes unifp: "isunified p" + shows "(qinterp (x#ats) (plusinf p)) = (qinterp ((- x)#ats) (minusinf (mirror p)))" +using unifp unified_islinform[OF unifp] +proof (induct p rule: islinform.induct, simp_all) + case (goal1 t z) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems + by (cases n, auto simp add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"] )} + ultimately show ?case by blast + +next + case (goal2 t z) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems + by (cases n, auto simp add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"] )} + ultimately show ?case by blast +next + case (goal3 d t) + + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + + have ?case using prems + by (cases n, simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"] )} + ultimately show ?case by blast +next + + case (goal4 d t) + + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + + have ?case using prems + by (cases n, simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"] )} + ultimately show ?case by blast +next + case (goal5 t z) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems + by (cases n, auto simp add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"] )} + ultimately show ?case by blast +qed + +(* relationship between aset abd bset *) +lemma aset_eq_bset_mirror: + assumes unifp: "isunified p" + shows "set (aset p) = set (map lin_neg (bset (mirror p)))" +using unifp +proof(induct p rule: mirror.induct) + case (1 c r z) + from prems have zz: "z = Cst 0" + by (cases z, auto) + from prems zz have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have neg1eqm1: "Cst 1 = lin_neg (Cst -1)" by (simp add: lin_neg_def) + have negm1eq1: "Cst -1 = lin_neg (Cst 1)" by (simp add: lin_neg_def) + show ?case using prems linr zz apply (auto simp add: lin_neg_lin_add_distrib lin_neg_idemp neg1eqm1) + by (simp add: negm1eq1 lin_neg_idemp sym[OF lin_neg_lin_add_distrib] lin_add_lin) +next + case (2 c r z) from prems have zz: "z = Cst 0" + by (cases z, auto) + from prems zz have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have neg1eqm1: "Cst 1 = lin_neg (Cst -1)" by (simp add: lin_neg_def) + have negm1eq1: "Cst -1 = lin_neg (Cst 1)" by (simp add: lin_neg_def) + show ?case using prems linr zz + by (auto simp add: lin_neg_lin_add_distrib lin_neg_idemp neg1eqm1) + (simp add: negm1eq1 lin_neg_idemp sym[OF lin_neg_lin_add_distrib] lin_add_lin lin_neg_lin) + +next + case (5 c r z) from prems have zz: "z = Cst 0" + by (cases z, auto) + from prems zz have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have neg1eqm1: "Cst 1 = lin_neg (Cst -1)" by (simp add: lin_neg_def) + have negm1eq1: "Cst -1 = lin_neg (Cst 1)" by (simp add: lin_neg_def) + show ?case using prems linr zz + by(auto simp add: lin_neg_lin_add_distrib lin_neg_idemp neg1eqm1) + +qed simp_all + +(* relationship between aset abd bset 2*) +lemma aset_eq_bset_mirror2: + assumes unifp: "isunified p" + shows "aset p = map lin_neg (bset (mirror p))" +using unifp +proof(induct p rule: mirror.induct) + case (1 c r z) + from prems have zz: "z = Cst 0" + by (cases z, auto) + from prems zz have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have neg1eqm1: "Cst 1 = lin_neg (Cst -1)" by (simp add: lin_neg_def) + have negm1eq1: "Cst -1 = lin_neg (Cst 1)" by (simp add: lin_neg_def) + show ?case using prems linr zz + apply (simp add: lin_neg_lin_add_distrib lin_neg_idemp neg1eqm1) + apply (simp add: negm1eq1 lin_neg_idemp sym[OF lin_neg_lin_add_distrib] lin_add_lin) + by arith +next + case (2 c r z) from prems have zz: "z = Cst 0" + by (cases z, auto) + from prems zz have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have neg1eqm1: "Cst 1 = lin_neg (Cst -1)" by (simp add: lin_neg_def) + have negm1eq1: "Cst -1 = lin_neg (Cst 1)" by (simp add: lin_neg_def) + show ?case using prems linr zz + by(auto simp add: lin_neg_lin_add_distrib lin_neg_idemp neg1eqm1) + (simp add: negm1eq1 lin_neg_idemp sym[OF lin_neg_lin_add_distrib] lin_add_lin lin_neg_lin) + +next + case (5 c r z) from prems have zz: "z = Cst 0" + by (cases z, auto) + from prems zz have lincnr: "islinintterm (Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have neg1eqm1: "Cst 1 = lin_neg (Cst -1)" by (simp add: lin_neg_def) + have negm1eq1: "Cst -1 = lin_neg (Cst 1)" by (simp add: lin_neg_def) + show ?case using prems linr zz + by(auto simp add: lin_neg_lin_add_distrib lin_neg_idemp neg1eqm1) + +qed simp_all + +(* mirror preserves divlcm *) +lemma divlcm_mirror_eq: + assumes unifp: "isunified p" + shows "divlcm p = divlcm (mirror p)" + using unifp +by (induct p rule: mirror.induct) auto + +(* mirror almost preserves semantics *) +lemma mirror_interp: + assumes unifp: "isunified p" + shows "(qinterp (x#ats) p) = (qinterp ((- x)#ats) (mirror p))" (is "?P x = ?MP (-x)") +using unifp unified_islinform[OF unifp] +proof (induct p rule: islinform.induct) + case (1 t z) + from prems have zz: "z = 0" by simp + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems zz + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +next + case (2 t z) + from prems have zz: "z = 0" by simp + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems zz + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +next + case (3 d t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case + using prems linr + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +next + + case (6 d t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case + using prems linr + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +next + case (7 t z) + from prems have zz: "z = 0" by simp + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems zz + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +qed simp_all + + +lemma mirror_interp2: + assumes unifp: "islinform p" + shows "(qinterp (x#ats) p) = (qinterp ((- x)#ats) (mirror p))" (is "?P x = ?MP (-x)") +using unifp +proof (induct p rule: islinform.induct) + case (1 t z) + from prems have zz: "z = 0" by simp + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems zz + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +next + case (2 t z) + from prems have zz: "z = 0" by simp + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems zz + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +next + case (3 d t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case + using prems linr + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +next + + case (6 d t) + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case + using prems linr + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +next + case (7 t z) + from prems have zz: "z = 0" by simp + from prems + have lint: "islinintterm t" by simp + then have "(\ i n r. t = Add (Mult (Cst i) (Var n) ) r) \ (\ i. t = Cst i)" + by (induct t rule: islinintterm.induct) auto + moreover{ assume "\ i. t = Cst i" then have ?case using prems by auto } + moreover + { assume "\ i n r. t = Add (Mult (Cst i) (Var n) ) r" + then obtain "i" "n" "r" where + inr_def: "t = Add (Mult (Cst i) (Var n) ) r" + by blast + with lint have lininr: "islinintterm (Add (Mult (Cst i) (Var n) ) r)" + by simp + have linr: "islinintterm r" + by (rule islinintterm_subt[OF lininr]) + have ?case using prems zz + by (cases n) (simp_all add: nth_pos2 + intterm_novar0[OF lininr, where x="x" and y="-x"]) + } + ultimately show ?case by blast +qed simp_all + +(* mirror preserves existence *) +lemma mirror_ex: + assumes unifp: "isunified p" + shows "(\ x. (qinterp (x#ats) p)) = (\ y. (qinterp (y#ats) (mirror p)))" + (is "(\ x. ?P x) = (\ y. ?MP y)") +proof + assume "\ x. ?P x" + then obtain "x" where px:"?P x" by blast + have "?MP (-x)" + using px + by(simp add: mirror_interp[OF unifp, where x="x"]) + then show "\ y. ?MP y" by blast +next + assume "\ y. ?MP y" + then obtain "y" where mpy: "?MP y" by blast + have "?P (-y)" + using mpy + by (simp add: mirror_interp[OF unifp, where x="-y"]) + then show "\ x. ?P x" by blast +qed + +lemma mirror_ex2: + assumes unifp: "isunified p" + shows "qinterp ats (QEx p) = qinterp ats (QEx (mirror p))" +using mirror_ex[OF unifp] by simp + + +(* Cooper's theorem in its plusinfinity version *) +lemma cooper_pi_eq: + assumes unifp : "isunified p" + shows "(\ x. qinterp (x#ats) p) = + ((\ j \ {1 .. (divlcm p)}. qinterp (-j#ats) (plusinf p)) \ + (\ j \ {1 .. (divlcm p)}. \ b \ set (aset p). + qinterp (((I_intterm (a#ats) b) - j)#ats) p))" + (is "(\ x. ?P x) = ((\ j\ {1 .. ?d}. ?PP (-j)) \ (\ j \ ?D. \ b\ ?A. ?P (?I a b - j)))") +proof- + have unifmp: "isunified (mirror p)" by (rule mirror_unified[OF unifp]) + have th1: + "(\ j\ {1 .. ?d}. ?PP (-j)) = (\ j\ {1..?d}. qinterp (j # ats) (minusinf (mirror p)))" + by (simp add: plusinf_eq_minusinf_mirror[OF unifp]) + have dth: "?d = divlcm (mirror p)" + by (rule divlcm_mirror_eq[OF unifp]) + have "(\ j \ ?D. \ b\ ?A. ?P (?I a b - j)) = + (\ j\ ?D. \ b \ set (map lin_neg (bset (mirror p))). ?P (?I a b - j))" + by (simp only: aset_eq_bset_mirror[OF unifp]) + also have "\ = (\ j\ ?D. \ b \ set (bset (mirror p)). ?P (?I a (lin_neg b) - j))" + by simp + also have "\ = (\ j\ ?D. \ b \ set (bset (mirror p)). ?P (-(?I a b + j)))" + proof + assume "\j\{1..divlcm p}. + \b\set (bset (mirror p)). qinterp ((I_intterm (a # ats) (lin_neg b) - j) # ats) p" + then + obtain "j" and "b" where + pbmj: "j\ ?D \ b\ set (bset (mirror p)) \ ?P (?I a (lin_neg b) - j)" by blast + then have linb: "islinintterm b" + by (auto simp add:bset_lin[OF unifmp]) + from linb pbmj have "?P (-(?I a b + j))" by (simp add: lin_neg_corr) + then show "\ j\ ?D. \ b \ set (bset (mirror p)). ?P (-(?I a b + j))" + using pbmj + by auto + next + assume "\ j\ ?D. \ b \ set (bset (mirror p)). ?P (-(?I a b + j))" + then obtain "j" and "b" where + pbmj: "j\ ?D \ b\ set (bset (mirror p)) \ ?P (-(?I a b + j))" + by blast + then have linb: "islinintterm b" + by (auto simp add:bset_lin[OF unifmp]) + from linb pbmj have "?P (?I a (lin_neg b) - j)" + by (simp add: lin_neg_corr) + then show "\ j\ ?D. \ b \ set (bset (mirror p)). ?P (?I a (lin_neg b) - j)" + using pbmj by auto + qed + finally + have bth: "(\ j\ ?D. \ b\ ?A. ?P (?I a b - j)) = + (\j\ ?D. \ b\set (bset (mirror p)). + qinterp ((I_intterm (a # ats) b + j) # ats) (mirror p))" + by (simp add: mirror_interp[OF unifp] zadd_ac) + from bth dth th1 + have "(\ x. ?P x) = (\ x. qinterp (x#ats) (mirror p))" + by (simp add: mirror_ex[OF unifp]) + also have "\ = ((\j\{1..divlcm (mirror p)}. qinterp (j # ats) (minusinf (mirror p))) \ + (\j\{1..divlcm (mirror p)}. + \b\set (bset (mirror p)). qinterp ((I_intterm (a # ats) b + j) # ats) (mirror p)))" + (is "(\ x. ?MP x) = ((\ j\ ?DM. ?MPM j) \ (\ j \ ?DM. \ b\ ?BM. ?MP (?I a b + j)))") + by (rule cooper_mi_eq[OF unifmp]) + also + have "\ = ((\ j\ ?D. ?PP (-j)) \ (\ j \ ?D. \ b\ ?BM. ?MP (?I a b + j)))" + using bth th1 dth by simp + finally show ?thesis using sym[OF bth] by simp +qed + + +(* substitution of a term into a Qfree formula, substitution of Bound 0 by i*) + +consts subst_it:: "intterm \ intterm \ intterm" +primrec +"subst_it i (Cst b) = Cst b" +"subst_it i (Var n) = (if n = 0 then i else Var n)" +"subst_it i (Neg it) = Neg (subst_it i it)" +"subst_it i (Add it1 it2) = Add (subst_it i it1) (subst_it i it2)" +"subst_it i (Sub it1 it2) = Sub (subst_it i it1) (subst_it i it2)" +"subst_it i (Mult it1 it2) = Mult (subst_it i it1) (subst_it i it2)" + + +(* subst_it preserves semantics *) +lemma subst_it_corr: +"I_intterm (a#ats) (subst_it i t) = I_intterm ((I_intterm (a#ats) i)#ats) t" +by (induct t rule: subst_it.induct, simp_all add: nth_pos2) + +consts subst_p:: "intterm \ QF \ QF" +primrec +"subst_p i (Le it1 it2) = Le (subst_it i it1) (subst_it i it2)" +"subst_p i (Lt it1 it2) = Lt (subst_it i it1) (subst_it i it2)" +"subst_p i (Ge it1 it2) = Ge (subst_it i it1) (subst_it i it2)" +"subst_p i (Gt it1 it2) = Gt (subst_it i it1) (subst_it i it2)" +"subst_p i (Eq it1 it2) = Eq (subst_it i it1) (subst_it i it2)" +"subst_p i (Divides d t) = Divides (subst_it i d) (subst_it i t)" +"subst_p i T = T" +"subst_p i F = F" +"subst_p i (And p q) = And (subst_p i p) (subst_p i q)" +"subst_p i (Or p q) = Or (subst_p i p) (subst_p i q)" +"subst_p i (Imp p q) = Imp (subst_p i p) (subst_p i q)" +"subst_p i (Equ p q) = Equ (subst_p i p) (subst_p i q)" +"subst_p i (NOT p) = (NOT (subst_p i p))" + +(* subs_p preserves correctness *) +lemma subst_p_corr: + assumes qf: "isqfree p" + shows "qinterp (a # ats) (subst_p i p) = qinterp ((I_intterm (a#ats) i)#ats) p " + using qf +by (induct p rule: subst_p.induct) (simp_all add: subst_it_corr) + +(* novar0 p is true if the fomula doese not depend on the quantified variable*) +consts novar0I:: "intterm \ bool" +primrec +"novar0I (Cst i) = True" +"novar0I (Var n) = (n > 0)" +"novar0I (Neg a) = (novar0I a)" +"novar0I (Add a b) = (novar0I a \ novar0I b)" +"novar0I (Sub a b) = (novar0I a \ novar0I b)" +"novar0I (Mult a b) = (novar0I a \ novar0I b)" + +consts novar0:: "QF \ bool" +recdef novar0 "measure size" +"novar0 (Lt a b) = (novar0I a \ novar0I b)" +"novar0 (Gt a b) = (novar0I a \ novar0I b)" +"novar0 (Le a b) = (novar0I a \ novar0I b)" +"novar0 (Ge a b) = (novar0I a \ novar0I b)" +"novar0 (Eq a b) = (novar0I a \ novar0I b)" +"novar0 (Divides a b) = (novar0I a \ novar0I b)" +"novar0 T = True" +"novar0 F = True" +"novar0 (NOT p) = novar0 p" +"novar0 (And p q) = (novar0 p \ novar0 q)" +"novar0 (Or p q) = (novar0 p \ novar0 q)" +"novar0 (Imp p q) = (novar0 p \ novar0 q)" +"novar0 (Equ p q) = (novar0 p \ novar0 q)" +"novar0 p = False" + +(* Interpretation of terms, that doese not depend on Var 0 *) +lemma I_intterm_novar0: + assumes nov0: "novar0I x" + shows "I_intterm (a#ats) x = I_intterm (b#ats) x" +using nov0 +by (induct x) (auto simp add: nth_pos2) + +(* substition is meaningless for term independent of Var 0*) +lemma subst_p_novar0_corr: +assumes qfp: "isqfree p" + and nov0: "novar0I i" + shows "qinterp (a#ats) (subst_p i p) = qinterp (I_intterm (b#ats) i#ats) p" +proof- + have "qinterp (a#ats) (subst_p i p) = qinterp (I_intterm (a#ats) i#ats) p" + by (rule subst_p_corr[OF qfp]) + moreover have "I_intterm (a#ats) i#ats = I_intterm (b#ats) i#ats" + by (simp add: I_intterm_novar0[OF nov0, where a="a" and b="b"]) + ultimately show ?thesis by simp +qed + +(* linearity and independence on Var 0*) +lemma lin_novar0: + assumes linx: "islinintterm x" + and nov0: "novar0I x" + shows "\ n > 0. islintn(n,x)" +using linx nov0 +by (induct x rule: islinintterm.induct) auto + +lemma lintnpos_novar0: + assumes npos: "n > 0" + and linx: "islintn(n,x)" + shows "novar0I x" +using npos linx +by (induct n x rule: islintn.induct) auto + +(* lin_add preserves independence on Var 0*) +lemma lin_add_novar0: + assumes nov0a: "novar0I a" + and nov0b : "novar0I b" + and lina : "islinintterm a" + and linb: "islinintterm b" + shows "novar0I (lin_add (a,b))" +proof- + have "\ na > 0. islintn(na, a)" by (rule lin_novar0[OF lina nov0a]) + then obtain "na" where na: "na > 0 \ islintn(na,a)" by blast + have "\ nb > 0. islintn(nb, b)" by (rule lin_novar0[OF linb nov0b]) + then obtain "nb" where nb: "nb > 0 \ islintn(nb,b)" by blast + from na have napos: "na > 0" by simp + from na have linna: "islintn(na,a)" by simp + from nb have nbpos: "nb > 0" by simp + from nb have linnb: "islintn(nb,b)" by simp + have "min na nb \ min na nb" by simp + then have "islintn (min na nb, lin_add(a,b))" by (simp add: lin_add_lint[OF linna linnb]) + moreover have "min na nb > 0" using napos nbpos by (simp add: min_def) + ultimately show ?thesis by (simp only: lintnpos_novar0) +qed + +(* lin__mul preserves independence on Var 0*) +lemma lin_mul_novar0: + assumes linx: "islinintterm x" + and nov0: "novar0I x" + shows "novar0I (lin_mul(i,x))" + using linx nov0 +proof (induct i x rule: lin_mul.induct, auto) + case (goal1 c c' n r) + from prems have lincnr: "islinintterm (Add (Mult (Cst c') (Var n)) r)" by simp + have "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + then show ?case using prems by simp +qed + +(* lin_neg preserves indepenednce on Var 0*) +lemma lin_neg_novar0: + assumes linx: "islinintterm x" + and nov0: "novar0I x" + shows "novar0I (lin_neg x)" +by (auto simp add: lin_mul_novar0 linx nov0 lin_neg_def) + +(* subterms of linear terms are independent on Var 0*) +lemma intterm_subt_novar0: + assumes lincnr: "islinintterm (Add (Mult (Cst c) (Var n)) r)" + shows "novar0I r" +proof- + have cnz: "c \ 0" by (rule islinintterm_cnz[OF lincnr]) + have "islintn(0,Add (Mult (Cst c) (Var n)) r)" using lincnr + by (simp only: islinintterm_eq_islint islint_def) + then have "islintn (n+1,r)" by auto + moreover have "n+1 >0 " by arith + ultimately show ?thesis + using lintnpos_novar0 + by auto +qed + +(* decrease the De-Bruijn indices*) +consts decrvarsI:: "intterm \ intterm" +primrec +"decrvarsI (Cst i) = (Cst i)" +"decrvarsI (Var n) = (Var (n - 1))" +"decrvarsI (Neg a) = (Neg (decrvarsI a))" +"decrvarsI (Add a b) = (Add (decrvarsI a) (decrvarsI b))" +"decrvarsI (Sub a b) = (Sub (decrvarsI a) (decrvarsI b))" +"decrvarsI (Mult a b) = (Mult (decrvarsI a) (decrvarsI b))" + +(* One can decrease the indics for terms and formulae independent on Var 0*) +lemma intterm_decrvarsI: + assumes nov0: "novar0I t" + shows "I_intterm (a#ats) t = I_intterm ats (decrvarsI t)" +using nov0 +by (induct t) (auto simp add: nth_pos2) + +consts decrvars:: "QF \ QF" +primrec +"decrvars (Lt a b) = (Lt (decrvarsI a) (decrvarsI b))" +"decrvars (Gt a b) = (Gt (decrvarsI a) (decrvarsI b))" +"decrvars (Le a b) = (Le (decrvarsI a) (decrvarsI b))" +"decrvars (Ge a b) = (Ge (decrvarsI a) (decrvarsI b))" +"decrvars (Eq a b) = (Eq (decrvarsI a) (decrvarsI b))" +"decrvars (Divides a b) = (Divides (decrvarsI a) (decrvarsI b))" +"decrvars T = T" +"decrvars F = F" +"decrvars (NOT p) = (NOT (decrvars p))" +"decrvars (And p q) = (And (decrvars p) (decrvars q))" +"decrvars (Or p q) = (Or (decrvars p) (decrvars q))" +"decrvars (Imp p q) = (Imp (decrvars p) (decrvars q))" +"decrvars (Equ p q) = (Equ (decrvars p) (decrvars q))" + +(* decrvars preserves quantifier freeness*) +lemma decrvars_qfree: "isqfree p \ isqfree (decrvars p)" +by (induct p rule: isqfree.induct, auto) + +lemma novar0_qfree: "novar0 p \ isqfree p" +by (induct p) auto + +lemma qinterp_novar0: + assumes nov0: "novar0 p" + shows "qinterp (a#ats) p = qinterp ats (decrvars p)" +using nov0 +by(induct p) (simp_all add: intterm_decrvarsI) + +(* All elements of bset p doese not depend on Var 0*) +lemma bset_novar0: + assumes unifp: "isunified p" + shows "\ b\ set (bset p). novar0I b " + using unifp +proof(induct p rule: bset.induct) + case (1 c r z) + from prems have zz: "z = Cst 0" by (cases "z", auto) + from prems zz have lincnr: "islinintterm(Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have novar0r: "novar0I r" by (rule intterm_subt_novar0[OF lincnr]) + from prems zz have "c = 1 \ c = -1" by auto + moreover + { + assume c1: "c=1" + have lin1: "islinintterm (Cst 1)" by simp + have novar01: "novar0I (Cst 1)" by simp + then have ?case + using prems zz novar0r lin1 novar01 + by (auto simp add: lin_add_novar0 lin_neg_novar0 linr lin_neg_lin) + } + moreover + { + assume c1: "c= -1" + have lin1: "islinintterm (Cst -1)" by simp + have novar01: "novar0I (Cst -1)" by simp + then have ?case + using prems zz novar0r lin1 novar01 + by (auto simp add: lin_add_novar0 lin_neg_novar0 linr lin_neg_lin) + } + ultimately show ?case by blast +next + case (2 c r z) + from prems have zz: "z = Cst 0" by (cases "z", auto) + from prems zz have lincnr: "islinintterm(Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have novar0r: "novar0I r" by (rule intterm_subt_novar0[OF lincnr]) + from prems zz have "c = 1 \ c = -1" by auto + moreover + { + assume c1: "c=1" + have lin1: "islinintterm (Cst 1)" by simp + have novar01: "novar0I (Cst 1)" by simp + then have ?case + using prems zz novar0r lin1 novar01 + by (auto simp add: lin_add_novar0 lin_neg_novar0 linr lin_neg_lin) + } + moreover + { + assume c1: "c= -1" + have lin1: "islinintterm (Cst -1)" by simp + have novar01: "novar0I (Cst -1)" by simp + then have ?case + using prems zz novar0r lin1 novar01 + by (auto simp add: lin_add_novar0 lin_neg_novar0 linr lin_neg_lin) + } + ultimately show ?case by blast +next + case (3 c r z) + from prems have zz: "z = Cst 0" by (cases "z", auto) + from prems zz have lincnr: "islinintterm(Add (Mult (Cst c) (Var 0)) r)" by simp + have linr: "islinintterm r" by (rule islinintterm_subt[OF lincnr]) + have novar0r: "novar0I r" by (rule intterm_subt_novar0[OF lincnr]) + from prems zz have "c = 1 \ c = -1" by auto + moreover + { + assume c1: "c=1" + have lin1: "islinintterm (Cst 1)" by simp + have novar01: "novar0I (Cst 1)" by simp + then have ?case + using prems zz novar0r lin1 novar01 + by (auto simp add: lin_add_novar0 lin_neg_novar0 linr lin_neg_lin) + } + moreover + { + assume c1: "c= -1" + have lin1: "islinintterm (Cst -1)" by simp + have novar01: "novar0I (Cst -1)" by simp + then have ?case + using prems zz novar0r lin1 novar01 + by (auto simp add: lin_add_novar0 lin_neg_novar0 linr lin_neg_lin) + } + ultimately show ?case by blast +qed auto + +(* substitution preserves independence on Var 0*) +lemma subst_it_novar0: + assumes nov0x: "novar0I x" + shows "novar0I (subst_it x t)" + using nov0x + by (induct t) auto + +lemma subst_p_novar0: + assumes nov0x:"novar0I x" + and qfp: "isqfree p" + shows "novar0 (subst_p x p)" + using nov0x qfp + by (induct p rule: novar0.induct) (simp_all add: subst_it_novar0) + +(* linearize preserves independence on Var 0 *) +lemma linearize_novar0: + assumes nov0t: "novar0I t " + shows "\ t'. linearize t = Some t' \ novar0I t'" +using nov0t +proof(induct t rule: novar0I.induct) + case (Neg a) + let ?la = "linearize a" + from prems have "\ a'. ?la = Some a'" by (cases ?la, auto) + then obtain "a'" where "?la = Some a'" by blast + with prems have nv0a':"novar0I a'" by simp + have "islinintterm a'" using prems by (simp add: linearize_linear) + with nv0a' have "novar0I (lin_neg a')" + by (simp add: lin_neg_novar0) + then + show ?case using prems by simp +next + case (Add a b) + let ?la = "linearize a" + let ?lb = "linearize b" + from prems have linab: "linearize (Add a b) = Some t'" by simp + then have "\ a'. ?la = Some a'" by (cases ?la) auto + then obtain "a'" where "?la = Some a'" by blast + with prems have nv0a':"novar0I a'" by simp + have lina': "islinintterm a'" using prems by (simp add: linearize_linear) + from linab have "\ b'. ?lb = Some b'" + by (cases ?la, auto simp add: measure_def inv_image_def) (cases ?lb, auto) + then obtain "b'" where "?lb = Some b'" by blast + with prems have nv0b':"novar0I b'" by simp + have linb': "islinintterm b'" using prems by (simp add: linearize_linear) + then show ?case using prems lina' linb' nv0a' nv0b' + by (auto simp add: measure_def inv_image_def lin_add_novar0) +next + case (Sub a b) + let ?la = "linearize a" + let ?lb = "linearize b" + from prems have linab: "linearize (Sub a b) = Some t'" by simp + then have "\ a'. ?la = Some a'" by (cases ?la) auto + then obtain "a'" where "?la = Some a'" by blast + with prems have nv0a':"novar0I a'" by simp + have lina': "islinintterm a'" using prems by (simp add: linearize_linear) + from linab have "\ b'. ?lb = Some b'" + by (cases ?la, auto simp add: measure_def inv_image_def) (cases ?lb, auto) + then obtain "b'" where "?lb = Some b'" by blast + with prems have nv0b':"novar0I b'" by simp + have linb': "islinintterm b'" using prems by (simp add: linearize_linear) + then show ?case using prems lina' linb' nv0a' nv0b' + by (auto simp add: + measure_def inv_image_def lin_add_novar0 lin_neg_novar0 lin_neg_lin) +next + case (Mult a b) + let ?la = "linearize a" + let ?lb = "linearize b" + from prems have linab: "linearize (Mult a b) = Some t'" by simp + then have "\ a'. ?la = Some a'" + by (cases ?la, auto simp add: measure_def inv_image_def) + then obtain "a'" where "?la = Some a'" by blast + with prems have nv0a':"novar0I a'" by simp + have lina': "islinintterm a'" using prems by (simp add: linearize_linear) + from prems linab have "\ b'. ?lb = Some b'" + apply (cases ?la, auto simp add: measure_def inv_image_def) + by (cases "a'",auto simp add: measure_def inv_image_def) (cases ?lb, auto)+ + then obtain "b'" where "?lb = Some b'" by blast + with prems have nv0b':"novar0I b'" by simp + have linb': "islinintterm b'" using prems by (simp add: linearize_linear) + then show ?case using prems lina' linb' nv0a' nv0b' + by (cases "a'",auto simp add: measure_def inv_image_def lin_mul_novar0) + (cases "b'",auto simp add: measure_def inv_image_def lin_mul_novar0) +qed auto + + +(* simplification of formulae *) +consts psimpl :: "QF \ QF" +recdef psimpl "measure size" +"psimpl (Le l r) = + (case (linearize (Sub l r)) of + None \ Le l r + | Some x \ (case x of + Cst i \ (if i \ 0 then T else F) + | _ \ (Le x (Cst 0))))" +"psimpl (Eq l r) = + (case (linearize (Sub l r)) of + None \ Eq l r + | Some x \ (case x of + Cst i \ (if i = 0 then T else F) + | _ \ (Eq x (Cst 0))))" + +"psimpl (Divides (Cst d) t) = + (case (linearize t) of + None \ (Divides (Cst d) t) + | Some c \ (case c of + Cst i \ (if d dvd i then T else F) + | _ \ (Divides (Cst d) c)))" + +"psimpl (And p q) = + (let p'= psimpl p + in (case p' of + F \ F + |T \ psimpl q + | _ \ let q' = psimpl q + in (case q' of + F \ F + | T \ p' + | _ \ (And p' q'))))" + +"psimpl (Or p q) = + (let p'= psimpl p + in (case p' of + T \ T + | F \ psimpl q + | _ \ let q' = psimpl q + in (case q' of + T \ T + | F \ p' + | _ \ (Or p' q'))))" + +"psimpl (Imp p q) = + (let p'= psimpl p + in (case p' of + F \ T + |T \ psimpl q + | NOT p1 \ let q' = psimpl q + in (case q' of + F \ p1 + | T \ T + | _ \ (Or p1 q')) + | _ \ let q' = psimpl q + in (case q' of + F \ NOT p' + | T \ T + | _ \ (Imp p' q'))))" + +"psimpl (Equ p q) = + (let p'= psimpl p ; q' = psimpl q + in (case p' of + T \ q' + | F \ (case q' of + T \ F + | F \ T + | NOT q1 \ q1 + | _ \ NOT q') + | NOT p1 \ (case q' of + T \ p' + | F \ p1 + | NOT q1 \ (Equ p1 q1) + | _ \ (Equ p' q')) + | _ \ (case q' of + T \ p' + | F \ NOT p' + | _ \ (Equ p' q'))))" + +"psimpl (NOT p) = + (let p' = psimpl p + in ( case p' of + F \ T + | T \ F + | NOT p1 \ p1 + | _ \ (NOT p')))" +"psimpl p = p" + +(* psimpl preserves semantics *) +lemma psimpl_corr: "qinterp ats p = qinterp ats (psimpl p)" +proof(induct p rule: psimpl.induct) + case (1 l r) + have "(\ lx. linearize (Sub l r) = Some lx) \ (linearize (Sub l r) = None)" by auto + moreover + { + assume lin: "\ lx. linearize (Sub l r) = Some lx" + from lin obtain "lx" where lx: "linearize (Sub l r) = Some lx" by blast + from lx have "I_intterm ats (Sub l r) = I_intterm ats lx" + by (rule linearize_corr[where t="Sub l r" and t'= "lx"]) + then have feq: "qinterp ats (Le l r) = qinterp ats (Le lx (Cst 0))" by (simp , arith) + from lx have lxlin: "islinintterm lx" by (rule linearize_linear) + from lxlin feq have ?case + proof- + have "(\ i. lx = Cst i) \ (\ (\ i. lx = Cst i))" by blast + moreover + { + assume lxcst: "\ i. lx = Cst i" + from lxcst obtain "i" where lxi: "lx = Cst i" by blast + with feq have "qinterp ats (Le l r) = (i \ 0)" by simp + then have ?case using prems by (simp add: measure_def inv_image_def) + } + moreover + { + assume "(\ (\ i. lx = Cst i))" + then have "(case lx of + Cst i \ (if i \ 0 then T else F) + | _ \ (Le lx (Cst 0))) = (Le lx (Cst 0))" + by (case_tac "lx::intterm", auto) + with prems lxlin feq have ?case by (auto simp add: measure_def inv_image_def) + } + ultimately show ?thesis by blast + qed + } + moreover + { + assume "linearize (Sub l r) = None" + then have ?case using prems by simp + } + ultimately show ?case by blast + +next + case (2 l r) + have "(\ lx. linearize (Sub l r) = Some lx) \ (linearize (Sub l r) = None)" by auto + moreover + { + assume lin: "\ lx. linearize (Sub l r) = Some lx" + from lin obtain "lx" where lx: "linearize (Sub l r) = Some lx" by blast + from lx have "I_intterm ats (Sub l r) = I_intterm ats lx" + by (rule linearize_corr[where t="Sub l r" and t'= "lx"]) + then have feq: "qinterp ats (Eq l r) = qinterp ats (Eq lx (Cst 0))" by (simp , arith) + from lx have lxlin: "islinintterm lx" by (rule linearize_linear) + from lxlin feq have ?case + proof- + have "(\ i. lx = Cst i) \ (\ (\ i. lx = Cst i))" by blast + moreover + { + assume lxcst: "\ i. lx = Cst i" + from lxcst obtain "i" where lxi: "lx = Cst i" by blast + with feq have "qinterp ats (Eq l r) = (i = 0)" by simp + then have ?case using prems by (simp add: measure_def inv_image_def) + } + moreover + { + assume "(\ (\ i. lx = Cst i))" + then have "(case lx of + Cst i \ (if i = 0 then T else F) + | _ \ (Eq lx (Cst 0))) = (Eq lx (Cst 0))" + by (case_tac "lx::intterm", auto) + with prems lxlin feq have ?case by (auto simp add: measure_def inv_image_def) + } + ultimately show ?thesis by blast + qed + } + moreover + { + assume "linearize (Sub l r) = None" + then have ?case using prems by simp + } + ultimately show ?case by blast + +next + + case (3 d t) + have "(\ lt. linearize t = Some lt) \ (linearize t = None)" by auto + moreover + { + assume lin: "\ lt. linearize t = Some lt" + from lin obtain "lt" where lt: "linearize t = Some lt" by blast + from lt have "I_intterm ats t = I_intterm ats lt" + by (rule linearize_corr[where t="t" and t'= "lt"]) + then have feq: "qinterp ats (Divides (Cst d) t) = qinterp ats (Divides (Cst d) lt)" by (simp) + from lt have ltlin: "islinintterm lt" by (rule linearize_linear) + from ltlin feq have ?case using prems apply simp by (case_tac "lt::intterm", simp_all) + } + moreover + { + assume "linearize t = None" + then have ?case using prems by simp + } + ultimately show ?case by blast + +next + case (4 f g) + + let ?sf = "psimpl f" + let ?sg = "psimpl g" + show ?case using prems + by (cases ?sf, simp_all add: Let_def measure_def inv_image_def) + (cases ?sg, simp_all)+ +next + case (5 f g) + let ?sf = "psimpl f" + let ?sg = "psimpl g" + show ?case using prems + apply (cases ?sf, simp_all add: Let_def measure_def inv_image_def) + apply (cases ?sg, simp_all) + apply (cases ?sg, simp_all) + apply (cases ?sg, simp_all) + apply (cases ?sg, simp_all) + apply (cases ?sg, simp_all) + apply (cases ?sg, simp_all) + apply (cases ?sg, simp_all) + apply blast + apply (cases ?sg, simp_all) + apply (cases ?sg, simp_all) + apply (cases ?sg, simp_all) + apply blast + apply (cases ?sg, simp_all) + by (cases ?sg, simp_all) (cases ?sg, simp_all) +next + case (6 f g) + let ?sf = "psimpl f" + let ?sg = "psimpl g" + show ?case using prems + apply(simp add: Let_def measure_def inv_image_def) + apply(cases ?sf,simp_all) + apply (simp_all add: Let_def measure_def inv_image_def) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply blast + apply blast + apply blast + apply blast + apply blast + apply blast + apply blast + apply blast + apply blast + apply blast + apply blast + apply blast + apply blast + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + apply(cases ?sg, simp_all) + done +next + case (7 f g) + let ?sf = "psimpl f" + let ?sg = "psimpl g" + show ?case + using prems + by (cases ?sf, simp_all add: Let_def) (cases ?sg, simp_all)+ +next + case (8 f) show ?case + using prems + apply (simp add: Let_def) + by (case_tac "psimpl f", simp_all) +qed simp_all + +(* psimpl preserves independence on Var 0*) +lemma psimpl_novar0: + assumes nov0p: "novar0 p" + shows "novar0 (psimpl p)" + using nov0p +proof (induct p rule: psimpl.induct) + case (1 l r) + let ?ls = "linearize (Sub l r)" + have "?ls = None \ (\ x. ?ls = Some x)" by auto + moreover + { + assume "?ls = None" then have ?case + using prems by (simp add: measure_def inv_image_def) + } + moreover { + assume "\ x. ?ls = Some x" + then obtain "x" where ls_d: "?ls = Some x" by blast + from prems have "novar0I l" by simp + moreover from prems have "novar0I r" by simp + ultimately have nv0s: "novar0I (Sub l r)" by simp + from prems have "novar0I x" + by (simp add: linearize_novar0[OF nv0s, where t'="x"]) + then have ?case + using prems + by (cases "x") (auto simp add: measure_def inv_image_def) + } + ultimately show ?case by blast +next + case (2 l r) + let ?ls = "linearize (Sub l r)" + have "?ls = None \ (\ x. ?ls = Some x)" by auto + moreover + { + assume "?ls = None" then have ?case + using prems by (simp add: measure_def inv_image_def) + } + moreover { + assume "\ x. ?ls = Some x" + then obtain "x" where ls_d: "?ls = Some x" by blast + from prems have "novar0I l" by simp + moreover from prems have "novar0I r" by simp + ultimately have nv0s: "novar0I (Sub l r)" by simp + from prems have "novar0I x" + by (simp add: linearize_novar0[OF nv0s, where t'="x"]) + then have ?case + using prems + by (cases "x") (auto simp add: measure_def inv_image_def) + } + ultimately show ?case by blast +next + case (3 d t) + let ?lt = "linearize t" + have "?lt = None \ (\ x. ?lt = Some x)" by auto + moreover + { assume "?lt = None" then have ?case using prems by simp } + moreover { + assume "\x. ?lt = Some x" + then obtain "x" where x_d: "?lt = Some x" by blast + from prems have nv0t: "novar0I t" by simp + with x_d have "novar0I x" + by (simp add: linearize_novar0[OF nv0t]) + with prems have ?case + by (cases "x") simp_all + } + ultimately show ?case by blast +next + case (4 f g) + let ?sf = "psimpl f" + let ?sg = "psimpl g" + show ?case + using prems + by (cases ?sf, simp_all add: Let_def measure_def inv_image_def) + (cases ?sg,simp_all)+ +next + case (5 f g) + let ?sf = "psimpl f" + let ?sg = "psimpl g" + show ?case + using prems + by (cases ?sf, simp_all add: Let_def measure_def inv_image_def) + (cases ?sg,simp_all)+ +next + case (6 f g) + let ?sf = "psimpl f" + let ?sg = "psimpl g" + show ?case + using prems + by (cases ?sf, simp_all add: Let_def measure_def inv_image_def) + (cases ?sg,simp_all)+ +next + case (7 f g) + let ?sf = "psimpl f" + let ?sg = "psimpl g" + show ?case + using prems + by (cases ?sf, simp_all add: Let_def measure_def inv_image_def) + (cases ?sg,simp_all)+ + +next + case (8 f) + let ?sf = "psimpl f" + from prems have nv0sf:"novar0 ?sf" by simp + show ?case using prems nv0sf + by (cases ?sf, auto simp add: Let_def measure_def inv_image_def) +qed simp_all + +(* implements a disj of p applied to all elements of the list*) +consts explode_disj :: "(intterm list \ QF) \ QF" +recdef explode_disj "measure (\(is,p). length is)" +"explode_disj ([],p) = F" +"explode_disj (i#is,p) = + (let pi = psimpl (subst_p i p) + in ( case pi of + T \ T + | F \ explode_disj (is,p) + | _ \ (let r = explode_disj (is,p) + in (case r of + T \ T + | F \ pi + | _ \ Or pi r))))" + +(* correctness theorem for one iteration of explode_disj *) +lemma explode_disj_disj: + assumes qfp: "isqfree p" + shows "(qinterp (x#xs) (explode_disj(i#is,p))) = + (qinterp (x#xs) (subst_p i p) \ (qinterp (x#xs) (explode_disj(is,p))))" + using qfp +proof- + let ?pi = "psimpl (subst_p i p)" + have pi: "qinterp (x#xs) ?pi = qinterp (x#xs) (subst_p i p)" + by (simp add: psimpl_corr[where p="(subst_p i p)"]) + let ?dp = "explode_disj(is,p)" + show ?thesis using pi + proof (cases) + assume "?pi= T \ ?pi = F" + then show ?thesis using pi by (case_tac "?pi::QF", auto) + + next + assume notTF: "\ (?pi = T \ ?pi = F)" + let ?dp = "explode_disj(is,p)" + have dp_cases: "explode_disj(i#is,p) = + (case (explode_disj(is,p)) of + T \ T + | F \ psimpl (subst_p i p) + | _ \ Or (psimpl (subst_p i p)) (explode_disj(is,p)))" using notTF + by (cases "?pi") + (simp_all add: Let_def cong del: QF.weak_case_cong) + show ?thesis using pi dp_cases notTF + proof(cases) + assume "?dp = T \ ?dp = F" + then show ?thesis + using pi dp_cases + by (cases "?dp") auto + next + assume "\ (?dp = T \ ?dp = F)" + then show ?thesis using pi dp_cases notTF + by (cases ?dp) auto + qed + qed +qed + +(* correctness theorem for explode_disj *) +lemma explode_disj_corr: + assumes qfp: "isqfree p" + shows "(\ x \ set xs. qinterp (a#ats) (subst_p x p)) = + (qinterp (a#ats) (explode_disj(xs,p)))" (is "(\ x \ set xs. ?P x) = (?DP a xs )") + using qfp + proof (induct xs) + case Nil show ?case by simp + next + case (Cons y ys) + have "(\ x \ set (y#ys). ?P x) = (?P y \ (\ x\ set ys. ?P x))" + by auto + also have "\ = (?P y \ ?DP a ys)" using "Cons.hyps" qfp by auto + also have "\ = ?DP a (y#ys)" using explode_disj_disj[OF qfp] by auto + finally show ?case by simp +qed + +(* explode_disj preserves independence on Var 0*) +lemma explode_disj_novar0: + assumes nov0xs: "\x \ set xs. novar0I x" + and qfp: "isqfree p" + shows "novar0 (explode_disj (xs,p))" + using nov0xs qfp +proof (induct xs, auto simp add: Let_def) + case (goal1 a as) + let ?q = "subst_p a p" + let ?qs = "psimpl ?q" + have "?qs = T \ ?qs = F \ (?qs \ T \ ?qs \ F)" by simp + moreover + { assume "?qs = T" then have ?case by simp } + moreover + { assume "?qs = F" then have ?case by simp } + moreover + { + assume qsnTF: "?qs \ T \ ?qs \ F" + let ?r = "explode_disj (as,p)" + have nov0qs: "novar0 ?qs" + using prems + by (auto simp add: psimpl_novar0 subst_p_novar0) + have "?r = T \ ?r = F \ (?r \ T \ ?r \ F)" by simp + moreover + { assume "?r = T" then have ?case by (cases ?qs) auto } + moreover + { assume "?r = F" then have ?case using nov0qs by (cases ?qs, auto) } + moreover + { assume "?r \ T \ ?r \ F" then have ?case using nov0qs prems qsnTF + by (cases ?qs, auto simp add: Let_def) (cases ?r,auto)+ + } + ultimately have ?case by blast + } + ultimately show ?case by blast +qed + +(* Some simple lemmas used for technical reasons *) +lemma eval_Or_cases: + "qinterp (a#ats) (case f of + T \ T + | F \ g + | _ \ (case g of + T \ T + | F \ f + | _ \ Or f g)) = (qinterp (a#ats) f \ qinterp (a#ats) g)" +proof- + let ?result = " + (case f of + T \ T + | F \ g + | _ \ (case g of + T \ T + | F \ f + | _ \ Or f g))" + have "f = T \ f = F \ (f \ T \ f\ F)" by auto + moreover + { + assume fT: "f = T" + then have ?thesis by auto + } + moreover + { + assume "f=F" + then have ?thesis by auto + } + moreover + { + assume fnT: "f\T" + and fnF: "f\F" + have "g = T \ g = F \ (g \ T \ g\ F)" by auto + moreover + { + assume "g=T" + then have ?thesis using fnT fnF by (cases f, auto) + } + moreover + { + assume "g=F" + then have ?thesis using fnT fnF by (cases f, auto) + } + moreover + { + assume gnT: "g\T" + and gnF: "g\F" + then have "?result = (case g of + T \ T + | F \ f + | _ \ Or f g)" + using fnT fnF + by (cases f, auto) + also have "\ = Or f g" + using gnT gnF + by (cases g, auto) + finally have "?result = Or f g" by simp + then + have ?thesis by simp + } + ultimately have ?thesis by blast + + } + + ultimately show ?thesis by blast +qed + +lemma or_case_novar0: + assumes fnTF: "f \ T \ f \ F" + and gnTF: "g \ T \ g \ F" + and f0: "novar0 f" + and g0: "novar0 g" + shows "novar0 + (case f of T \ T | F \ g + | _ \ (case g of T \ T | F \ f | _ \ Or f g))" +using fnTF gnTF f0 g0 +by (cases f, auto) (cases g, auto)+ + + +(* An implementation of sets trough lists *) +constdefs list_insert :: "'a \ 'a list \ 'a list" + "list_insert x xs \ (if x mem xs then xs else x#xs)" + +lemma list_insert_set: "set (list_insert x xs) = set (x#xs)" +by(induct xs) (auto simp add: list_insert_def) + +consts list_union :: "('a list \ 'a list) \ 'a list" + +recdef list_union "measure (\(xs,ys). length xs)" +"list_union ([], ys) = ys" +"list_union (xs, []) = xs" +"list_union (x#xs,ys) = list_insert x (list_union (xs,ys))" + +lemma list_union_set: "set (list_union(xs,ys)) = set (xs@ys)" + by(induct xs ys rule: list_union.induct, auto simp add:list_insert_set) + + +consts list_set ::"'a list \ 'a list" +primrec + "list_set [] = []" + "list_set (x#xs) = list_insert x (list_set xs)" + +lemma list_set_set: "set xs = set (list_set xs)" +by (induct xs) (auto simp add: list_insert_set) + +consts iupto :: "int \ int \ int list" +recdef iupto "measure (\ (i,j). nat (j - i +1))" +"iupto(i,j) = (if j intterm list \ intterm list" +recdef all_sums "measure (\(i,is). length is)" +"all_sums (j,[]) = []" +"all_sums (j,i#is) = (map (\x. lin_add (i,(Cst x))) (iupto(1,j))@(all_sums (j,is)))" +(* all_sums preserves independence on Var 0*) +lemma all_sums_novar0: + assumes nov0xs: "\ x\ set xs. novar0I x" + and linxs: "\ x\ set xs. islinintterm x " + shows "\ x\ set (all_sums (d,xs)). novar0I x" + using nov0xs linxs +proof(induct d xs rule: all_sums.induct) + case 1 show ?case by simp +next + case (2 j a as) + have lina: "islinintterm a" using "2.prems" by auto + have nov0a: "novar0I a" using "2.prems" by auto + let ?ys = "map (\x. lin_add (a,(Cst x))) (iupto(1,j))" + have nov0ys: "\ y\ set ?ys. novar0I y" + proof- + have linx: "\ x \ set (iupto(1,j)). islinintterm (Cst x)" by simp + have nov0x: "\ x \ set (iupto(1,j)). novar0I (Cst x)" by simp + with nov0a lina linx have "\ x\ set (iupto(1,j)). novar0I (lin_add (a,Cst x))" + by (simp add: lin_add_novar0) + then show ?thesis by auto + qed + from "2.prems" + have linas: "\u\set as. islinintterm u" by auto + from "2.prems" have nov0as: "\u\set as. novar0I u" by auto + from "2.hyps" linas nov0as have nov0alls: "\u\set (all_sums (j, as)). novar0I u" by simp + from nov0alls nov0ys have + cs: "(\ u\ set (?ys@ (all_sums (j,as))). novar0I u)" + by (simp only: sym[OF list_all_iff]) auto + + have "all_sums(j,a#as) = ?ys@(all_sums(j,as))" + by simp + then + have "?case = (\ x\ set (?ys@ (all_sums (j,as))). novar0I x)" + by auto + with cs show ?case by blast +qed + +(* correctness theorem for all_sums*) +lemma all_sums_ex: + "(\ j\ {1..d}. \ b\ (set xs). P (lin_add(b,Cst j))) = + (\ x\ set (all_sums (d,xs)). P x)" +proof(induct d xs rule: all_sums.induct) + case (1 a) show ?case by simp +next + case (2 a y ys) + have "(\ x\ set (map (\x. lin_add (y,(Cst x))) (iupto(1,a))) . P x) = + (\ j\ set (iupto(1,a)). P (lin_add(y,Cst j)))" + by auto + also have "\ = (\ j\ {1..a}. P (lin_add(y,Cst j)))" + by (simp only : iupto_set) + finally + have dsj1:"(\j\{1..a}. P (lin_add (y, Cst j))) = (\x\set (map (\x. lin_add (y, Cst x)) (iupto (1, a))). P x)" by simp + + from prems have "(\ j\ {1..a}. \ b\ (set (y#ys)). P (lin_add(b,Cst j))) = + ((\ j\ {1..a}. P (lin_add(y,Cst j))) \ (\ j\ {1..a}. \ b \ set ys. P (lin_add(b,Cst j))))" by auto + also + have " \ = ((\ j\ {1..a}. P (lin_add(y,Cst j))) \ (\ x\ set (all_sums(a, ys)). P x))" using prems by simp + also have "\ = ((\x\set (map (\x. lin_add (y, Cst x)) (iupto (1, a))). P x) \ (\x\set (all_sums (a, ys)). P x))" using dsj1 by simp + also have "\ = (\ x\ (set (map (\x. lin_add (y, Cst x)) (iupto (1, a)))) \ (set (all_sums(a, ys))). P x)" by blast + finally show ?case by simp +qed + + + +(* explode_minf (p,B) assumes that p is unified and B = bset p, it computes the rhs of cooper_mi_eq*) + +consts explode_minf :: "(QF \ intterm list) \ QF" +recdef explode_minf "measure size" +"explode_minf (q,B) = + (let d = divlcm q; + pm = minusinf q; + dj1 = explode_disj ((map Cst (iupto (1, d))),pm) + in (case dj1 of + T \ T + | F \ explode_disj (all_sums (d,B),q) + | _ \ (let dj2 = explode_disj (all_sums (d,B),q) + in ( case dj2 of + T \ T + | F \ dj1 + | _ \ Or dj1 dj2))))" + +(* The result of the rhs of cooper's theorem doese not depend on Var 0*) +lemma explode_minf_novar0: + assumes unifp : "isunified p" + and bst: "set (bset p) = set B" + shows "novar0 (explode_minf (p,B))" +proof- + let ?d = "divlcm p" + let ?pm = "minusinf p" + let ?dj1 = "explode_disj (map Cst (iupto(1,?d)),?pm)" + + have qfpm: "isqfree ?pm" using unified_islinform[OF unifp] minusinf_qfree by simp + have dpos: "?d >0" using unified_islinform[OF unifp] divlcm_pos by simp + have "\ x\ set (map Cst (iupto(1,?d))). novar0I x" by auto + then have dj1_nov0: "novar0 ?dj1" using qfpm explode_disj_novar0 by simp + + let ?dj2 = "explode_disj (all_sums (?d,B),p)" + have + bstlin: "\b\set B. islinintterm b" + using bset_lin[OF unifp] bst + by simp + + have bstnov0: "\b\set B. novar0I b" + using bst bset_novar0[OF unifp] by simp + have allsnov0: "\x\set (all_sums(?d,B)). novar0I x " + by (simp add:all_sums_novar0[OF bstnov0 bstlin] ) + then have dj2_nov0: "novar0 ?dj2" + using explode_disj_novar0 unified_isqfree[OF unifp] bst by simp + have "?dj1 = T \ ?dj1 = F \ (?dj1 \ T \ ?dj1 \ F)" by auto + moreover + { assume "?dj1 = T" then have ?thesis by simp } + moreover + { assume "?dj1 = F" then have ?thesis using bst dj2_nov0 by (simp add: Let_def)} + moreover + { + assume dj1nFT:"?dj1 \ T \ ?dj1 \ F" + + have "?dj2 = T \ ?dj2 = F \ (?dj2 \ T \ ?dj2 \ F)" by auto + moreover + { assume "?dj2 = T" then have ?thesis by (cases ?dj1) simp_all } + moreover + { assume "?dj2 = F" then have ?thesis using dj1_nov0 bst + by (cases ?dj1) (simp_all add: Let_def)} + moreover + { + assume dj2_nTF:"?dj2 \ T \ ?dj2 \ F" + let ?res = "\f. \g. (case f of T \ T | F \ g + | _ \ (case g of T \ T| F \ f| _ \ Or f g))" + have expth: "explode_minf (p,B) = ?res ?dj1 ?dj2" + by (simp add: Let_def del: iupto.simps split del: split_if + cong del: QF.weak_case_cong) + then have ?thesis + using prems or_case_novar0 [OF dj1nFT dj2_nTF dj1_nov0 dj2_nov0] + by (simp add: Let_def del: iupto.simps cong del: QF.weak_case_cong) + } + ultimately have ?thesis by blast + } + ultimately show ?thesis by blast +qed + +(* explode_minf computes the rhs of cooper's thm*) +lemma explode_minf_corr: + assumes unifp : "isunified p" + and bst: "set (bset p) = set B" + shows "(\ x . qinterp (x#ats) p) = (qinterp (a#ats) (explode_minf (p,B)))" + (is "(\ x. ?P x) = (?EXP a p)") +proof- + let ?d = "divlcm p" + let ?pm = "minusinf p" + let ?dj1 = "explode_disj (map Cst (iupto(1,?d)),?pm)" + have qfpm: "isqfree ?pm" using unified_islinform[OF unifp] minusinf_qfree by simp + have nnfp: "isnnf p" by (rule unified_isnnf[OF unifp]) + + have "(\j\{1..?d}. qinterp (j # ats) (minusinf p)) + = (\j\ set (iupto(1,?d)). qinterp (j#ats) (minusinf p))" + (is "(\ j\ {1..?d}. ?QM j) = \") + by (simp add: sym[OF iupto_set] ) + also + have "\ =(\j\ set (iupto(1,?d)). qinterp ((I_intterm (a#ats) (Cst j))#ats) (minusinf p))" + by simp + also have + "\ = (\j\ set (map Cst (iupto(1,?d))). qinterp ((I_intterm (a#ats) j)#ats) (minusinf p))" by simp + also have + "\ = + (\j\ set (map Cst (iupto(1,?d))). qinterp (a#ats) (subst_p j (minusinf p)))" + by (simp add: subst_p_corr[OF qfpm]) + finally have dj1_thm: + "(\ j\ {1..?d}. ?QM j) = (qinterp (a#ats) ?dj1)" + by (simp only: explode_disj_corr[OF qfpm]) + let ?dj2 = "explode_disj (all_sums (?d,B),p)" + have + bstlin: "\b\set B. islinintterm b" + using bst by (simp add: bset_lin[OF unifp]) + have bstnov0: "\b\set B. novar0I b" + using bst by (simp add: bset_novar0[OF unifp]) + have allsnov0: "\x\set (all_sums(?d,B)). novar0I x " + by (simp add:all_sums_novar0[OF bstnov0 bstlin] ) + have "(\ j\ {1..?d}. \ b\ set B. ?P (I_intterm (a#ats) b + j)) = + (\ j\ {1..?d}. \ b\ set B. ?P (I_intterm (a#ats) (lin_add(b,Cst j))))" + using bst by (auto simp add: lin_add_corr bset_lin[OF unifp]) + also have "\ = (\ x \ set (all_sums (?d, B)). ?P (I_intterm (a#ats) x))" + by (simp add: all_sums_ex[where P="\ t. ?P (I_intterm (a#ats) t)"]) + finally + have "(\ j\ {1..?d}. \ b\ set B. ?P (I_intterm (a#ats) b + j)) = + (\ x \ set (all_sums (?d, B)). qinterp (a#ats) (subst_p x p))" + using allsnov0 prems linform_isqfree unified_islinform[OF unifp] + by (simp add: all_sums_ex subst_p_corr) + also have "\ = (qinterp (a#ats) ?dj2)" + using linform_isqfree unified_islinform[OF unifp] + by (simp add: explode_disj_corr) + finally have dj2th: + "(\ j\ {1..?d}. \ b\ set B. ?P (I_intterm (a#ats) b + j)) = + (qinterp (a#ats) ?dj2)" by simp + let ?result = "\f. \g. + (case f of + T \ T + | F \ g + | _ \ (case g of + T \ T + | F \ f + | _ \ Or f g))" + have "?EXP a p = qinterp (a#ats) (?result ?dj1 ?dj2)" + by (simp only: explode_minf.simps Let_def) + also + have "\ = (qinterp (a#ats) ?dj1 \ qinterp (a#ats) ?dj2)" + by (rule eval_Or_cases[where f="?dj1" and g="?dj2" and a="a" and ats="ats"]) + also + have "\ = ((\ j\ {1..?d}. ?QM j) \ + (\ j\ {1..?d}. \ b\ set B. ?P (I_intterm (a#ats) b + j)))" + by (simp add: dj1_thm dj2th) + also + have "\ = (\ x. ?P x)" + using bst sym[OF cooper_mi_eq[OF unifp]] by simp + finally show ?thesis by simp +qed + + +lemma explode_minf_corr2: + assumes unifp : "isunified p" + and bst: "set (bset p) = set B" + shows "(qinterp ats (QEx p)) = (qinterp ats (decrvars(explode_minf (p,B))))" + (is "?P = (?Qe p)") +proof- + have "?P = (\x. qinterp (x#ats) p)" by simp + also have "\ = (qinterp (a # ats) (explode_minf (p,B)))" + using unifp bst explode_minf_corr by simp + finally have ex: "?P = (qinterp (a # ats) (explode_minf (p,B)))" . + have nv0: "novar0 (explode_minf (p,B))" + by (rule explode_minf_novar0[OF unifp]) + show ?thesis + using qinterp_novar0[OF nv0] ex by simp +qed + +(* An implementation of cooper's method for both plus/minus/infinity *) + +(* unify the formula *) +constdefs unify:: "QF \ (QF \ intterm list)" + "unify p \ + (let q = unitycoeff p; + B = list_set(bset q); + A = list_set (aset q) + in + if (length B \ length A) + then (q,B) + else (mirror q, map lin_neg A))" + +(* unify behaves like unitycoeff *) +lemma unify_ex: + assumes linp: "islinform p" + shows "qinterp ats (QEx p) = qinterp ats (QEx (fst (unify p)))" +proof- + have "length (list_set(bset (unitycoeff p))) \ length (list_set (aset (unitycoeff p))) \ length (list_set(bset (unitycoeff p))) > length (list_set (aset (unitycoeff p)))" by arith + moreover + { + assume "length (list_set(bset (unitycoeff p))) \ length (list_set (aset (unitycoeff p)))" + then have "fst (unify p) = unitycoeff p" using unify_def by (simp add: Let_def) + then have ?thesis using unitycoeff_corr[OF linp] + by simp + } + moreover + { + assume "length (list_set(bset (unitycoeff p))) > length (list_set (aset (unitycoeff p)))" + then have unif: "fst(unify p) = mirror (unitycoeff p)" + using unify_def by (simp add: Let_def) + let ?q ="unitycoeff p" + have unifq: "isunified ?q" by(rule unitycoeff_unified[OF linp]) + have linq: "islinform ?q" by (rule unified_islinform[OF unifq]) + have "qinterp ats (QEx ?q) = qinterp ats (QEx (mirror ?q))" + by (rule mirror_ex2[OF unifq]) + moreover have "qinterp ats (QEx p) = qinterp ats (QEx ?q)" + using unitycoeff_corr linp by simp + ultimately have ?thesis using prems unif by simp + } + ultimately show ?thesis by blast +qed + +(* unify's result is a unified formula *) +lemma unify_unified: + assumes linp: "islinform p" + shows "isunified (fst (unify p))" + using linp unitycoeff_unified mirror_unified unify_def unified_islinform + by (auto simp add: Let_def) + + +(* unify preserves quantifier-freeness*) +lemma unify_qfree: + assumes linp: "islinform p" + shows "isqfree (fst(unify p))" + using linp unify_unified unified_isqfree by simp + +lemma unify_bst: + assumes linp: " islinform p" + and unif: "unify p = (q,B)" + shows "set B = set (bset q)" +proof- + let ?q = "unitycoeff p" + let ?a = "aset ?q" + let ?b = "bset ?q" + let ?la = "list_set ?a" + let ?lb = "list_set ?b" + have " length ?lb \ length ?la \ length ?lb > length ?la" by arith + moreover + { + assume "length ?lb \ length ?la" + then + have "unify p = (?q,?lb)"using unify_def prems by (simp add: Let_def) + then + have ?thesis using prems by (simp add: sym[OF list_set_set]) + } + moreover + { assume "length ?lb > length ?la" + have r: "unify p = (mirror ?q,map lin_neg ?la)"using unify_def prems by (simp add: Let_def) + have lin: "\ x\ set (bset (mirror ?q)). islinintterm x" + using bset_lin mirror_unified unitycoeff_unified[OF linp] by auto + with r prems aset_eq_bset_mirror lin_neg_idemp unitycoeff_unified linp + have "set B = set (map lin_neg (map lin_neg (bset (mirror (unitycoeff p)))))" + by (simp add: sym[OF list_set_set]) + also have "\ = set (map (\x. lin_neg (lin_neg x)) (bset (mirror (unitycoeff p))))" + by auto + also have "\ = set (bset (mirror (unitycoeff p)))" + using lin lin_neg_idemp by (auto simp add: map_idI) + finally + have ?thesis using r prems aset_eq_bset_mirror lin_neg_idemp unitycoeff_unified linp + by (simp add: sym[OF list_set_set])} + ultimately show ?thesis by blast +qed + +lemma explode_minf_unify_novar0: + assumes linp: "islinform p" + shows "novar0 (explode_minf (unify p))" +proof- + have "\ q B. unify p = (q,B)" by simp + then obtain "q" "B" where qB_def: "unify p = (q,B)" by blast + have unifq: "isunified q" using unify_unified[OF linp] qB_def by simp + have bst: "set B = set (bset q)" using unify_bst linp qB_def by simp + from unifq bst explode_minf_novar0 show ?thesis + using qB_def by simp +qed + +lemma explode_minf_unify_corr2: + assumes linp: "islinform p" + shows "qinterp ats (QEx p) = qinterp ats (decrvars(explode_minf(unify p)))" +proof- + have "\ q B. unify p = (q,B)" by simp + then obtain "q" "B" where qB_def: "unify p = (q,B)" by blast + have unifq: "isunified q" using unify_unified[OF linp] qB_def by simp + have bst: "set (bset q) = set B" using unify_bst linp qB_def by simp + from explode_minf_corr2[OF unifq bst] unify_ex[OF linp] show ?thesis + using qB_def by simp +qed +(* An implementation of cooper's method *) +constdefs cooper:: "QF \ QF option" +"cooper p \ lift_un (\q. decrvars(explode_minf (unify q))) (linform (nnf p))" + +(* cooper eliminates quantifiers *) +lemma cooper_qfree: "(\ q q'. \isqfree q ; cooper q = Some q'\ \ isqfree q')" +proof- + fix "q" "q'" + assume qfq: "isqfree q" + and qeq: "cooper q = Some q'" + from qeq have "\p. linform (nnf q) = Some p" + by (cases "linform (nnf q)") (simp_all add: cooper_def) + then obtain "p" where p_def: "linform (nnf q) = Some p" by blast + have linp: "islinform p" using p_def linform_lin nnf_isnnf qfq + by auto + have nnfq: "isnnf (nnf q)" using nnf_isnnf qfq by simp + then have nnfp: "isnnf p" using linform_nnf[OF nnfq] p_def by auto + have qfp: "isqfree p" using linp linform_isqfree by simp + have "cooper q = Some (decrvars(explode_minf (unify p)))" using p_def + by (simp add: cooper_def del: explode_minf.simps) + then have "q' = decrvars (explode_minf (unify p))" using qeq by simp + with linp qfp nnfp unify_unified unify_qfree unified_islinform + show "isqfree q'" + using novar0_qfree explode_minf_unify_novar0 decrvars_qfree + by simp +qed + +(* cooper preserves semantics *) +lemma cooper_corr: "(\ q q' ats. \isqfree q ; cooper q = Some q'\ \ (qinterp ats (QEx q)) = (qinterp ats q'))" (is "\ q q' ats. \ _ ; _ \ \ (?P ats (QEx q) = ?P ats q')") +proof- + fix "q" "q'" "ats" + assume qfq: "isqfree q" + and qeq: "cooper q = Some q'" + from qeq have "\p. linform (nnf q) = Some p" + by (cases "linform (nnf q)") (simp_all add: cooper_def) + then obtain "p" where p_def: "linform (nnf q) = Some p" by blast + have linp: "islinform p" using p_def linform_lin nnf_isnnf qfq by auto + have qfp: "isqfree p" using linp linform_isqfree by simp + have nnfq: "isnnf (nnf q)" using nnf_isnnf qfq by simp + then have nnfp: "isnnf p" using linform_nnf[OF nnfq] p_def by auto + have "\ ats. ?P ats q = ?P ats (nnf q)" using nnf_corr qfq by auto + then have qeqp: "\ ats. ?P ats q = ?P ats p" + using linform_corr p_def nnf_isnnf qfq + by auto + + have "cooper q = Some (decrvars (explode_minf (unify p)))" using p_def + by (simp add: cooper_def del: explode_minf.simps) + then have decr: "q' = decrvars(explode_minf (unify p))" using qeq by simp + have eqq:"?P ats (QEx q) = ?P ats (QEx p)" using qeqp by auto + with decr explode_minf_unify_corr2 unified_islinform unify_unified linp + show "?P ats (QEx q) = ?P ats q'" by simp +qed + +(* A decision procedure for Presburger Arithmetics *) +constdefs pa:: "QF \ QF option" +"pa p \ lift_un psimpl (qelim(cooper, p))" + +lemma psimpl_qfree: "isqfree p \ isqfree (psimpl p)" +apply(induct p rule: isqfree.induct) +apply(auto simp add: Let_def measure_def inv_image_def) +apply (simp_all cong del: QF.weak_case_cong add: Let_def) +apply (case_tac "psimpl p", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl p", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl p", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl p", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) +apply (case_tac "psimpl q", auto) + +apply (case_tac "psimpl p", auto) +apply (case_tac "lift_bin (\x y. lin_add (x, lin_neg y), linearize y, + linearize z)", auto) +apply (case_tac "a",auto) +apply (case_tac "lift_bin (\x y. lin_add (x, lin_neg y), linearize ac, + linearize ad)", auto) +apply (case_tac "a",auto) +apply (case_tac "ae", auto) +apply (case_tac "linearize af", auto) +by (case_tac "a", auto) + +(* pa eliminates quantifiers *) +theorem pa_qfree: "\ p'. pa p = Some p' \ isqfree p'" +proof(simp only: pa_def) +fix "p'" +assume qep: "lift_un psimpl (qelim (cooper, p)) = Some p'" +then have "\ q. qelim (cooper, p) = Some q" + by (cases "qelim(cooper, p)") auto +then obtain "q" where q_def: "qelim (cooper, p) = Some q" by blast +have "\q q'. \isqfree q; cooper q = Some q'\ \ isqfree q'" using cooper_qfree by blast +with q_def +have "isqfree q" using qelim_qfree by blast +then have "isqfree (psimpl q)" using psimpl_qfree + by auto +then show "isqfree p'" + using prems + by simp + +qed + +(* pa preserves semantics *) +theorem pa_corr: + "\ p'. pa p = Some p' \ (qinterp ats p = qinterp ats p')" +proof(simp only: pa_def) + fix "p'" + assume qep: "lift_un psimpl (qelim(cooper, p)) = Some p'" + then have "\ q. qelim (cooper, p) = Some q" + by (cases "qelim(cooper, p)") auto +then obtain "q" where q_def: "qelim (cooper, p) = Some q" by blast + have cp1:"\q q' ats. + \isqfree q; cooper q = Some q'\ \ qinterp ats (QEx q) = qinterp ats q'" + using cooper_corr by blast + moreover have cp2: "\q q'. \isqfree q; cooper q = Some q'\ \ isqfree q'" + using cooper_qfree by blast + ultimately have "qinterp ats p = qinterp ats q" using qelim_corr qep psimpl_corr q_def + by blast + then have "qinterp ats p = qinterp ats (psimpl q)" using psimpl_corr q_def + by auto + then show "qinterp ats p = qinterp ats p'" using prems + by simp +qed + +lemma [code]: "linearize (Mult i j) = + (case linearize i of + None \ None + | Some li \ (case li of + Cst b \ (case linearize j of + None \ None + | (Some lj) \ Some (lin_mul(b,lj))) + | _ \ (case linearize j of + None \ None + | (Some lj) \ (case lj of + Cst b \ Some (lin_mul (b,li)) + | _ \ None))))" +by (simp add: measure_def inv_image_def) + +lemma [code]: "psimpl (And p q) = + (let p'= psimpl p + in (case p' of + F \ F + |T \ psimpl q + | _ \ let q' = psimpl q + in (case q' of + F \ F + | T \ p' + | _ \ (And p' q'))))" + +by (simp add: measure_def inv_image_def) + +lemma [code]: "psimpl (Or p q) = + (let p'= psimpl p + in (case p' of + T \ T + | F \ psimpl q + | _ \ let q' = psimpl q + in (case q' of + T \ T + | F \ p' + | _ \ (Or p' q'))))" + +by (simp add: measure_def inv_image_def) + +lemma [code]: "psimpl (Imp p q) = + (let p'= psimpl p + in (case p' of + F \ T + |T \ psimpl q + | NOT p1 \ let q' = psimpl q + in (case q' of + F \ p1 + | T \ T + | _ \ (Or p1 q')) + | _ \ let q' = psimpl q + in (case q' of + F \ NOT p' + | T \ T + | _ \ (Imp p' q'))))" +by (simp add: measure_def inv_image_def) + +declare zdvd_iff_zmod_eq_0 [code] + +(* +generate_code ("presburger.ML") test = "pa" +*) +use "rcooper.ML" +oracle rpresburger_oracle ("term") = RCooper.rpresburger_oracle +use "rpresbtac.ML" +setup RPresburger.setup + +end