New decision procedure for Presburger arithmetic.
authorberghofe
Tue, 25 Mar 2003 09:47:05 +0100
changeset 13876 68f4ed8311ac
parent 13875 12997e3ddd8d
child 13877 a6b825ee48d9
New decision procedure for Presburger arithmetic.
src/HOL/Integ/Presburger.thy
src/HOL/Integ/cooper_dec.ML
src/HOL/Integ/cooper_proof.ML
src/HOL/Integ/presburger.ML
src/HOL/Integ/qelim.ML
src/HOL/Presburger.thy
src/HOL/Tools/Presburger/cooper_dec.ML
src/HOL/Tools/Presburger/cooper_proof.ML
src/HOL/Tools/Presburger/presburger.ML
src/HOL/Tools/Presburger/qelim.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Integ/Presburger.thy	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,1002 @@
+(*  Title:      HOL/Integ/Presburger.thy
+    ID:         $Id$
+    Author:     Amine Chaieb, Tobias Nipkow and Stefan Berghofer, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+File containing necessary theorems for the proof
+generation for Cooper Algorithm  
+*)
+
+theory Presburger = NatSimprocs
+files
+  ("cooper_dec.ML")
+  ("cooper_proof.ML")
+  ("qelim.ML")
+  ("presburger.ML"):
+
+(* Theorem for unitifying the coeffitients of x in an existential formula*)
+
+theorem unity_coeff_ex: "(\<exists>x::int. P (l * x)) = (\<exists>x. l dvd (1*x+0) \<and> P x)"
+  apply (rule iffI)
+  apply (erule exE)
+  apply (rule_tac x = "l * x" in exI)
+  apply simp
+  apply (erule exE)
+  apply (erule conjE)
+  apply (erule dvdE)
+  apply (rule_tac x = k in exI)
+  apply simp
+  done
+
+lemma uminus_dvd_conv: "(d dvd (t::int)) = (-d dvd t)"
+apply(unfold dvd_def)
+apply(rule iffI)
+apply(clarsimp)
+apply(rename_tac k)
+apply(rule_tac x = "-k" in exI)
+apply simp
+apply(clarsimp)
+apply(rename_tac k)
+apply(rule_tac x = "-k" in exI)
+apply simp
+done
+
+lemma uminus_dvd_conv': "(d dvd (t::int)) = (d dvd -t)"
+apply(unfold dvd_def)
+apply(rule iffI)
+apply(clarsimp)
+apply(rule_tac x = "-k" in exI)
+apply simp
+apply(clarsimp)
+apply(rule_tac x = "-k" in exI)
+apply simp
+done
+
+
+
+(*Theorems for the combination of proofs of the equality of P and P_m for integers x less than some integer z.*)
+
+theorem eq_minf_conjI: "\<exists>z1::int. \<forall>x. x < z1 \<longrightarrow> (A1 x = A2 x) \<Longrightarrow>
+  \<exists>z2::int. \<forall>x. x < z2 \<longrightarrow> (B1 x = B2 x) \<Longrightarrow>
+  \<exists>z::int. \<forall>x. x < z \<longrightarrow> ((A1 x \<and> B1 x) = (A2 x \<and> B2 x))"
+  apply (erule exE)+
+  apply (rule_tac x = "min z1 z2" in exI)
+  apply simp
+  done
+
+
+theorem eq_minf_disjI: "\<exists>z1::int. \<forall>x. x < z1 \<longrightarrow> (A1 x = A2 x) \<Longrightarrow>
+  \<exists>z2::int. \<forall>x. x < z2 \<longrightarrow> (B1 x = B2 x) \<Longrightarrow>
+  \<exists>z::int. \<forall>x. x < z \<longrightarrow> ((A1 x \<or> B1 x) = (A2 x \<or> B2 x))"
+
+  apply (erule exE)+
+  apply (rule_tac x = "min z1 z2" in exI)
+  apply simp
+  done
+
+
+(*Theorems for the combination of proofs of the equality of P and P_m for integers x greather than some integer z.*)
+
+theorem eq_pinf_conjI: "\<exists>z1::int. \<forall>x. z1 < x \<longrightarrow> (A1 x = A2 x) \<Longrightarrow>
+  \<exists>z2::int. \<forall>x. z2 < x \<longrightarrow> (B1 x = B2 x) \<Longrightarrow>
+  \<exists>z::int. \<forall>x. z < x \<longrightarrow> ((A1 x \<and> B1 x) = (A2 x \<and> B2 x))"
+  apply (erule exE)+
+  apply (rule_tac x = "max z1 z2" in exI)
+  apply simp
+  done
+
+
+theorem eq_pinf_disjI: "\<exists>z1::int. \<forall>x. z1 < x \<longrightarrow> (A1 x = A2 x) \<Longrightarrow>
+  \<exists>z2::int. \<forall>x. z2 < x \<longrightarrow> (B1 x = B2 x) \<Longrightarrow>
+  \<exists>z::int. \<forall>x. z < x  \<longrightarrow> ((A1 x \<or> B1 x) = (A2 x \<or> B2 x))"
+  apply (erule exE)+
+  apply (rule_tac x = "max z1 z2" in exI)
+  apply simp
+  done
+(*=============================================================================*)
+(*Theorems for the combination of proofs of the modulo D property for P
+pluusinfinity*)
+(* FIXME : This is THE SAME theorem as for the minusinf version, but with +k.. instead of -k.. In the future replace these both with only one*)
+
+theorem modd_pinf_conjI: "\<forall>(x::int) k. A x = A (x+k*d) \<Longrightarrow>
+  \<forall>(x::int) k. B x = B (x+k*d) \<Longrightarrow>
+  \<forall>(x::int) (k::int). (A x \<and> B x) = (A (x+k*d) \<and> B (x+k*d))"
+  by simp
+
+
+theorem modd_pinf_disjI: "\<forall>(x::int) k. A x = A (x+k*d) \<Longrightarrow>
+  \<forall>(x::int) k. B x = B (x+k*d) \<Longrightarrow>
+  \<forall>(x::int) (k::int). (A x \<or> B x) = (A (x+k*d) \<or> B (x+k*d))"
+  by simp
+
+(*=============================================================================*)
+(*This is one of the cases where the simplifed formula is prooved to habe some property
+(in relation to P_m) but we need to proove the property for the original formula (P_m)*)
+(*FIXME : This is exaclty the same thm as for minusinf.*)
+lemma pinf_simp_eq: "ALL x. P(x) = Q(x) ==> (EX (x::int). P(x)) --> (EX (x::int). F(x))  ==> (EX (x::int). Q(x)) --> (EX (x::int). F(x)) "
+by blast
+
+
+
+(*=============================================================================*)
+(*Theorems for the combination of proofs of the modulo D property for P
+minusinfinity*)
+
+theorem modd_minf_conjI: "\<forall>(x::int) k. A x = A (x-k*d) \<Longrightarrow>
+  \<forall>(x::int) k. B x = B (x-k*d) \<Longrightarrow>
+  \<forall>(x::int) (k::int). (A x \<and> B x) = (A (x-k*d) \<and> B (x-k*d))"
+  by simp
+
+
+theorem modd_minf_disjI: "\<forall>(x::int) k. A x = A (x-k*d) \<Longrightarrow>
+  \<forall>(x::int) k. B x = B (x-k*d) \<Longrightarrow>
+  \<forall>(x::int) (k::int). (A x \<or> B x) = (A (x-k*d) \<or> B (x-k*d))"
+  by simp
+
+(*=============================================================================*)
+(*This is one of the cases where the simplifed formula is prooved to habe some property
+(in relation to P_m) but we need to proove the property for the original formula (P_m)*)
+
+lemma minf_simp_eq: "ALL x. P(x) = Q(x) ==> (EX (x::int). P(x)) --> (EX (x::int). F(x))  ==> (EX (x::int). Q(x)) --> (EX (x::int). F(x)) "
+by blast
+
+(*=============================================================================*)
+
+(*theorem needed for prooving at runtime divide properties using the arithmetic tatic
+(who knows only about modulo = 0)*)
+
+lemma zdvd_iff_zmod_eq_0: "(m dvd n) = (n mod m = (0::int))"
+by(simp add:dvd_def zmod_eq_0_iff)
+
+(*=============================================================================*)
+
+
+
+(*Theorems used for the combination of proof for the backwards direction of cooper's
+theorem. they rely exclusively on Predicate calculus.*)
+
+lemma not_ast_p_disjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P1(x) --> P1(x + d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P2(x) --> P2(x + d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->(P1(x) \<or> P2(x)) --> (P1(x + d) \<or> P2(x + d))) "
+by blast
+
+
+
+lemma not_ast_p_conjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a- j)) --> P1(x) --> P1(x + d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P2(x) --> P2(x + d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->(P1(x) \<and> P2(x)) --> (P1(x + d)
+\<and> P2(x + d))) "
+by blast
+
+lemma not_ast_p_Q_elim: "
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->P(x) --> P(x + d))
+==> ( P = Q )
+==> (ALL x. ~(EX (j::int) : {1..d}. EX (a::int) : A. P(a - j)) -->P(x) --> P(x + d))"
+by blast
+(*=============================================================================*)
+
+
+(*Theorems used for the combination of proof for the backwards direction of cooper's
+theorem. they rely exclusively on Predicate calculus.*)
+
+lemma not_bst_p_disjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P1(x) --> P1(x - d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P2(x) --> P2(x - d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->(P1(x) \<or> P2(x)) --> (P1(x - d)
+\<or> P2(x-d))) "
+by blast
+
+
+
+lemma not_bst_p_conjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P1(x) --> P1(x - d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P2(x) --> P2(x - d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->(P1(x) \<and> P2(x)) --> (P1(x - d)
+\<and> P2(x-d))) "
+by blast
+
+lemma not_bst_p_Q_elim: "
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->P(x) --> P(x - d)) 
+==> ( P = Q )
+==> (ALL x. ~(EX (j::int) : {1..d}. EX (b::int) : B. P(b+j)) -->P(x) --> P(x - d))"
+by blast
+(*=============================================================================*)
+(*The Theorem for the second proof step- about bset. it is trivial too. *)
+lemma bst_thm: " (EX (j::int) : {1..d}. EX (b::int) : B. P (b+j) )--> (EX x::int. P (x)) "
+by blast
+
+(*The Theorem for the second proof step- about aset. it is trivial too. *)
+lemma ast_thm: " (EX (j::int) : {1..d}. EX (a::int) : A. P (a - j) )--> (EX x::int. P (x)) "
+by blast
+
+
+(*=============================================================================*)
+(*This is the first direction of cooper's theorem*)
+lemma cooper_thm: "(R --> (EX x::int. P x))  ==> (Q -->(EX x::int.  P x )) ==> ((R|Q) --> (EX x::int. P x )) "
+by blast
+
+(*=============================================================================*)
+(*The full cooper's theoorem in its equivalence Form- Given the premisses it is trivial
+too, it relies exclusively on prediacte calculus.*)
+lemma cooper_eq_thm: "(R --> (EX x::int. P x))  ==> (Q -->(EX x::int.  P x )) ==> ((~Q)
+--> (EX x::int. P x ) --> R) ==> (EX x::int. P x) = R|Q "
+by blast
+
+(*=============================================================================*)
+(*Some of the atomic theorems generated each time the atom does not depend on x, they
+are trivial.*)
+
+lemma  fm_eq_minf: "EX z::int. ALL x. x < z --> (P = P) "
+by blast
+
+lemma  fm_modd_minf: "ALL (x::int). ALL (k::int). (P = P)"
+by blast
+
+lemma not_bst_p_fm: "ALL (x::int). Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> fm --> fm"
+by blast
+
+
+
+lemma  fm_eq_pinf: "EX z::int. ALL x. z < x --> (P = P) "
+by blast
+
+(* The next 2 thms are the same as the minusinf version*)
+lemma  fm_modd_pinf: "ALL (x::int). ALL (k::int). (P = P)"
+by blast
+
+lemma not_ast_p_fm: "ALL (x::int). Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> fm --> fm"
+by blast
+
+
+(* Theorems to be deleted from simpset when proving simplified formulaes*)
+lemma P_eqtrue: "(P=True) = P"
+  by rules
+
+lemma P_eqfalse: "(P=False) = (~P)"
+  by rules
+
+(*=============================================================================*)
+
+(*Theorems for the generation of the bachwards direction of cooper's theorem*)
+(*These are the 6 interesting atomic cases which have to be proved relying on the
+properties of B-set ant the arithmetic and contradiction proofs*)
+
+lemma not_bst_p_lt: "0 < (d::int) ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ( 0 < -x + a) --> (0 < -(x - d) + a )"
+by arith
+
+lemma not_bst_p_gt: "\<lbrakk> (g::int) \<in> B; g = -a \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> (0 < (x) + a) --> ( 0 < (x - d) + a)"
+apply clarsimp
+apply(rule ccontr)
+apply(drule_tac x = "x+a" in bspec)
+apply(simp add:atLeastAtMost_iff)
+apply(drule_tac x = "-a" in bspec)
+apply assumption
+apply(simp)
+done
+
+lemma not_bst_p_eq: "\<lbrakk> 0 < d; (g::int) \<in> B; g = -a - 1 \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> (0 = x + a) --> (0 = (x - d) + a )"
+apply clarsimp
+apply(subgoal_tac "x = -a")
+ prefer 2 apply arith
+apply(drule_tac x = "1" in bspec)
+apply(simp add:atLeastAtMost_iff)
+apply(drule_tac x = "-a- 1" in bspec)
+apply assumption
+apply(simp)
+done
+
+
+lemma not_bst_p_ne: "\<lbrakk> 0 < d; (g::int) \<in> B; g = -a \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ~(0 = x + a) --> ~(0 = (x - d) + a)"
+apply clarsimp
+apply(subgoal_tac "x = -a+d")
+ prefer 2 apply arith
+apply(drule_tac x = "d" in bspec)
+apply(simp add:atLeastAtMost_iff)
+apply(drule_tac x = "-a" in bspec)
+apply assumption
+apply(simp)
+done
+
+
+lemma not_bst_p_dvd: "(d1::int) dvd d ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> d1 dvd (x + a) --> d1 dvd ((x - d) + a )"
+apply(clarsimp simp add:dvd_def)
+apply(rename_tac m)
+apply(rule_tac x = "m - k" in exI)
+apply(simp add:int_distrib)
+done
+
+lemma not_bst_p_ndvd: "(d1::int) dvd d ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ~(d1 dvd (x + a)) --> ~(d1 dvd ((x - d) + a ))"
+apply(clarsimp simp add:dvd_def)
+apply(rename_tac m)
+apply(erule_tac x = "m + k" in allE)
+apply(simp add:int_distrib)
+done
+
+
+
+(*Theorems for the generation of the bachwards direction of cooper's theorem*)
+(*These are the 6 interesting atomic cases which have to be proved relying on the
+properties of A-set ant the arithmetic and contradiction proofs*)
+
+lemma not_ast_p_gt: "0 < (d::int) ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ( 0 < x + t) --> (0 < (x + d) + t )"
+by arith
+
+
+lemma not_ast_p_lt: "\<lbrakk>0 < d ;(t::int) \<in> A \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> (0 < -x + t) --> ( 0 < -(x + d) + t)"
+  apply clarsimp
+  apply (rule ccontr)
+  apply (drule_tac x = "t-x" in bspec)
+  apply simp
+  apply (drule_tac x = "t" in bspec)
+  apply assumption
+  apply simp
+  done
+
+lemma not_ast_p_eq: "\<lbrakk> 0 < d; (g::int) \<in> A; g = -t + 1 \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> (0 = x + t) --> (0 = (x + d) + t )"
+  apply clarsimp
+  apply (drule_tac x="1" in bspec)
+  apply simp
+  apply (drule_tac x="- t + 1" in bspec)
+  apply assumption
+  apply(subgoal_tac "x = -t")
+  prefer 2 apply arith
+  apply simp
+  done
+
+lemma not_ast_p_ne: "\<lbrakk> 0 < d; (g::int) \<in> A; g = -t \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ~(0 = x + t) --> ~(0 = (x + d) + t)"
+  apply clarsimp
+  apply (subgoal_tac "x = -t-d")
+  prefer 2 apply arith
+  apply (drule_tac x = "d" in bspec)
+  apply simp
+  apply (drule_tac x = "-t" in bspec)
+  apply assumption
+  apply simp
+  done
+
+lemma not_ast_p_dvd: "(d1::int) dvd d ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> d1 dvd (x + t) --> d1 dvd ((x + d) + t )"
+  apply(clarsimp simp add:dvd_def)
+  apply(rename_tac m)
+  apply(rule_tac x = "m + k" in exI)
+  apply(simp add:int_distrib)
+  done
+
+lemma not_ast_p_ndvd: "(d1::int) dvd d ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ~(d1 dvd (x + t)) --> ~(d1 dvd ((x + d) + t ))"
+  apply(clarsimp simp add:dvd_def)
+  apply(rename_tac m)
+  apply(erule_tac x = "m - k" in allE)
+  apply(simp add:int_distrib)
+  done
+
+
+
+(*=============================================================================*)
+(*These are the atomic cases for the proof generation for the modulo D property for P
+plusinfinity*)
+(*They are fully based on arithmetics*)
+
+lemma  dvd_modd_pinf: "((d::int) dvd d1) ==>
+ (ALL (x::int). ALL (k::int). (((d::int) dvd (x + t)) = (d dvd (x+k*d1 + t))))"
+  apply(clarsimp simp add:dvd_def)
+  apply(rule iffI)
+  apply(clarsimp)
+  apply(rename_tac n m)
+  apply(rule_tac x = "m + n*k" in exI)
+  apply(simp add:int_distrib)
+  apply(clarsimp)
+  apply(rename_tac n m)
+  apply(rule_tac x = "m - n*k" in exI)
+  apply(simp add:int_distrib zmult_ac)
+  done
+
+lemma  not_dvd_modd_pinf: "((d::int) dvd d1) ==>
+ (ALL (x::int). ALL k. (~((d::int) dvd (x + t))) = (~(d dvd (x+k*d1 + t))))"
+  apply(clarsimp simp add:dvd_def)
+  apply(rule iffI)
+  apply(clarsimp)
+  apply(rename_tac n m)
+  apply(erule_tac x = "m - n*k" in allE)
+  apply(simp add:int_distrib zmult_ac)
+  apply(clarsimp)
+  apply(rename_tac n m)
+  apply(erule_tac x = "m + n*k" in allE)
+  apply(simp add:int_distrib zmult_ac)
+  done
+
+(*=============================================================================*)
+(*These are the atomic cases for the proof generation for the equivalence of P and P
+plusinfinity for integers x greather than some integer z.*)
+(*They are fully based on arithmetics*)
+
+lemma  eq_eq_pinf: "EX z::int. ALL x. z < x --> (( 0 = x +t ) = False )"
+  apply(rule_tac x = "-t" in exI)
+  apply simp
+  done
+
+lemma  neq_eq_pinf: "EX z::int. ALL x.  z < x --> ((~( 0 = x +t )) = True )"
+  apply(rule_tac x = "-t" in exI)
+  apply simp
+  done
+
+lemma  le_eq_pinf: "EX z::int. ALL x.  z < x --> ( 0 < x +t  = True )"
+  apply(rule_tac x = "-t" in exI)
+  apply simp
+  done
+
+lemma  len_eq_pinf: "EX z::int. ALL x. z < x  --> (0 < -x +t  = False )"
+  apply(rule_tac x = "t" in exI)
+  apply simp
+  done
+
+lemma  dvd_eq_pinf: "EX z::int. ALL x.  z < x --> ((d dvd (x + t)) = (d dvd (x + t))) "
+by simp
+
+lemma  not_dvd_eq_pinf: "EX z::int. ALL x. z < x  --> ((~(d dvd (x + t))) = (~(d dvd (x + t)))) "
+by simp
+
+
+
+
+(*=============================================================================*)
+(*These are the atomic cases for the proof generation for the modulo D property for P
+minusinfinity*)
+(*They are fully based on arithmetics*)
+
+lemma  dvd_modd_minf: "((d::int) dvd d1) ==>
+ (ALL (x::int). ALL (k::int). (((d::int) dvd (x + t)) = (d dvd (x-k*d1 + t))))"
+apply(clarsimp simp add:dvd_def)
+apply(rule iffI)
+apply(clarsimp)
+apply(rename_tac n m)
+apply(rule_tac x = "m - n*k" in exI)
+apply(simp add:int_distrib)
+apply(clarsimp)
+apply(rename_tac n m)
+apply(rule_tac x = "m + n*k" in exI)
+apply(simp add:int_distrib zmult_ac)
+done
+
+
+lemma  not_dvd_modd_minf: "((d::int) dvd d1) ==>
+ (ALL (x::int). ALL k. (~((d::int) dvd (x + t))) = (~(d dvd (x-k*d1 + t))))"
+apply(clarsimp simp add:dvd_def)
+apply(rule iffI)
+apply(clarsimp)
+apply(rename_tac n m)
+apply(erule_tac x = "m + n*k" in allE)
+apply(simp add:int_distrib zmult_ac)
+apply(clarsimp)
+apply(rename_tac n m)
+apply(erule_tac x = "m - n*k" in allE)
+apply(simp add:int_distrib zmult_ac)
+done
+
+
+(*=============================================================================*)
+(*These are the atomic cases for the proof generation for the equivalence of P and P
+minusinfinity for integers x less than some integer z.*)
+(*They are fully based on arithmetics*)
+
+lemma  eq_eq_minf: "EX z::int. ALL x. x < z --> (( 0 = x +t ) = False )"
+apply(rule_tac x = "-t" in exI)
+apply simp
+done
+
+lemma  neq_eq_minf: "EX z::int. ALL x. x < z --> ((~( 0 = x +t )) = True )"
+apply(rule_tac x = "-t" in exI)
+apply simp
+done
+
+lemma  le_eq_minf: "EX z::int. ALL x. x < z --> ( 0 < x +t  = False )"
+apply(rule_tac x = "-t" in exI)
+apply simp
+done
+
+
+lemma  len_eq_minf: "EX z::int. ALL x. x < z --> (0 < -x +t  = True )"
+apply(rule_tac x = "t" in exI)
+apply simp
+done
+
+lemma  dvd_eq_minf: "EX z::int. ALL x. x < z --> ((d dvd (x + t)) = (d dvd (x + t))) "
+by simp
+
+lemma  not_dvd_eq_minf: "EX z::int. ALL x. x < z --> ((~(d dvd (x + t))) = (~(d dvd (x + t)))) "
+by simp
+
+
+(*=============================================================================*)
+(*This Theorem combines whithnesses about P minusinfinity to schow one component of the
+equivalence proof for cooper's theorem*)
+
+(* FIXME: remove once they are part of the distribution *)
+theorem int_ge_induct[consumes 1,case_names base step]:
+  assumes ge: "k \<le> (i::int)" and
+        base: "P(k)" and
+        step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
+  shows "P i"
+proof -
+  { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k <= i \<Longrightarrow> P i"
+    proof (induct n)
+      case 0
+      hence "i = k" by arith
+      thus "P i" using base by simp
+    next
+      case (Suc n)
+      hence "n = nat((i - 1) - k)" by arith
+      moreover
+      have ki1: "k \<le> i - 1" using Suc.prems by arith
+      ultimately
+      have "P(i - 1)" by(rule Suc.hyps)
+      from step[OF ki1 this] show ?case by simp
+    qed
+  }
+  from this ge show ?thesis by fast
+qed
+
+theorem int_gr_induct[consumes 1,case_names base step]:
+  assumes gr: "k < (i::int)" and
+        base: "P(k+1)" and
+        step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
+  shows "P i"
+apply(rule int_ge_induct[of "k + 1"])
+  using gr apply arith
+ apply(rule base)
+apply(rule step)
+ apply simp+
+done
+
+lemma decr_lemma: "0 < (d::int) \<Longrightarrow> x - (abs(x-z)+1) * d < z"
+apply(induct rule: int_gr_induct)
+ apply simp
+ apply arith
+apply (simp add:int_distrib)
+apply arith
+done
+
+lemma incr_lemma: "0 < (d::int) \<Longrightarrow> z < x + (abs(x-z)+1) * d"
+apply(induct rule: int_gr_induct)
+ apply simp
+ apply arith
+apply (simp add:int_distrib)
+apply arith
+done
+
+lemma  minusinfinity:
+  assumes "0 < d" and
+    P1eqP1: "ALL x k. P1 x = P1(x - k*d)" and
+    ePeqP1: "EX z::int. ALL x. x < z \<longrightarrow> (P x = P1 x)"
+  shows "(EX x. P1 x) \<longrightarrow> (EX x. P x)"
+proof
+  assume eP1: "EX x. P1 x"
+  then obtain x where P1: "P1 x" ..
+  from ePeqP1 obtain z where P1eqP: "ALL x. x < z \<longrightarrow> (P x = P1 x)" ..
+  let ?w = "x - (abs(x-z)+1) * d"
+  show "EX x. P x"
+  proof
+    have w: "?w < z" by(rule decr_lemma)
+    have "P1 x = P1 ?w" using P1eqP1 by blast
+    also have "\<dots> = P(?w)" using w P1eqP by blast
+    finally show "P ?w" using P1 by blast
+  qed
+qed
+
+(*=============================================================================*)
+(*This Theorem combines whithnesses about P minusinfinity to schow one component of the
+equivalence proof for cooper's theorem*)
+
+lemma plusinfinity:
+  assumes "0 < d" and
+    P1eqP1: "ALL (x::int) (k::int). P1 x = P1 (x + k * d)" and
+    ePeqP1: "EX z::int. ALL x. z < x  --> (P x = P1 x)"
+  shows "(EX x::int. P1 x) --> (EX x::int. P x)"
+proof
+  assume eP1: "EX x. P1 x"
+  then obtain x where P1: "P1 x" ..
+  from ePeqP1 obtain z where P1eqP: "ALL x. z < x \<longrightarrow> (P x = P1 x)" ..
+  let ?w = "x + (abs(x-z)+1) * d"
+  show "EX x. P x"
+  proof
+    have w: "z < ?w" by(rule incr_lemma)
+    have "P1 x = P1 ?w" using P1eqP1 by blast
+    also have "\<dots> = P(?w)" using w P1eqP by blast
+    finally show "P ?w" using P1 by blast
+  qed
+qed
+ 
+
+
+(*=============================================================================*)
+(*Theorem for periodic function on discrete sets*)
+
+lemma minf_vee:
+  assumes dpos: "(0::int) < d" and modd: "ALL x k. P x = P(x - k*d)"
+  shows "(EX x. P x) = (EX j : {1..d}. P j)"
+  (is "?LHS = ?RHS")
+proof
+  assume ?LHS
+  then obtain x where P: "P x" ..
+  have "x mod d = x - (x div d)*d"
+    by(simp add:zmod_zdiv_equality zmult_ac eq_zdiff_eq)
+  hence Pmod: "P x = P(x mod d)" using modd by simp
+  show ?RHS
+  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 ?RHS ..
+  next
+    assume not0: "x mod d \<noteq> 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 \<le> x mod d" by(rule pos_mod_sign)
+      moreover have "x mod d < d" by(rule pos_mod_bound)
+      ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff)
+    qed
+    ultimately show ?RHS ..
+  qed
+next
+  assume ?RHS thus ?LHS by blast
+qed
+
+(*=============================================================================*)
+(*Theorem for periodic function on discrete sets*)
+lemma pinf_vee:
+  assumes dpos: "0 < (d::int)" and modd: "ALL (x::int) (k::int). P x = P (x+k*d)"
+  shows "(EX x::int. P x) = (EX (j::int) : {1..d} . P j)"
+  (is "?LHS = ?RHS")
+proof
+  assume ?LHS
+  then obtain x where P: "P x" ..
+  have "x mod d = x + (-(x div d))*d"
+    by(simp add:zmod_zdiv_equality zmult_ac eq_zdiff_eq)
+  hence Pmod: "P x = P(x mod d)" using modd by (simp only:)
+  show ?RHS
+  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 ?RHS ..
+  next
+    assume not0: "x mod d \<noteq> 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 \<le> x mod d" by(rule pos_mod_sign)
+      moreover have "x mod d < d" by(rule pos_mod_bound)
+      ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff)
+    qed
+    ultimately show ?RHS ..
+  qed
+next
+  assume ?RHS thus ?LHS by blast
+qed
+
+lemma decr_mult_lemma:
+  assumes dpos: "(0::int) < d" and
+          minus: "ALL x::int. P x \<longrightarrow> P(x - d)" and
+          knneg: "0 <= k"
+  shows "ALL x. P x \<longrightarrow> P(x - k*d)"
+using knneg
+proof (induct rule:int_ge_induct)
+  case base thus ?case by simp
+next
+  case (step i)
+  show ?case
+  proof
+    fix x
+    have "P x \<longrightarrow> P (x - i * d)" using step.hyps by blast
+    also have "\<dots> \<longrightarrow> P(x - (i + 1) * d)"
+      using minus[THEN spec, of "x - i * d"]
+      by (simp add:int_distrib zdiff_zdiff_eq[symmetric])
+    ultimately show "P x \<longrightarrow> P(x - (i + 1) * d)" by blast
+  qed
+qed
+
+lemma incr_mult_lemma:
+  assumes dpos: "(0::int) < d" and
+          plus: "ALL x::int. P x \<longrightarrow> P(x + d)" and
+          knneg: "0 <= k"
+  shows "ALL x. P x \<longrightarrow> P(x + k*d)"
+using knneg
+proof (induct rule:int_ge_induct)
+  case base thus ?case by simp
+next
+  case (step i)
+  show ?case
+  proof
+    fix x
+    have "P x \<longrightarrow> P (x + i * d)" using step.hyps by blast
+    also have "\<dots> \<longrightarrow> P(x + (i + 1) * d)"
+      using plus[THEN spec, of "x + i * d"]
+      by (simp add:int_distrib zadd_ac)
+    ultimately show "P x \<longrightarrow> P(x + (i + 1) * d)" by blast
+  qed
+qed
+
+lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x))
+==> (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)) --> (EX (x::int). P x) 
+==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) 
+==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D))))
+==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))"
+apply(rule iffI)
+prefer 2
+apply(drule minusinfinity)
+apply assumption+
+apply(fastsimp)
+apply clarsimp
+apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)")
+apply(frule_tac x = x and z=z in decr_lemma)
+apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)")
+prefer 2
+apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
+prefer 2 apply arith
+ apply fastsimp
+apply(drule (1) minf_vee)
+apply blast
+apply(blast dest:decr_mult_lemma)
+done
+
+(* Cooper Thm `, plus infinity version*)
+lemma cppi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. z < x --> (P x = P1 x))
+==> (EX (j::int) : {1..D}. EX (a::int) : A. P (a - j)) --> (EX (x::int). P x) 
+==> ALL x.~(EX (j::int) : {1..D}. EX (a::int) : A. P(a - j)) --> P (x) --> P (x + D) 
+==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x+k*D))))
+==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (a::int) : A. P (a - j)))"
+  apply(rule iffI)
+  prefer 2
+  apply(drule plusinfinity)
+  apply assumption+
+  apply(fastsimp)
+  apply clarsimp
+  apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x + k*D)")
+  apply(frule_tac x = x and z=z in incr_lemma)
+  apply(subgoal_tac "P1(x + (\<bar>x - z\<bar> + 1) * D)")
+  prefer 2
+  apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
+  prefer 2 apply arith
+  apply fastsimp
+  apply(drule (1) pinf_vee)
+  apply blast
+  apply(blast dest:incr_mult_lemma)
+  done
+
+
+(*=============================================================================*)
+
+(*Theorems for the quantifier elminination Functions.*)
+
+lemma qe_ex_conj: "(EX (x::int). A x) = R
+		==> (EX (x::int). P x) = (Q & (EX x::int. A x))
+		==> (EX (x::int). P x) = (Q & R)"
+by blast
+
+lemma qe_ex_nconj: "(EX (x::int). P x) = (True & Q)
+		==> (EX (x::int). P x) = Q"
+by blast
+
+lemma qe_conjI: "P1 = P2 ==> Q1 = Q2 ==> (P1 & Q1) = (P2 & Q2)"
+by blast
+
+lemma qe_disjI: "P1 = P2 ==> Q1 = Q2 ==> (P1 | Q1) = (P2 | Q2)"
+by blast
+
+lemma qe_impI: "P1 = P2 ==> Q1 = Q2 ==> (P1 --> Q1) = (P2 --> Q2)"
+by blast
+
+lemma qe_eqI: "P1 = P2 ==> Q1 = Q2 ==> (P1 = Q1) = (P2 = Q2)"
+by blast
+
+lemma qe_Not: "P = Q ==> (~P) = (~Q)"
+by blast
+
+lemma qe_ALL: "(EX x. ~P x) = R ==> (ALL x. P x) = (~R)"
+by blast
+
+(* Theorems for proving NNF *)
+
+lemma nnf_im: "((~P) = P1) ==> (Q=Q1) ==> ((P --> Q) = (P1 | Q1))"
+by blast
+
+lemma nnf_eq: "((P & Q) = (P1 & Q1)) ==> (((~P) & (~Q)) = (P2 & Q2)) ==> ((P = Q) = ((P1 & Q1)|(P2 & Q2)))"
+by blast
+
+lemma nnf_nn: "(P = Q) ==> ((~~P) = Q)"
+  by blast
+lemma nnf_ncj: "((~P) = P1) ==> ((~Q) = Q1) ==> ((~(P & Q)) = (P1 | Q1))"
+by blast
+
+lemma nnf_ndj: "((~P) = P1) ==> ((~Q) = Q1) ==> ((~(P | Q)) = (P1 & Q1))"
+by blast
+lemma nnf_nim: "(P = P1) ==> ((~Q) = Q1) ==> ((~(P --> Q)) = (P1 & Q1))"
+by blast
+lemma nnf_neq: "((P & (~Q)) = (P1 & Q1)) ==> (((~P) & Q) = (P2 & Q2)) ==> ((~(P = Q)) = ((P1 & Q1)|(P2 & Q2)))"
+by blast
+lemma nnf_sdj: "((A & (~B)) = (A1 & B1)) ==> ((C & (~D)) = (C1 & D1)) ==> (A = (~C)) ==> ((~((A & B) | (C & D))) = ((A1 & B1) | (C1 & D1)))"
+by blast
+
+
+lemma qe_exI2: "A = B ==> (EX (x::int). A(x)) = (EX (x::int). B(x))"
+  by simp
+
+lemma qe_exI: "(!!x::int. A x = B x) ==> (EX (x::int). A(x)) = (EX (x::int). B(x))"
+  by rules
+
+lemma qe_ALLI: "(!!x::int. A x = B x) ==> (ALL (x::int). A(x)) = (ALL (x::int). B(x))"
+  by rules
+
+lemma cp_expand: "(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (b::int) : B. (P1 (j) | P(b+j)))
+==>(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (b::int) : B. (P1 (j) | P(b+j))) "
+by blast
+
+lemma cppi_expand: "(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (a::int) : A. (P1 (j) | P(a - j)))
+==>(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (a::int) : A. (P1 (j) | P(a - j))) "
+by blast
+
+
+lemma simp_from_to: "{i..j::int} = (if j < i then {} else insert i {i+1..j})"
+apply(simp add:atLeastAtMost_def atLeast_def atMost_def)
+apply(fastsimp)
+done
+
+(* Theorems required for the adjustcoeffitienteq*)
+
+lemma ac_dvd_eq: assumes not0: "0 ~= (k::int)"
+shows "((m::int) dvd (c*n+t)) = (k*m dvd ((k*c)*n+(k*t)))" (is "?P = ?Q")
+proof
+  assume ?P
+  thus ?Q
+    apply(simp add:dvd_def)
+    apply clarify
+    apply(rename_tac d)
+    apply(drule_tac f = "op * k" in arg_cong)
+    apply(simp only:int_distrib)
+    apply(rule_tac x = "d" in exI)
+    apply(simp only:zmult_ac)
+    done
+next
+  assume ?Q
+  then obtain d where "k * c * n + k * t = (k*m)*d" by(fastsimp simp:dvd_def)
+  hence "(c * n + t) * k = (m*d) * k" by(simp add:int_distrib zmult_ac)
+  hence "((c * n + t) * k) div k = ((m*d) * k) div k" by(rule arg_cong[of _ _ "%t. t div k"])
+  hence "c*n+t = m*d" by(simp add: zdiv_zmult_self1[OF not0[symmetric]])
+  thus ?P by(simp add:dvd_def)
+qed
+
+lemma ac_lt_eq: assumes gr0: "0 < (k::int)"
+shows "((m::int) < (c*n+t)) = (k*m <((k*c)*n+(k*t)))" (is "?P = ?Q")
+proof
+  assume P: ?P
+  show ?Q using zmult_zless_mono2[OF P gr0] by(simp add: int_distrib zmult_ac)
+next
+  assume ?Q
+  hence "0 < k*(c*n + t - m)" by(simp add: int_distrib zmult_ac)
+  with gr0 have "0 < (c*n + t - m)" by(simp add:int_0_less_mult_iff)
+  thus ?P by(simp)
+qed
+
+lemma ac_eq_eq : assumes not0: "0 ~= (k::int)" shows "((m::int) = (c*n+t)) = (k*m =((k*c)*n+(k*t)) )" (is "?P = ?Q")
+proof
+  assume ?P
+  thus ?Q
+    apply(drule_tac f = "op * k" in arg_cong)
+    apply(simp only:int_distrib)
+    done
+next
+  assume ?Q
+  hence "m * k = (c*n + t) * k" by(simp add:int_distrib zmult_ac)
+  hence "((m) * k) div k = ((c*n + t) * k) div k" by(rule arg_cong[of _ _ "%t. t div k"])
+  thus ?P by(simp add: zdiv_zmult_self1[OF not0[symmetric]])
+qed
+
+lemma ac_pi_eq: assumes gr0: "0 < (k::int)" shows "(~((0::int) < (c*n + t))) = (0 < ((-k)*c)*n + ((-k)*t + k))"
+proof -
+  have "(~ (0::int) < (c*n + t)) = (0<1-(c*n + t))" by arith
+  also have  "(1-(c*n + t)) = (-1*c)*n + (-t+1)" by(simp add: int_distrib zmult_ac)
+  also have "0<(-1*c)*n + (-t+1) = (0 < (k*(-1*c)*n) + (k*(-t+1)))" by(rule ac_lt_eq[of _ 0,OF gr0,simplified])
+  also have "(k*(-1*c)*n) + (k*(-t+1)) = ((-k)*c)*n + ((-k)*t + k)" by(simp add: int_distrib zmult_ac)
+  finally show ?thesis .
+qed
+
+lemma binminus_uminus_conv: "(a::int) - b = a + (-b)"
+by arith
+
+lemma  linearize_dvd: "(t::int) = t1 ==> (d dvd t) = (d dvd t1)"
+by simp
+
+lemma lf_lt: "(l::int) = ll ==> (r::int) = lr ==> (l < r) =(ll < lr)"
+by simp
+
+lemma lf_eq: "(l::int) = ll ==> (r::int) = lr ==> (l = r) =(ll = lr)"
+by simp
+
+lemma lf_dvd: "(l::int) = ll ==> (r::int) = lr ==> (l dvd r) =(ll dvd lr)"
+by simp
+
+(* Theorems for transforming predicates on nat to predicates on int*)
+
+theorem all_nat: "(\<forall>x::nat. P x) = (\<forall>x::int. 0 <= x \<longrightarrow> P (nat x))"
+  by (simp split add: split_nat)
+
+theorem ex_nat: "(\<exists>x::nat. P x) = (\<exists>x::int. 0 <= x \<and> P (nat x))"
+  apply (simp split add: split_nat)
+  apply (rule iffI)
+  apply (erule exE)
+  apply (rule_tac x = "int x" in exI)
+  apply simp
+  apply (erule exE)
+  apply (rule_tac x = "nat x" in exI)
+  apply (erule conjE)
+  apply (erule_tac x = "nat x" in allE)
+  apply simp
+  done
+
+theorem zdiff_int_split: "P (int (x - y)) =
+  ((y \<le> x \<longrightarrow> P (int x - int y)) \<and> (x < y \<longrightarrow> P 0))"
+  apply (case_tac "y \<le> x")
+  apply (simp_all add: zdiff_int)
+  done
+
+theorem zdvd_int: "(x dvd y) = (int x dvd int y)"
+  apply (simp only: dvd_def ex_nat int_int_eq [symmetric] zmult_int [symmetric]
+    nat_0_le cong add: conj_cong)
+  apply (rule iffI)
+  apply rules
+  apply (erule exE)
+  apply (case_tac "x=0")
+  apply (rule_tac x=0 in exI)
+  apply simp
+  apply (case_tac "0 \<le> k")
+  apply rules
+  apply (simp add: linorder_not_le)
+  apply (drule zmult_zless_mono2_neg [OF iffD2 [OF zero_less_int_conv]])
+  apply assumption
+  apply (simp add: zmult_ac)
+  done
+
+theorem number_of1: "(0::int) <= number_of n \<Longrightarrow> (0::int) <= number_of (n BIT b)"
+  by simp
+
+theorem number_of2: "(0::int) <= number_of bin.Pls" by simp
+
+theorem Suc_plus1: "Suc n = n + 1" by simp
+
+(* specific instances of congruence rules, to prevent simplifier from looping *)
+
+theorem imp_le_cong: "(0 <= x \<Longrightarrow> P = P') \<Longrightarrow> (0 <= (x::nat) \<longrightarrow> P) = (0 <= x \<longrightarrow> P')"
+  by simp
+
+theorem conj_le_cong: "(0 <= x \<Longrightarrow> P = P') \<Longrightarrow> (0 <= (x::nat) \<and> P) = (0 <= x \<and> P')"
+  by simp
+
+use "cooper_dec.ML"
+use "cooper_proof.ML"
+use "qelim.ML"
+use "presburger.ML"
+
+setup "Presburger.setup"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Integ/cooper_dec.ML	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,773 @@
+(*  Title:      HOL/Integ/cooper_dec.ML
+    ID:         $Id$
+    Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+File containing the implementation of Cooper Algorithm
+decision procedure (intensively inspired from J.Harrison)
+*)
+ 
+signature COOPER_DEC = 
+sig
+  exception COOPER
+  val is_arith_rel : term -> bool
+  val mk_numeral : int -> term
+  val dest_numeral : term -> int
+  val zero : term
+  val one : term
+  val linear_cmul : int -> term -> term
+  val linear_add : string list -> term -> term -> term 
+  val linear_sub : string list -> term -> term -> term 
+  val linear_neg : term -> term
+  val lint : string list -> term -> term
+  val linform : string list -> term -> term
+  val formlcm : term -> term -> int
+  val adjustcoeff : term -> int -> term -> term
+  val unitycoeff : term -> term -> term
+  val divlcm : term -> term -> int
+  val bset : term -> term -> term list
+  val aset : term -> term -> term list
+  val linrep : string list -> term -> term -> term -> term
+  val list_disj : term list -> term
+  val simpl : term -> term
+  val fv : term -> string list
+  val negate : term -> term
+  val operations : (string * (int * int -> bool)) list
+end;
+
+structure  CooperDec : COOPER_DEC =
+struct
+
+(* ========================================================================= *) 
+(* Cooper's algorithm for Presburger arithmetic.                             *) 
+(* ========================================================================= *) 
+exception COOPER;
+
+(* ------------------------------------------------------------------------- *) 
+(* Lift operations up to numerals.                                           *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+(*Assumption : The construction of atomar formulas in linearl arithmetic is based on 
+relation operations of Type : [int,int]---> bool *) 
+ 
+(* ------------------------------------------------------------------------- *) 
+ 
+ 
+(*Function is_arith_rel returns true if and only if the term is an atomar presburger 
+formula *) 
+fun is_arith_rel tm = case tm of 
+	 Const(p,Type ("fun",[Type ("Numeral.bin", []),Type ("fun",[Type ("Numeral.bin", 
+	 []),Type ("bool",[])] )])) $ _ $_ => true 
+	|Const(p,Type ("fun",[Type ("IntDef.int", []),Type ("fun",[Type ("IntDef.int", 
+	 []),Type ("bool",[])] )])) $ _ $_ => true 
+	|_ => false; 
+ 
+(*Function is_arith_rel returns true if and only if the term is an operation of the 
+form [int,int]---> int*) 
+ 
+(*Transform a natural number to a term*) 
+ 
+fun mk_numeral 0 = Const("0",HOLogic.intT)
+   |mk_numeral 1 = Const("1",HOLogic.intT)
+   |mk_numeral n = (HOLogic.number_of_const HOLogic.intT) $ (HOLogic.mk_bin n); 
+ 
+(*Transform an Term to an natural number*)	  
+	  
+fun dest_numeral (Const("0",Type ("IntDef.int", []))) = 0
+   |dest_numeral (Const("1",Type ("IntDef.int", []))) = 1
+   |dest_numeral (Const ("Numeral.number_of",_) $ n)= HOLogic.dest_binum n; 
+(*Some terms often used for pattern matching*) 
+ 
+val zero = mk_numeral 0; 
+val one = mk_numeral 1; 
+ 
+(*Tests if a Term is representing a number*) 
+ 
+fun is_numeral t = (t = zero) orelse (t = one) orelse (can dest_numeral t); 
+ 
+(*maps a unary natural function on a term containing an natural number*) 
+ 
+fun numeral1 f n = mk_numeral (f(dest_numeral n)); 
+ 
+(*maps a binary natural function on 2 term containing  natural numbers*) 
+ 
+fun numeral2 f m n = mk_numeral(f(dest_numeral m) (dest_numeral n)); 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Operations on canonical linear terms c1 * x1 + ... + cn * xn + k          *) 
+(*                                                                           *) 
+(* Note that we're quite strict: the ci must be present even if 1            *) 
+(* (but if 0 we expect the monomial to be omitted) and k must be there       *) 
+(* even if it's zero. Thus, it's a constant iff not an addition term.        *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+ 
+fun linear_cmul n tm =  if n = 0 then zero else let fun times n k = n*k in  
+  ( case tm of  
+     (Const("op +",T)  $  (Const ("op *",T1 ) $c1 $  x1) $ rest) => 
+       Const("op +",T) $ ((Const("op *",T1) $ (numeral1 (times n) c1) $ x1)) $ (linear_cmul n rest) 
+    |_ =>  numeral1 (times n) tm) 
+    end ; 
+ 
+ 
+ 
+ 
+(* Whether the first of two items comes earlier in the list  *) 
+fun earlier [] x y = false 
+	|earlier (h::t) x y =if h = y then false 
+              else if h = x then true 
+              	else earlier t x y ; 
+ 
+fun earlierv vars (Bound i) (Bound j) = i < j 
+   |earlierv vars (Bound _) _ = true 
+   |earlierv vars _ (Bound _)  = false 
+   |earlierv vars (Free (x,_)) (Free (y,_)) = earlier vars x y; 
+ 
+ 
+fun linear_add vars tm1 tm2 = 
+  let fun addwith x y = x + y in
+ (case (tm1,tm2) of 
+	((Const ("op +",T1) $ ( Const("op *",T2) $ c1 $  x1) $ rest1),(Const 
+	("op +",T3)$( Const("op *",T4) $ c2 $  x2) $ rest2)) => 
+         if x1 = x2 then 
+              let val c = (numeral2 (addwith) c1 c2) 
+	      in 
+              if c = zero then (linear_add vars rest1  rest2)  
+	      else (Const("op +",T1) $ (Const("op *",T2) $ c $ x1) $ (linear_add vars  rest1 rest2)) 
+              end 
+	   else 
+		if earlierv vars x1 x2 then (Const("op +",T1) $  
+		(Const("op *",T2)$ c1 $ x1) $ (linear_add vars rest1 tm2)) 
+    	       else (Const("op +",T1) $ (Const("op *",T2) $ c2 $ x2) $ (linear_add vars tm1 rest2)) 
+   	|((Const("op +",T1) $ (Const("op *",T2) $ c1 $ x1) $ rest1) ,_) => 
+    	  (Const("op +",T1)$ (Const("op *",T2) $ c1 $ x1) $ (linear_add vars 
+	  rest1 tm2)) 
+   	|(_, (Const("op +",T1) $(Const("op *",T2) $ c2 $ x2) $ rest2)) => 
+      	  (Const("op +",T1) $ (Const("op *",T2) $ c2 $ x2) $ (linear_add vars tm1 
+	  rest2)) 
+   	| (_,_) => numeral2 (addwith) tm1 tm2) 
+	 
+	end; 
+ 
+(*To obtain the unary - applyed on a formula*) 
+ 
+fun linear_neg tm = linear_cmul (0 - 1) tm; 
+ 
+(*Substraction of two terms *) 
+ 
+fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2); 
+ 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Linearize a term.                                                         *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+(* linearises a term from the point of view of Variable Free (x,T). 
+After this fuction the all expressions containig ths variable will have the form  
+ c*Free(x,T) + t where c is a constant ant t is a Term which is not containing 
+ Free(x,T)*) 
+  
+fun lint vars tm = if is_numeral tm then tm else case tm of 
+   (Free (x,T)) =>  (HOLogic.mk_binop "op +" ((HOLogic.mk_binop "op *" ((mk_numeral 1),Free (x,T))), zero)) 
+  |(Bound i) =>  (Const("op +",HOLogic.intT -->HOLogic.intT -->HOLogic.intT) $ 
+  (Const("op *",HOLogic.intT -->HOLogic.intT -->HOLogic.intT) $ (mk_numeral 1) $ (Bound i)) $ zero) 
+  |(Const("uminus",_) $ t ) => (linear_neg (lint vars t)) 
+  |(Const("op +",_) $ s $ t) => (linear_add vars (lint vars s) (lint vars t)) 
+  |(Const("op -",_) $ s $ t) => (linear_sub vars (lint vars s) (lint vars t)) 
+  |(Const ("op *",_) $ s $ t) => 
+        let val s' = lint vars s  
+            val t' = lint vars t  
+        in 
+        if is_numeral s' then (linear_cmul (dest_numeral s') t') 
+        else if is_numeral t' then (linear_cmul (dest_numeral t') s') 
+ 
+         else (warning "lint: apparent nonlinearity"; raise COOPER)
+         end 
+  |_ =>   error "lint: unknown term"; 
+   
+ 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Linearize the atoms in a formula, and eliminate non-strict inequalities.  *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun mkatom vars p t = Const(p,HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ zero $ (lint vars t); 
+ 
+fun linform vars (Const ("Divides.op dvd",_) $ c $ t) =  
+      let val c' = (mk_numeral(abs(dest_numeral c)))  
+      in (HOLogic.mk_binrel "Divides.op dvd" (c,lint vars t)) 
+      end 
+  |linform vars  (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ s $ t ) = (mkatom vars "op =" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ t $ s) ) 
+  |linform vars  (Const("op <",_)$ s $t ) = (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ t $ s))
+  |linform vars  (Const("op >",_) $ s $ t ) = (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ s $ t)) 
+  |linform vars  (Const("op <=",_)$ s $ t ) = 
+        (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ (Const("op +",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $t $(mk_numeral 1)) $ s)) 
+  |linform vars  (Const("op >=",_)$ s $ t ) = 
+        (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> 
+	HOLogic.intT) $ (Const("op +",HOLogic.intT --> HOLogic.intT --> 
+	HOLogic.intT) $s $(mk_numeral 1)) $ t)) 
+ 
+   |linform vars  fm =  fm; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Post-NNF transformation eliminating negated inequalities.                 *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun posineq fm = case fm of  
+ (Const ("Not",_)$(Const("op <",_)$ c $ t)) =>
+   (HOLogic.mk_binrel "op <"  (zero , (linear_sub [] (mk_numeral 1) (linear_add [] c t ) ))) 
+  | ( Const ("op &",_) $ p $ q)  => HOLogic.mk_conj (posineq p,posineq q)
+  | ( Const ("op |",_) $ p $ q ) => HOLogic.mk_disj (posineq p,posineq q)
+  | _ => fm; 
+  
+
+(* ------------------------------------------------------------------------- *) 
+(* Find the LCM of the coefficients of x.                                    *) 
+(* ------------------------------------------------------------------------- *) 
+(*gcd calculates gcd (a,b) and helps lcm_num calculating lcm (a,b)*) 
+ 
+fun gcd a b = if a=0 then b else gcd (b mod a) a ; 
+fun lcm_num a b = (abs a*b) div (gcd (abs a) (abs b)); 
+ 
+fun formlcm x fm = case fm of 
+    (Const (p,_)$ _ $(Const ("op +", _)$(Const ("op *",_)$ c $ y ) $z ) ) =>  if 
+    (is_arith_rel fm) andalso (x = y) then abs(dest_numeral c) else 1 
+  | ( Const ("Not", _) $p) => formlcm x p 
+  | ( Const ("op &",_) $ p $ q) => lcm_num (formlcm x p) (formlcm x q) 
+  | ( Const ("op |",_) $ p $ q )=> lcm_num (formlcm x p) (formlcm x q) 
+  |  _ => 1; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Adjust all coefficients of x in formula; fold in reduction to +/- 1.      *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun adjustcoeff x l fm = 
+     case fm of  
+      (Const(p,_) $d $( Const ("op +", _)$(Const ("op *",_) $ 
+      c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then  
+        let val m = l div (dest_numeral c) 
+            val n = (if p = "op <" then abs(m) else m) 
+            val xtm = HOLogic.mk_binop "op *" ((mk_numeral (m div n)), x) 
+	in
+        (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
+	end 
+	else fm 
+  |( Const ("Not", _) $ p) => HOLogic.Not $ (adjustcoeff x l p) 
+  |( Const ("op &",_) $ p $ q) => HOLogic.conj$(adjustcoeff x l p) $(adjustcoeff x l q) 
+  |( Const ("op |",_) $ p $ q) => HOLogic.disj $(adjustcoeff x l p)$ (adjustcoeff x l q) 
+  |_ => fm; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Hence make coefficient of x one in existential formula.                   *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun unitycoeff x fm = 
+  let val l = formlcm x fm 
+      val fm' = adjustcoeff x l fm in
+     if l = 1 then fm' else 
+     let val xp = (HOLogic.mk_binop "op +"  
+     		((HOLogic.mk_binop "op *" ((mk_numeral 1), x )), zero)) in 
+      HOLogic.conj $(HOLogic.mk_binrel "Divides.op dvd" ((mk_numeral l) , xp )) $ (adjustcoeff x l fm) 
+      end 
+  end; 
+ 
+(* adjustcoeffeq l fm adjusts the coeffitients c_i of x  overall in fm to l*)
+(* Here l must be a multiple of all c_i otherwise the obtained formula is not equivalent*)
+(*
+fun adjustcoeffeq x l fm = 
+    case fm of  
+      (Const(p,_) $d $( Const ("op +", _)$(Const ("op *",_) $ 
+      c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then  
+        let val m = l div (dest_numeral c) 
+            val n = (if p = "op <" then abs(m) else m)  
+            val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x))
+            in (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
+	    end 
+	else fm 
+  |( Const ("Not", _) $ p) => HOLogic.Not $ (adjustcoeffeq x l p) 
+  |( Const ("op &",_) $ p $ q) => HOLogic.conj$(adjustcoeffeq x l p) $(adjustcoeffeq x l q) 
+  |( Const ("op |",_) $ p $ q) => HOLogic.disj $(adjustcoeffeq x l p)$ (adjustcoeffeq x l q) 
+  |_ => fm;
+ 
+
+*)
+
+(* ------------------------------------------------------------------------- *) 
+(* The "minus infinity" version.                                             *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun minusinf x fm = case fm of  
+    (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ (c1 ) $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) => 
+  	 if (is_arith_rel fm) andalso (x=y) andalso (c2 = one) andalso (c1 =zero) then HOLogic.false_const  
+	 				 else fm 
+ 
+  |(Const("op <",_) $ c $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z 
+  )) => 
+        if (x =y) andalso (pm1 = one) andalso (c = zero) then HOLogic.false_const else HOLogic.true_const 
+	 
+  |(Const ("Not", _) $ p) => HOLogic.Not $ (minusinf x p) 
+  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (minusinf x p) $ (minusinf x q) 
+  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (minusinf x p) $ (minusinf x q) 
+  |_ => fm; 
+
+(* ------------------------------------------------------------------------- *)
+(* The "Plus infinity" version.                                             *)
+(* ------------------------------------------------------------------------- *)
+
+fun plusinf x fm = case fm of
+    (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ (c1 ) $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	 if (is_arith_rel fm) andalso (x=y) andalso (c2 = one) andalso (c1 =zero) then HOLogic.false_const
+	 				 else fm
+
+  |(Const("op <",_) $ c $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z
+  )) =>
+        if (x =y) andalso (pm1 = one) andalso (c = zero) then HOLogic.true_const else HOLogic.false_const
+
+  |(Const ("Not", _) $ p) => HOLogic.Not $ (plusinf x p)
+  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (plusinf x p) $ (plusinf x q)
+  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (plusinf x p) $ (plusinf x q)
+  |_ => fm;
+ 
+(* ------------------------------------------------------------------------- *) 
+(* The LCM of all the divisors that involve x.                               *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun divlcm x (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z ) ) =  
+        if x = y then abs(dest_numeral d) else 1 
+  |divlcm x ( Const ("Not", _) $ p) = divlcm x p 
+  |divlcm x ( Const ("op &",_) $ p $ q) = lcm_num (divlcm x p) (divlcm x q) 
+  |divlcm x ( Const ("op |",_) $ p $ q ) = lcm_num (divlcm x p) (divlcm x q) 
+  |divlcm x  _ = 1; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Construct the B-set.                                                      *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun bset x fm = case fm of 
+   (Const ("Not", _) $ p) => if (is_arith_rel p) then  
+          (case p of  
+	      (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $c2 $y) $a ) )  
+	             => if (is_arith_rel p) andalso (x=	y) andalso (c2 = one) andalso (c1 = zero)  
+	                then [linear_neg a] 
+			else  bset x p 
+   	  |_ =>[]) 
+			 
+			else bset x p 
+  |(Const ("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +",_) $ (Const ("op *",_) $c2 $ x) $ a)) =>  if (c1 =zero) andalso (c2 = one) then [linear_neg(linear_add [] a (mk_numeral 1))]  else [] 
+  |(Const ("op <",_) $ c1$ (Const ("op +",_) $(Const ("op *",_)$ c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = one) then [linear_neg a] else [] 
+  |(Const ("op &",_) $ p $ q) => (bset x p) union (bset x q) 
+  |(Const ("op |",_) $ p $ q) => (bset x p) union (bset x q) 
+  |_ => []; 
+ 
+(* ------------------------------------------------------------------------- *)
+(* Construct the A-set.                                                      *)
+(* ------------------------------------------------------------------------- *)
+
+fun aset x fm = case fm of
+   (Const ("Not", _) $ p) => if (is_arith_rel p) then
+          (case p of
+	      (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $c2 $y) $a ) )
+	             => if (x=	y) andalso (c2 = one) andalso (c1 = zero)
+	                then [linear_neg a]
+			else  []
+   	  |_ =>[])
+
+			else aset x p
+  |(Const ("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +",_) $ (Const ("op *",_) $c2 $ x) $ a)) =>  if (c1 =zero) andalso (c2 = one) then [linear_sub [] (mk_numeral 1) a]  else []
+  |(Const ("op <",_) $ c1$ (Const ("op +",_) $(Const ("op *",_)$ c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = (mk_numeral (~1))) then [a] else []
+  |(Const ("op &",_) $ p $ q) => (aset x p) union (aset x q)
+  |(Const ("op |",_) $ p $ q) => (aset x p) union (aset x q)
+  |_ => [];
+
+
+(* ------------------------------------------------------------------------- *) 
+(* Replace top variable with another linear form, retaining canonicality.    *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun linrep vars x t fm = case fm of  
+   ((Const(p,_)$ d $ (Const("op +",_)$(Const("op *",_)$ c $ y) $ z))) => 
+      if (x = y) andalso (is_arith_rel fm)  
+      then  
+        let val ct = linear_cmul (dest_numeral c) t  
+	in (HOLogic.mk_binrel p (d, linear_add vars ct z)) 
+	end 
+	else fm 
+  |(Const ("Not", _) $ p) => HOLogic.Not $ (linrep vars x t p) 
+  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (linrep vars x t p) $ (linrep vars x t q) 
+  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (linrep vars x t p) $ (linrep vars x t q) 
+  |_ => fm; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Evaluation of constant expressions.                                       *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+val operations = 
+  [("op =",op=), ("op <",op<), ("op >",op>), ("op <=",op<=) , ("op >=",op>=), 
+   ("Divides.op dvd",fn (x,y) =>((y mod x) = 0))]; 
+ 
+fun applyoperation (Some f) (a,b) = f (a, b) 
+    |applyoperation _ (_, _) = false; 
+ 
+(*Evaluation of constant atomic formulas*) 
+ 
+fun evalc_atom at = case at of  
+      (Const (p,_) $ s $ t) =>(  
+         case assoc (operations,p) of 
+             Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)  
+              handle _ => at) 
+             | _ =>  at) 
+     |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
+         case assoc (operations,p) of 
+             Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then 
+	     HOLogic.false_const else HOLogic.true_const)  
+              handle _ => at) 
+             | _ =>  at) 
+     | _ =>  at; 
+ 
+(*Function onatoms apllys function f on the atomic formulas involved in a.*) 
+ 
+fun onatoms f a = if (is_arith_rel a) then f a else case a of 
+ 
+  	(Const ("Not",_) $ p) => if is_arith_rel p then HOLogic.Not $ (f p) 
+				 
+				else HOLogic.Not $ (onatoms f p) 
+  	|(Const ("op &",_) $ p $ q) => HOLogic.conj $ (onatoms f p) $ (onatoms f q) 
+  	|(Const ("op |",_) $ p $ q) => HOLogic.disj $ (onatoms f p) $ (onatoms f q) 
+  	|(Const ("op -->",_) $ p $ q) => HOLogic.imp $ (onatoms f p) $ (onatoms f q) 
+  	|((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q) => (Const ("op =", [HOLogic.boolT, HOLogic.boolT] ---> HOLogic.boolT)) $ (onatoms f p) $ (onatoms f q) 
+  	|(Const("All",_) $ Abs(x,T,p)) => Const("All", [HOLogic.intT --> 
+	HOLogic.boolT] ---> HOLogic.boolT)$ Abs (x ,T, (onatoms f p)) 
+  	|(Const("Ex",_) $ Abs(x,T,p)) => Const("Ex", [HOLogic.intT --> HOLogic.boolT]---> HOLogic.boolT) $ Abs( x ,T, (onatoms f p)) 
+  	|_ => a; 
+ 
+val evalc = onatoms evalc_atom; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Hence overall quantifier elimination.                                     *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+(*Applyes a function iteratively on the list*) 
+ 
+fun end_itlist f []     = error "end_itlist" 
+   |end_itlist f [x]    = x 
+   |end_itlist f (h::t) = f h (end_itlist f t); 
+ 
+ 
+(*list_disj[conj] makes a disj[conj] of a given list. used with conjucts or disjuncts 
+it liearises iterated conj[disj]unctions. *) 
+ 
+fun disj_help p q = HOLogic.disj $ p $ q ; 
+ 
+fun list_disj l = 
+  if l = [] then HOLogic.false_const else end_itlist disj_help l; 
+   
+fun conj_help p q = HOLogic.conj $ p $ q ; 
+ 
+fun list_conj l = 
+  if l = [] then HOLogic.true_const else end_itlist conj_help l; 
+   
+(*Simplification of Formulas *) 
+ 
+(*Function q_bnd_chk checks if a quantified Formula makes sens : Means if in 
+the body of the existential quantifier there are bound variables to the 
+existential quantifier.*) 
+ 
+fun has_bound fm =let fun has_boundh fm i = case fm of 
+		 Bound n => (i = n) 
+		 |Abs (_,_,p) => has_boundh p (i+1) 
+		 |t1 $ t2 => (has_boundh t1 i) orelse (has_boundh t2 i) 
+		 |_ =>false
+
+in  case fm of 
+	Bound _ => true 
+       |Abs (_,_,p) => has_boundh p 0 
+       |t1 $ t2 => (has_bound t1 ) orelse (has_bound t2 ) 
+       |_ =>false
+end;
+ 
+(*has_sub_abs checks if in a given Formula there are subformulas which are quantifed 
+too. Is no used no more.*) 
+ 
+fun has_sub_abs fm = case fm of  
+		 Abs (_,_,_) => true 
+		 |t1 $ t2 => (has_bound t1 ) orelse (has_bound t2 ) 
+		 |_ =>false ; 
+		  
+(*update_bounds called with i=0 udates the numeration of bounded variables because the 
+formula will not be quantified any more.*) 
+ 
+fun update_bounds fm i = case fm of 
+		 Bound n => if n >= i then Bound (n-1) else fm 
+		 |Abs (x,T,p) => Abs(x,T,(update_bounds p (i+1))) 
+		 |t1 $ t2 => (update_bounds t1 i) $ (update_bounds t2 i) 
+		 |_ => fm ; 
+ 
+(*psimpl : Simplification of propositions (general purpose)*) 
+fun psimpl1 fm = case fm of 
+    Const("Not",_) $ Const ("False",_) => HOLogic.true_const 
+  | Const("Not",_) $ Const ("True",_) => HOLogic.false_const 
+  | Const("op &",_) $ Const ("False",_) $ q => HOLogic.false_const 
+  | Const("op &",_) $ p $ Const ("False",_)  => HOLogic.false_const 
+  | Const("op &",_) $ Const ("True",_) $ q => q 
+  | Const("op &",_) $ p $ Const ("True",_) => p 
+  | Const("op |",_) $ Const ("False",_) $ q => q 
+  | Const("op |",_) $ p $ Const ("False",_)  => p 
+  | Const("op |",_) $ Const ("True",_) $ q => HOLogic.true_const 
+  | Const("op |",_) $ p $ Const ("True",_)  => HOLogic.true_const 
+  | Const("op -->",_) $ Const ("False",_) $ q => HOLogic.true_const 
+  | Const("op -->",_) $ Const ("True",_) $  q => q 
+  | Const("op -->",_) $ p $ Const ("True",_)  => HOLogic.true_const 
+  | Const("op -->",_) $ p $ Const ("False",_)  => HOLogic.Not $  p 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ Const ("True",_) $ q => q 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ Const ("True",_) => p 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ Const ("False",_) $ q => HOLogic.Not $  q 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ Const ("False",_)  => HOLogic.Not $  p 
+  | _ => fm; 
+ 
+fun psimpl fm = case fm of 
+   Const ("Not",_) $ p => psimpl1 (HOLogic.Not $ (psimpl p)) 
+  | Const("op &",_) $ p $ q => psimpl1 (HOLogic.mk_conj (psimpl p,psimpl q)) 
+  | Const("op |",_) $ p $ q => psimpl1 (HOLogic.mk_disj (psimpl p,psimpl q)) 
+  | Const("op -->",_) $ p $ q => psimpl1 (HOLogic.mk_imp(psimpl p,psimpl q)) 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q => psimpl1 (HOLogic.mk_eq(psimpl p,psimpl q)) 
+  | _ => fm; 
+ 
+ 
+(*simpl : Simplification of Terms involving quantifiers too. 
+ This function is able to drop out some quantified expressions where there are no 
+ bound varaibles.*) 
+  
+fun simpl1 fm  = 
+  case fm of 
+    Const("All",_) $Abs(x,_,p) => if (has_bound fm ) then fm  
+    				else (update_bounds p 0) 
+  | Const("Ex",_) $ Abs (x,_,p) => if has_bound fm then fm  
+    				else (update_bounds p 0) 
+  | _ => psimpl1 fm; 
+ 
+fun simpl fm = case fm of 
+    Const ("Not",_) $ p => simpl1 (HOLogic.Not $(simpl p))  
+  | Const ("op &",_) $ p $ q => simpl1 (HOLogic.mk_conj (simpl p ,simpl q))  
+  | Const ("op |",_) $ p $ q => simpl1 (HOLogic.mk_disj (simpl p ,simpl q ))  
+  | Const ("op -->",_) $ p $ q => simpl1 (HOLogic.mk_imp(simpl p ,simpl q ))  
+  | Const("op =", Type ("fun",[Type ("bool", []),_]))$ p $ q => simpl1 
+  (HOLogic.mk_eq(simpl p ,simpl q ))  
+  | Const ("All",Ta) $ Abs(Vn,VT,p) => simpl1(Const("All",Ta) $ 
+  Abs(Vn,VT,simpl p ))  
+  | Const ("Ex",Ta)  $ Abs(Vn,VT,p) => simpl1(Const("Ex",Ta)  $ 
+  Abs(Vn,VT,simpl p ))  
+  | _ => fm; 
+ 
+(* ------------------------------------------------------------------------- *) 
+ 
+(* Puts fm into NNF*) 
+ 
+fun  nnf fm = if (is_arith_rel fm) then fm  
+else (case fm of 
+  ( Const ("op &",_) $ p $ q)  => HOLogic.conj $ (nnf p) $(nnf q) 
+  | (Const("op |",_) $ p $q) => HOLogic.disj $ (nnf p)$(nnf q) 
+  | (Const ("op -->",_)  $ p $ q) => HOLogic.disj $ (nnf (HOLogic.Not $ p)) $ (nnf q) 
+  | ((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q) =>(HOLogic.disj $ (HOLogic.conj $ (nnf p) $ (nnf q)) $ (HOLogic.conj $ (nnf (HOLogic.Not $ p) ) $ (nnf(HOLogic.Not $ q)))) 
+  | (Const ("Not",_)) $ ((Const ("Not",_)) $ p) => (nnf p) 
+  | (Const ("Not",_)) $ (( Const ("op &",_)) $ p $ q) =>HOLogic.disj $ (nnf(HOLogic.Not $ p)) $ (nnf(HOLogic.Not $q)) 
+  | (Const ("Not",_)) $ (( Const ("op |",_)) $ p $ q) =>HOLogic.conj $ (nnf(HOLogic.Not $ p)) $ (nnf(HOLogic.Not $ q)) 
+  | (Const ("Not",_)) $ (( Const ("op -->",_)) $ p $ q ) =>HOLogic.conj $ (nnf p) $(nnf(HOLogic.Not $ q)) 
+  | (Const ("Not",_)) $ ((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q ) =>(HOLogic.disj $ (HOLogic.conj $(nnf p) $ (nnf(HOLogic.Not $ q))) $ (HOLogic.conj $(nnf(HOLogic.Not $ p)) $ (nnf q))) 
+  | _ => fm); 
+ 
+ 
+(* Function remred to remove redundancy in a list while keeping the order of appearance of the 
+elements. but VERY INEFFICIENT!! *) 
+ 
+fun remred1 el [] = [] 
+    |remred1 el (h::t) = if el=h then (remred1 el t) else h::(remred1 el t); 
+     
+fun remred [] = [] 
+    |remred (x::l) =  x::(remred1 x (remred l)); 
+ 
+(*Makes sure that all free Variables are of the type integer but this function is only 
+used temporarily, this job must be done by the parser later on.*) 
+ 
+fun mk_uni_vars T  (node $ rest) = (case node of 
+    Free (name,_) => Free (name,T) $ (mk_uni_vars T rest) 
+    |_=> (mk_uni_vars T node) $ (mk_uni_vars T rest )  ) 
+    |mk_uni_vars T (Free (v,_)) = Free (v,T) 
+    |mk_uni_vars T tm = tm; 
+ 
+fun mk_uni_int T (Const ("0",T2)) = if T = T2 then (mk_numeral 0) else (Const ("0",T2)) 
+    |mk_uni_int T (Const ("1",T2)) = if T = T2 then (mk_numeral 1) else (Const ("1",T2)) 
+    |mk_uni_int T (node $ rest) = (mk_uni_int T node) $ (mk_uni_int T rest )  
+    |mk_uni_int T (Abs(AV,AT,p)) = Abs(AV,AT,mk_uni_int T p) 
+    |mk_uni_int T tm = tm; 
+ 
+
+(* Minusinfinity Version*) 
+fun coopermi vars1 fm = 
+  case fm of 
+   Const ("Ex",_) $ Abs(x0,T,p0) => let 
+    val (xn,p1) = variant_abs (x0,T,p0) 
+    val x = Free (xn,T)  
+    val vars = (xn::vars1) 
+    val p = unitycoeff x  (posineq (simpl p1))
+    val p_inf = simpl (minusinf x p) 
+    val bset = bset x p 
+    val js = 1 upto divlcm x p  
+    fun p_element j b = linrep vars x (linear_add vars b (mk_numeral j)) p  
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) bset)  
+   in (list_disj (map stage js))
+    end 
+  | _ => error "cooper: not an existential formula"; 
+ 
+
+
+(* The plusinfinity version of cooper*)
+fun cooperpi vars1 fm =
+  case fm of
+   Const ("Ex",_) $ Abs(x0,T,p0) => let 
+    val (xn,p1) = variant_abs (x0,T,p0)
+    val x = Free (xn,T)
+    val vars = (xn::vars1)
+    val p = unitycoeff x  (posineq (simpl p1))
+    val p_inf = simpl (plusinf x p)
+    val aset = aset x p
+    val js = 1 upto divlcm x p
+    fun p_element j a = linrep vars x (linear_sub vars a (mk_numeral j)) p
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) aset)
+   in (list_disj (map stage js))
+   end
+  | _ => error "cooper: not an existential formula";
+  
+
+
+(*Cooper main procedure*) 
+  
+fun cooper vars1 fm =
+  case fm of
+   Const ("Ex",_) $ Abs(x0,T,p0) => let 
+    val (xn,p1) = variant_abs (x0,T,p0)
+    val x = Free (xn,T)
+    val vars = (xn::vars1)
+    val p = unitycoeff x  (posineq (simpl p1))
+    val ast = aset x p
+    val bst = bset x p
+    val js = 1 upto divlcm x p
+    val (p_inf,f,S ) = 
+    if (length bst) < (length ast) 
+     then (minusinf x p,linear_add,bst)
+     else (plusinf x p, linear_sub,ast)
+    fun p_element j a = linrep vars x (f vars a (mk_numeral j)) p
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) S)
+   in (list_disj (map stage js))
+   end
+  | _ => error "cooper: not an existential formula";
+
+
+
+ 
+(*Function itlist applys a double parametred function f : 'a->'b->b iteratively to a List l : 'a 
+list With End condition b. ict calculates f(e1,f(f(e2,f(e3,...(...f(en,b))..))))) 
+ assuming l = [e1,e2,...,en]*) 
+ 
+fun itlist f l b = case l of 
+    [] => b 
+  | (h::t) => f h (itlist f t b); 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Free variables in terms and formulas.	                             *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun fvt tml = case tml of 
+    [] => [] 
+  | Free(x,_)::r => x::(fvt r) 
+ 
+fun fv fm = fvt (term_frees fm); 
+ 
+ 
+(* ========================================================================= *) 
+(* Quantifier elimination.                                                   *) 
+(* ========================================================================= *) 
+(*conj[/disj]uncts lists iterated conj[disj]unctions*) 
+ 
+fun disjuncts fm = case fm of 
+    Const ("op |",_) $ p $ q => (disjuncts p) @ (disjuncts q) 
+  | _ => [fm]; 
+ 
+fun conjuncts fm = case fm of 
+    Const ("op &",_) $p $ q => (conjuncts p) @ (conjuncts q) 
+  | _ => [fm]; 
+ 
+ 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Lift procedure given literal modifier, formula normalizer & basic quelim. *) 
+(* ------------------------------------------------------------------------- *) 
+   
+fun lift_qelim afn nfn qfn isat = 
+ let   fun qelim x vars p = 
+  let val cjs = conjuncts p 
+      val (ycjs,ncjs) = partition (has_bound) cjs in 
+      (if ycjs = [] then p else 
+                          let val q = (qfn vars ((HOLogic.exists_const HOLogic.intT 
+			  ) $ Abs(x,HOLogic.intT,(list_conj ycjs)))) in 
+                          (itlist conj_help ncjs q)  
+			  end) 
+       end 
+    
+  fun qelift vars fm = if (isat fm) then afn vars fm 
+    else  
+    case fm of 
+      Const ("Not",_) $ p => HOLogic.Not $ (qelift vars p) 
+    | Const ("op &",_) $ p $q => HOLogic.conj $ (qelift vars p) $ (qelift vars q) 
+    | Const ("op |",_) $ p $ q => HOLogic.disj $ (qelift vars p) $ (qelift vars q) 
+    | Const ("op -->",_) $ p $ q => HOLogic.imp $ (qelift vars p) $ (qelift vars q) 
+    | Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q => HOLogic.mk_eq ((qelift vars p),(qelift vars q)) 
+    | Const ("All",QT) $ Abs(x,T,p) => HOLogic.Not $(qelift vars (Const ("Ex",QT) $ Abs(x,T,(HOLogic.Not $ p)))) 
+    | Const ("Ex",_) $ Abs (x,T,p)  => let  val djs = disjuncts(nfn(qelift (x::vars) p)) in 
+    			list_disj(map (qelim x vars) djs) end 
+    | _ => fm 
+ 
+  in (fn fm => simpl(qelift (fv fm) fm)) 
+  end; 
+ 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Cleverer (proposisional) NNF with conditional and literal modification.   *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+(*Function Negate used by cnnf, negates a formula p*) 
+ 
+fun negate (Const ("Not",_) $ p) = p 
+    |negate p = (HOLogic.Not $ p); 
+ 
+fun cnnf lfn = 
+  let fun cnnfh fm = case  fm of 
+      (Const ("op &",_) $ p $ q) => HOLogic.mk_conj(cnnfh p,cnnfh q) 
+    | (Const ("op |",_) $ p $ q) => HOLogic.mk_disj(cnnfh p,cnnfh q) 
+    | (Const ("op -->",_) $ p $q) => HOLogic.mk_disj(cnnfh(HOLogic.Not $ p),cnnfh q) 
+    | (Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q) => HOLogic.mk_disj( 
+    		HOLogic.mk_conj(cnnfh p,cnnfh q), 
+		HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $q))) 
+
+    | (Const ("Not",_) $ (Const("Not",_) $ p)) => cnnfh p 
+    | (Const ("Not",_) $ (Const ("op &",_) $ p $ q)) => HOLogic.mk_disj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $ q)) 
+    | (Const ("Not",_) $(Const ("op |",_) $ (Const ("op &",_) $ p $ q) $  
+    			(Const ("op &",_) $ p1 $ r))) => if p1 = negate p then 
+		         HOLogic.mk_disj(  
+			   cnnfh (HOLogic.mk_conj(p,cnnfh(HOLogic.Not $ q))), 
+			   cnnfh (HOLogic.mk_conj(p1,cnnfh(HOLogic.Not $ r)))) 
+			 else  HOLogic.mk_conj(
+			  cnnfh (HOLogic.mk_disj(cnnfh (HOLogic.Not $ p),cnnfh(HOLogic.Not $ q))), 
+			   cnnfh (HOLogic.mk_disj(cnnfh (HOLogic.Not $ p1),cnnfh(HOLogic.Not $ r)))
+			 ) 
+    | (Const ("Not",_) $ (Const ("op |",_) $ p $ q)) => HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $ q)) 
+    | (Const ("Not",_) $ (Const ("op -->",_) $ p $q)) => HOLogic.mk_conj(cnnfh p,cnnfh(HOLogic.Not $ q)) 
+    | (Const ("Not",_) $ (Const ("op =",Type ("fun",[Type ("bool", []),_]))  $ p $ q)) => HOLogic.mk_disj(HOLogic.mk_conj(cnnfh p,cnnfh(HOLogic.Not $ q)),HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh q)) 
+    | _ => lfn fm  
+  in cnnfh o simpl
+  end; 
+ 
+(*End- function the quantifierelimination an decion procedure of presburger formulas.*)   
+val integer_qelim = simpl o evalc o (lift_qelim linform (simpl o (cnnf posineq o evalc)) cooper is_arith_rel) ; 
+
+end;
+ 
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Integ/cooper_proof.ML	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,1488 @@
+(*  Title:      HOL/Integ/cooper_proof.ML
+    ID:         $Id$
+    Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+File containing the implementation of the proof
+generation for Cooper Algorithm
+*)
+
+signature COOPER_PROOF =
+sig
+  val qe_Not : thm
+  val qe_conjI : thm
+  val qe_disjI : thm
+  val qe_impI : thm
+  val qe_eqI : thm
+  val qe_exI : thm
+  val qe_get_terms : thm -> term * term
+  val cooper_prv : Sign.sg -> term -> term -> string list -> thm
+  val proof_of_evalc : Sign.sg -> term -> thm
+  val proof_of_cnnf : Sign.sg -> term -> (term -> thm) -> thm
+  val proof_of_linform : Sign.sg -> string list -> term -> thm
+end;
+
+structure CooperProof : COOPER_PROOF =
+struct
+
+open CooperDec;
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---         Protocoling part                                  ---*)
+(*---                                                           ---*)
+(*---           includes the protocolling datastructure         ---*)
+(*---                                                           ---*)
+(*---          and the protocolling fuctions                    ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+val presburger_ss = simpset_of (theory "Presburger")
+  addsimps [zdiff_def] delsimps [symmetric zdiff_def];
+val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT;
+
+(*Theorems that will be used later for the proofgeneration*)
+
+val zdvd_iff_zmod_eq_0 = thm "zdvd_iff_zmod_eq_0";
+val unity_coeff_ex = thm "unity_coeff_ex";
+
+(* Thorems for proving the adjustment of the coeffitients*)
+
+val ac_lt_eq =  thm "ac_lt_eq";
+val ac_eq_eq = thm "ac_eq_eq";
+val ac_dvd_eq = thm "ac_dvd_eq";
+val ac_pi_eq = thm "ac_pi_eq";
+
+(* The logical compination of the sythetised properties*)
+val qe_Not = thm "qe_Not";
+val qe_conjI = thm "qe_conjI";
+val qe_disjI = thm "qe_disjI";
+val qe_impI = thm "qe_impI";
+val qe_eqI = thm "qe_eqI";
+val qe_exI = thm "qe_exI";
+val qe_ALLI = thm "qe_ALLI";
+
+(*Modulo D property for Pminusinf an Plusinf *)
+val fm_modd_minf = thm "fm_modd_minf";
+val not_dvd_modd_minf = thm "not_dvd_modd_minf";
+val dvd_modd_minf = thm "dvd_modd_minf";
+
+val fm_modd_pinf = thm "fm_modd_pinf";
+val not_dvd_modd_pinf = thm "not_dvd_modd_pinf";
+val dvd_modd_pinf = thm "dvd_modd_pinf";
+
+(* the minusinfinity proprty*)
+
+val fm_eq_minf = thm "fm_eq_minf";
+val neq_eq_minf = thm "neq_eq_minf";
+val eq_eq_minf = thm "eq_eq_minf";
+val le_eq_minf = thm "le_eq_minf";
+val len_eq_minf = thm "len_eq_minf";
+val not_dvd_eq_minf = thm "not_dvd_eq_minf";
+val dvd_eq_minf = thm "dvd_eq_minf";
+
+(* the Plusinfinity proprty*)
+
+val fm_eq_pinf = thm "fm_eq_pinf";
+val neq_eq_pinf = thm "neq_eq_pinf";
+val eq_eq_pinf = thm "eq_eq_pinf";
+val le_eq_pinf = thm "le_eq_pinf";
+val len_eq_pinf = thm "len_eq_pinf";
+val not_dvd_eq_pinf = thm "not_dvd_eq_pinf";
+val dvd_eq_pinf = thm "dvd_eq_pinf";
+
+(*Logical construction of the Property*)
+val eq_minf_conjI = thm "eq_minf_conjI";
+val eq_minf_disjI = thm "eq_minf_disjI";
+val modd_minf_disjI = thm "modd_minf_disjI";
+val modd_minf_conjI = thm "modd_minf_conjI";
+
+val eq_pinf_conjI = thm "eq_pinf_conjI";
+val eq_pinf_disjI = thm "eq_pinf_disjI";
+val modd_pinf_disjI = thm "modd_pinf_disjI";
+val modd_pinf_conjI = thm "modd_pinf_conjI";
+
+(*A/B - set Theorem *)
+
+val bst_thm = thm "bst_thm";
+val ast_thm = thm "ast_thm";
+
+(*Cooper Backwards...*)
+(*Bset*)
+val not_bst_p_fm = thm "not_bst_p_fm";
+val not_bst_p_ne = thm "not_bst_p_ne";
+val not_bst_p_eq = thm "not_bst_p_eq";
+val not_bst_p_gt = thm "not_bst_p_gt";
+val not_bst_p_lt = thm "not_bst_p_lt";
+val not_bst_p_ndvd = thm "not_bst_p_ndvd";
+val not_bst_p_dvd = thm "not_bst_p_dvd";
+
+(*Aset*)
+val not_ast_p_fm = thm "not_ast_p_fm";
+val not_ast_p_ne = thm "not_ast_p_ne";
+val not_ast_p_eq = thm "not_ast_p_eq";
+val not_ast_p_gt = thm "not_ast_p_gt";
+val not_ast_p_lt = thm "not_ast_p_lt";
+val not_ast_p_ndvd = thm "not_ast_p_ndvd";
+val not_ast_p_dvd = thm "not_ast_p_dvd";
+
+(*Logical construction of the prop*)
+(*Bset*)
+val not_bst_p_conjI = thm "not_bst_p_conjI";
+val not_bst_p_disjI = thm "not_bst_p_disjI";
+val not_bst_p_Q_elim = thm "not_bst_p_Q_elim";
+
+(*Aset*)
+val not_ast_p_conjI = thm "not_ast_p_conjI";
+val not_ast_p_disjI = thm "not_ast_p_disjI";
+val not_ast_p_Q_elim = thm "not_ast_p_Q_elim";
+
+(*Cooper*)
+val cppi_eq = thm "cppi_eq";
+val cpmi_eq = thm "cpmi_eq";
+
+(*Others*)
+val simp_from_to = thm "simp_from_to";
+val P_eqtrue = thm "P_eqtrue";
+val P_eqfalse = thm "P_eqfalse";
+
+(*For Proving NNF*)
+
+val nnf_nn = thm "nnf_nn";
+val nnf_im = thm "nnf_im";
+val nnf_eq = thm "nnf_eq";
+val nnf_sdj = thm "nnf_sdj";
+val nnf_ncj = thm "nnf_ncj";
+val nnf_nim = thm "nnf_nim";
+val nnf_neq = thm "nnf_neq";
+val nnf_ndj = thm "nnf_ndj";
+
+(*For Proving term linearizition*)
+val linearize_dvd = thm "linearize_dvd";
+val lf_lt = thm "lf_lt";
+val lf_eq = thm "lf_eq";
+val lf_dvd = thm "lf_dvd";
+
+
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the cooperprocedure.*)
+(* ------------------------------------------------------------------------- *)
+
+
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the adjustcoeff step.*)
+(* ------------------------------------------------------------------------- *)
+datatype CpLog = No
+                |Simp of term*CpLog
+		|Blast of CpLog*CpLog
+		|Aset of (term*term*(term list)*term)
+		|Bset of (term*term*(term list)*term)
+		|Minusinf of CpLog*CpLog
+		|Cooper of term*CpLog*CpLog*CpLog
+		|Eq_minf of term*term
+		|Modd_minf of term*term
+		|Eq_minf_conjI of CpLog*CpLog
+		|Modd_minf_conjI of CpLog*CpLog	
+		|Modd_minf_disjI of CpLog*CpLog
+		|Eq_minf_disjI of CpLog*CpLog	
+		|Not_bst_p of term*term*term*term*CpLog
+		|Not_bst_p_atomic of term
+		|Not_bst_p_conjI of CpLog*CpLog
+		|Not_bst_p_disjI of CpLog*CpLog
+		|Not_ast_p of term*term*term*term*CpLog
+		|Not_ast_p_atomic of term
+		|Not_ast_p_conjI of CpLog*CpLog
+		|Not_ast_p_disjI of CpLog*CpLog
+		|CpLogError;
+
+
+
+datatype ACLog = ACAt of int*term
+                |ACPI of int*term
+                |ACfm of term
+                |ACNeg of ACLog
+		|ACConst of string*ACLog*ACLog;
+
+
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the CNNF step.*)
+(* ------------------------------------------------------------------------- *)
+
+
+datatype NNFLog = NNFAt of term
+                |NNFSimp of NNFLog
+                |NNFNN of NNFLog
+		|NNFConst of string*NNFLog*NNFLog;
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the linform  step.*)
+(* ------------------------------------------------------------------------- *)
+
+
+datatype LfLog = LfAt of term
+                |LfAtdvd of term
+                |Lffm of term
+                |LfConst of string*LfLog*LfLog
+		|LfNot of LfLog
+		|LfQ of string*string*typ*LfLog;
+
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the evaluation- evalc-  step.*)
+(* ------------------------------------------------------------------------- *)
+
+
+datatype EvalLog = EvalAt of term
+                |Evalfm of term
+		|EvalConst of string*EvalLog*EvalLog;
+
+(* ------------------------------------------------------------------------- *)
+(*This function norm_zero_one  replaces the occurences of Numeral1 and Numeral0*)
+(*Respectively by their abstract representation Const("1",..) and COnst("0",..)*)
+(*this is necessary because the theorems use this representation.*)
+(* This function should be elminated in next versions...*)
+(* ------------------------------------------------------------------------- *)
+
+fun norm_zero_one fm = case fm of
+  (Const ("op *",_) $ c $ t) => 
+    if c = one then (norm_zero_one t)
+    else if (dest_numeral c = ~1) 
+         then (Const("uminus",HOLogic.intT --> HOLogic.intT) $ (norm_zero_one t))
+         else (HOLogic.mk_binop "op *" (norm_zero_one c,norm_zero_one t))
+  |(node $ rest) => ((norm_zero_one node)$(norm_zero_one rest))
+  |(Abs(x,T,p)) => (Abs(x,T,(norm_zero_one p)))
+  |_ => fm;
+
+
+(* ------------------------------------------------------------------------- *)
+(* Intended to tell that here we changed the structure of the formula with respect to the posineq theorem : ~(0 < t) = 0 < 1-t*)
+(* ------------------------------------------------------------------------- *)
+fun adjustcoeffeq_wp  x l fm = 
+    case fm of  
+  (Const("Not",_)$(Const("op <",_) $(Const("0",_)) $(rt as (Const ("op +", _)$(Const ("op *",_) $    c $ y ) $z )))) => 
+  if (x = y) 
+  then let  
+       val m = l div (dest_numeral c) 
+       val n = abs (m)
+       val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x)) 
+       val rs = (HOLogic.mk_binrel "op <" (zero,linear_sub [] one (HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
+       in (ACPI(n,fm),rs)
+       end
+  else  let val rs = (HOLogic.mk_binrel "op <" (zero,linear_sub [] one rt )) 
+        in (ACPI(1,fm),rs)
+        end
+
+  |(Const(p,_) $d $( Const ("op +", _)$(Const ("op *",_) $ 
+      c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then  
+        let val m = l div (dest_numeral c) 
+           val n = (if p = "op <" then abs(m) else m)  
+           val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x))
+           val rs = (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
+	   in (ACAt(n,fm),rs)
+	   end
+        else (ACfm(fm),fm) 
+  |( Const ("Not", _) $ p) => let val (rsp,rsr) = adjustcoeffeq_wp x l p 
+                              in (ACNeg(rsp),HOLogic.Not $ rsr) 
+                              end
+  |( Const ("op &",_) $ p $ q) =>let val (rspp,rspr) = adjustcoeffeq_wp x l p
+                                     val (rsqp,rsqr) = adjustcoeffeq_wp x l q
+
+                                  in (ACConst ("CJ",rspp,rsqp), HOLogic.mk_conj (rspr,rsqr)) 
+                                  end 
+  |( Const ("op |",_) $ p $ q) =>let val (rspp,rspr) = adjustcoeffeq_wp x l p
+                                     val (rsqp,rsqr) = adjustcoeffeq_wp x l q
+
+                                  in (ACConst ("DJ",rspp,rsqp), HOLogic.mk_disj (rspr,rsqr)) 
+                                  end
+
+  |_ => (ACfm(fm),fm);
+
+
+(*_________________________________________*)
+(*-----------------------------------------*)
+(* Protocol generation for the liform step *)
+(*_________________________________________*)
+(*-----------------------------------------*)
+
+
+fun linform_wp fm = 
+  let fun at_linform_wp at =
+    case at of
+      (Const("op <=",_)$s$t) => LfAt(at)
+      |(Const("op <",_)$s$t) => LfAt(at)
+      |(Const("op =",_)$s$t) => LfAt(at)
+      |(Const("Divides.op dvd",_)$s$t) => LfAtdvd(at)
+  in
+  if is_arith_rel fm 
+  then at_linform_wp fm 
+  else case fm of
+    (Const("Not",_) $ A) => LfNot(linform_wp A)
+   |(Const("op &",_)$ A $ B) => LfConst("CJ",linform_wp A, linform_wp B)
+   |(Const("op |",_)$ A $ B) => LfConst("DJ",linform_wp A, linform_wp B)
+   |(Const("op -->",_)$ A $ B) => LfConst("IM",linform_wp A, linform_wp B)
+   |(Const("op =",Type ("fun",[Type ("bool", []),_]))$ A $ B) => LfConst("EQ",linform_wp A, linform_wp B)
+   |Const("Ex",_)$Abs(x,T,p) => 
+     let val (xn,p1) = variant_abs(x,T,p)
+     in LfQ("Ex",xn,T,linform_wp p1)
+     end 
+   |Const("All",_)$Abs(x,T,p) => 
+     let val (xn,p1) = variant_abs(x,T,p)
+     in LfQ("All",xn,T,linform_wp p1)
+     end 
+end;
+
+
+(* ------------------------------------------------------------------------- *)
+(*For simlified formulas we just notice the original formula, for whitch we habe been
+intendes to make the proof.*)
+(* ------------------------------------------------------------------------- *)
+fun simpl_wp (fm,pr) = let val fm2 = simpl fm
+				in (fm2,Simp(fm,pr))
+				end;
+
+	
+(* ------------------------------------------------------------------------- *)
+(*Help function for the generation of the proof EX.P_{minus \infty} --> EX. P(x) *)
+(* ------------------------------------------------------------------------- *)
+fun minusinf_wph x fm = let fun mk_atomar_minusinf_proof x fm = (Modd_minf(x,fm),Eq_minf(x,fm))
+  
+	      fun combine_minusinf_proofs opr (ppr1,ppr2) (qpr1,qpr2) = case opr of 
+		 "CJ" => (Modd_minf_conjI(ppr1,qpr1),Eq_minf_conjI(ppr2,qpr2))
+		|"DJ" => (Modd_minf_disjI(ppr1,qpr1),Eq_minf_disjI(ppr2,qpr2))
+	in 
+ 
+ case fm of 
+ (Const ("Not", _) $  (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+     if (x=y) andalso (c1= zero) andalso (c2= one) then (HOLogic.true_const ,(mk_atomar_minusinf_proof x fm))
+        else (fm ,(mk_atomar_minusinf_proof x fm))
+ |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	 if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one)
+	 then (HOLogic.false_const ,(mk_atomar_minusinf_proof x fm))
+	 				 else (fm,(mk_atomar_minusinf_proof x fm)) 
+ |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y ) $ z )) =>
+       if (y=x) andalso (c1 = zero) then 
+        if c2 = one then (HOLogic.false_const,(mk_atomar_minusinf_proof x fm)) else
+	(HOLogic.true_const,(mk_atomar_minusinf_proof x fm))
+	else (fm,(mk_atomar_minusinf_proof x fm))
+  
+  |(Const("Not",_)$(Const ("Divides.op dvd",_) $_ )) => (fm,mk_atomar_minusinf_proof x fm)
+  
+  |(Const ("Divides.op dvd",_) $_ ) => (fm,mk_atomar_minusinf_proof x fm)
+  
+  |(Const ("op &",_) $ p $ q) => let val (pfm,ppr) = minusinf_wph x p
+  				    val (qfm,qpr) = minusinf_wph x q
+				    val pr = (combine_minusinf_proofs "CJ" ppr qpr)
+				     in 
+				     (HOLogic.conj $ pfm $qfm , pr)
+				     end 
+  |(Const ("op |",_) $ p $ q) => let val (pfm,ppr) = minusinf_wph x p
+  				     val (qfm,qpr) = minusinf_wph x q
+				     val pr = (combine_minusinf_proofs "DJ" ppr qpr)
+				     in 
+				     (HOLogic.disj $ pfm $qfm , pr)
+				     end 
+
+  |_ => (fm,(mk_atomar_minusinf_proof x fm))
+  
+  end;					 
+(* ------------------------------------------------------------------------- *)	    (* Protokol for the Proof of the property of the minusinfinity formula*)
+(* Just combines the to protokols *)
+(* ------------------------------------------------------------------------- *)
+fun minusinf_wp x fm  = let val (fm2,pr) = (minusinf_wph x fm)
+                       in (fm2,Minusinf(pr))
+                        end;
+
+(* ------------------------------------------------------------------------- *)
+(*Help function for the generation of the proof EX.P_{plus \infty} --> EX. P(x) *)
+(* ------------------------------------------------------------------------- *)
+
+fun plusinf_wph x fm = let fun mk_atomar_plusinf_proof x fm = (Modd_minf(x,fm),Eq_minf(x,fm))
+  
+	      fun combine_plusinf_proofs opr (ppr1,ppr2) (qpr1,qpr2) = case opr of 
+		 "CJ" => (Modd_minf_conjI(ppr1,qpr1),Eq_minf_conjI(ppr2,qpr2))
+		|"DJ" => (Modd_minf_disjI(ppr1,qpr1),Eq_minf_disjI(ppr2,qpr2))
+	in 
+ 
+ case fm of 
+ (Const ("Not", _) $  (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+     if (x=y) andalso (c1= zero) andalso (c2= one) then (HOLogic.true_const ,(mk_atomar_plusinf_proof x fm))
+        else (fm ,(mk_atomar_plusinf_proof x fm))
+ |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	 if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one)
+	 then (HOLogic.false_const ,(mk_atomar_plusinf_proof x fm))
+	 				 else (fm,(mk_atomar_plusinf_proof x fm)) 
+ |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y ) $ z )) =>
+       if (y=x) andalso (c1 = zero) then 
+        if c2 = one then (HOLogic.true_const,(mk_atomar_plusinf_proof x fm)) else
+	(HOLogic.false_const,(mk_atomar_plusinf_proof x fm))
+	else (fm,(mk_atomar_plusinf_proof x fm))
+  
+  |(Const("Not",_)$(Const ("Divides.op dvd",_) $_ )) => (fm,mk_atomar_plusinf_proof x fm)
+  
+  |(Const ("Divides.op dvd",_) $_ ) => (fm,mk_atomar_plusinf_proof x fm)
+  
+  |(Const ("op &",_) $ p $ q) => let val (pfm,ppr) = plusinf_wph x p
+  				    val (qfm,qpr) = plusinf_wph x q
+				    val pr = (combine_plusinf_proofs "CJ" ppr qpr)
+				     in 
+				     (HOLogic.conj $ pfm $qfm , pr)
+				     end 
+  |(Const ("op |",_) $ p $ q) => let val (pfm,ppr) = plusinf_wph x p
+  				     val (qfm,qpr) = plusinf_wph x q
+				     val pr = (combine_plusinf_proofs "DJ" ppr qpr)
+				     in 
+				     (HOLogic.disj $ pfm $qfm , pr)
+				     end 
+
+  |_ => (fm,(mk_atomar_plusinf_proof x fm))
+  
+  end;					 
+(* ------------------------------------------------------------------------- *)	    (* Protokol for the Proof of the property of the minusinfinity formula*)
+(* Just combines the to protokols *)
+(* ------------------------------------------------------------------------- *)
+fun plusinf_wp x fm  = let val (fm2,pr) = (plusinf_wph x fm)
+                       in (fm2,Minusinf(pr))
+                        end;
+
+
+(* ------------------------------------------------------------------------- *)
+(*Protocol that we here uses Bset.*)
+(* ------------------------------------------------------------------------- *)
+fun bset_wp x fm = let val bs = bset x fm in
+				(bs,Bset(x,fm,bs,mk_numeral (divlcm x fm)))
+				end;
+
+(* ------------------------------------------------------------------------- *)
+(*Protocol that we here uses Aset.*)
+(* ------------------------------------------------------------------------- *)
+fun aset_wp x fm = let val ast = aset x fm in
+				(ast,Aset(x,fm,ast,mk_numeral (divlcm x fm)))
+				end;
+ 
+
+
+(* ------------------------------------------------------------------------- *)
+(*function list to Set, constructs a set containing all elements of a given list.*)
+(* ------------------------------------------------------------------------- *)
+fun list_to_set T1 l = let val T = (HOLogic.mk_setT T1) in 
+	case l of 
+		[] => Const ("{}",T)
+		|(h::t) => Const("insert", T1 --> (T --> T)) $ h $(list_to_set T1 t)
+		end;
+		
+
+(*====================================================================*)
+(* ------------------------------------------------------------------------- *)
+(* ------------------------------------------------------------------------- *)
+(*Protocol for the proof of the backward direction of the cooper theorem.*)
+(* Helpfunction - Protokols evereything about the proof reconstruction*)
+(* ------------------------------------------------------------------------- *)
+fun not_bst_p_wph fm = case fm of
+	Const("Not",_) $ R => if (is_arith_rel R) then (Not_bst_p_atomic (fm)) else CpLogError
+	|Const("op &",_) $ ls $ rs => Not_bst_p_conjI((not_bst_p_wph ls),(not_bst_p_wph rs))
+	|Const("op |",_) $ ls $ rs => Not_bst_p_disjI((not_bst_p_wph ls),(not_bst_p_wph rs))
+	|_ => Not_bst_p_atomic (fm);
+(* ------------------------------------------------------------------------- *)	
+(* Main protocoling function for the backward direction gives the Bset and the divlcm and the Formula herself. Needed as inherited attributes for the proof reconstruction*)
+(* ------------------------------------------------------------------------- *)
+fun not_bst_p_wp x fm = let val prt = not_bst_p_wph fm
+			    val D = mk_numeral (divlcm x fm)
+			    val B = map norm_zero_one (bset x fm)
+			in (Not_bst_p (x,fm,D,(list_to_set HOLogic.intT B) , prt))
+			end;
+(*====================================================================*)
+(* ------------------------------------------------------------------------- *)
+(* ------------------------------------------------------------------------- *)
+(*Protocol for the proof of the backward direction of the cooper theorem.*)
+(* Helpfunction - Protokols evereything about the proof reconstruction*)
+(* ------------------------------------------------------------------------- *)
+fun not_ast_p_wph fm = case fm of
+	Const("Not",_) $ R => if (is_arith_rel R) then (Not_ast_p_atomic (fm)) else CpLogError
+	|Const("op &",_) $ ls $ rs => Not_ast_p_conjI((not_ast_p_wph ls),(not_ast_p_wph rs))
+	|Const("op |",_) $ ls $ rs => Not_ast_p_disjI((not_ast_p_wph ls),(not_ast_p_wph rs))
+	|_ => Not_ast_p_atomic (fm);
+(* ------------------------------------------------------------------------- *)	
+(* Main protocoling function for the backward direction gives the Bset and the divlcm and the Formula herself. Needed as inherited attributes for the proof reconstruction*)
+(* ------------------------------------------------------------------------- *)
+fun not_ast_p_wp x fm = let val prt = not_ast_p_wph fm
+			    val D = mk_numeral (divlcm x fm)
+			    val B = map norm_zero_one (aset x fm)
+			in (Not_ast_p (x,fm,D,(list_to_set HOLogic.intT B) , prt))
+			end;
+
+(*======================================================*)
+(* Protokolgeneration for the formula evaluation process*)
+(*======================================================*)
+
+fun evalc_wp fm = 
+  let fun evalc_atom_wp at =case at of  
+    (Const (p,_) $ s $ t) =>(  
+    case assoc (operations,p) of 
+        Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then EvalAt(HOLogic.mk_eq(at,HOLogic.true_const)) else EvalAt(HOLogic.mk_eq(at, HOLogic.false_const)))  
+		   handle _ => Evalfm(at)) 
+        | _ =>  Evalfm(at)) 
+     |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
+       case assoc (operations,p) of 
+         Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then 
+	  EvalAt(HOLogic.mk_eq(at, HOLogic.false_const))  else EvalAt(HOLogic.mk_eq(at,HOLogic.true_const)))  
+		      handle _ => Evalfm(at)) 
+         | _ => Evalfm(at)) 
+     | _ => Evalfm(at)  
+ 
+  in
+   case fm of
+    (Const("op &",_)$A$B) => EvalConst("CJ",evalc_wp A,evalc_wp B)
+   |(Const("op |",_)$A$B) => EvalConst("DJ",evalc_wp A,evalc_wp B) 
+   |(Const("op -->",_)$A$B) => EvalConst("IM",evalc_wp A,evalc_wp B) 
+   |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => EvalConst("EQ",evalc_wp A,evalc_wp B) 
+   |_ => evalc_atom_wp fm
+  end;
+
+
+
+(*======================================================*)
+(* Protokolgeneration for the NNF Transformation        *)
+(*======================================================*)
+
+fun cnnf_wp f = 
+  let fun hcnnf_wp fm =
+    case fm of
+    (Const ("op &",_) $ p $ q) => NNFConst("CJ",hcnnf_wp p,hcnnf_wp q) 
+    | (Const ("op |",_) $ p $ q) =>  NNFConst("DJ",hcnnf_wp p,hcnnf_wp q)
+    | (Const ("op -->",_) $ p $q) => NNFConst("IM",hcnnf_wp (HOLogic.Not $ p),hcnnf_wp q)
+    | (Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q) => NNFConst("EQ",hcnnf_wp (HOLogic.mk_conj(p,q)),hcnnf_wp (HOLogic.mk_conj((HOLogic.Not $ p), (HOLogic.Not $ q)))) 
+
+    | (Const ("Not",_) $ (Const("Not",_) $ p)) => NNFNN(hcnnf_wp p) 
+    | (Const ("Not",_) $ (Const ("op &",_) $ p $ q)) => NNFConst ("NCJ",(hcnnf_wp(HOLogic.Not $ p)),(hcnnf_wp(HOLogic.Not $ q))) 
+    | (Const ("Not",_) $(Const ("op |",_) $ (A as (Const ("op &",_) $ p $ q)) $  
+    			(B as (Const ("op &",_) $ p1 $ r)))) => if p1 = negate p then 
+		         NNFConst("SDJ",  
+			   NNFConst("CJ",hcnnf_wp p,hcnnf_wp(HOLogic.Not $ q)),
+			   NNFConst("CJ",hcnnf_wp p1,hcnnf_wp(HOLogic.Not $ r)))
+			 else  NNFConst ("NDJ",(hcnnf_wp(HOLogic.Not $ A)),(hcnnf_wp(HOLogic.Not $ B))) 
+
+    | (Const ("Not",_) $ (Const ("op |",_) $ p $ q)) => NNFConst ("NDJ",(hcnnf_wp(HOLogic.Not $ p)),(hcnnf_wp(HOLogic.Not $ q))) 
+    | (Const ("Not",_) $ (Const ("op -->",_) $ p $q)) =>  NNFConst ("NIM",(hcnnf_wp(p)),(hcnnf_wp(HOLogic.Not $ q))) 
+    | (Const ("Not",_) $ (Const ("op =",Type ("fun",[Type ("bool", []),_]))  $ p $ q)) =>NNFConst ("NEQ",(NNFConst("CJ",hcnnf_wp p,hcnnf_wp(HOLogic.Not $ q))),(NNFConst("CJ",hcnnf_wp(HOLogic.Not $ p),hcnnf_wp q))) 
+    | _ => NNFAt(fm)  
+  in NNFSimp(hcnnf_wp f)
+end; 
+   
+
+
+
+
+
+(* ------------------------------------------------------------------------- *)
+(*Cooper decision Procedure with proof protocoling*)
+(* ------------------------------------------------------------------------- *)
+
+fun coopermi_wp vars fm =
+  case fm of
+   Const ("Ex",_) $ Abs(xo,T,po) => let 
+    val (xn,np) = variant_abs(xo,T,po) 
+    val x = (Free(xn , T))
+    val p = np     (* Is this a legal proof for the P=NP Problem??*)
+    val (p_inf,miprt) = simpl_wp (minusinf_wp x p)
+    val (bset,bsprt) = bset_wp x p
+    val nbst_p_prt = not_bst_p_wp x p
+    val dlcm = divlcm x p 
+    val js = 1 upto dlcm 
+    fun p_element j b = linrep vars x (linear_add vars b (mk_numeral j)) p 
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) bset) 
+   in (list_disj (map stage js),Cooper(mk_numeral dlcm,miprt,bsprt,nbst_p_prt))
+   end
+   
+  | _ => (error "cooper: not an existential formula",No);
+				
+fun cooperpi_wp vars fm =
+  case fm of
+   Const ("Ex",_) $ Abs(xo,T,po) => let 
+    val (xn,np) = variant_abs(xo,T,po) 
+    val x = (Free(xn , T))
+    val p = np     (* Is this a legal proof for the P=NP Problem??*)
+    val (p_inf,piprt) = simpl_wp (plusinf_wp x p)
+    val (aset,asprt) = aset_wp x p
+    val nast_p_prt = not_ast_p_wp x p
+    val dlcm = divlcm x p 
+    val js = 1 upto dlcm 
+    fun p_element j a = linrep vars x (linear_sub vars a (mk_numeral j)) p 
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) aset) 
+   in (list_disj (map stage js),Cooper(mk_numeral dlcm,piprt,asprt,nast_p_prt))
+   end
+  | _ => (error "cooper: not an existential formula",No);
+				
+
+
+
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---      Interpretation and Proofgeneration Part              ---*)
+(*---                                                           ---*)
+(*---      Protocole interpretation functions                   ---*)
+(*---                                                           ---*)
+(*---      and proofgeneration functions                        ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+(* ------------------------------------------------------------------------- *)
+(* Returns both sides of an equvalence in the theorem*)
+(* ------------------------------------------------------------------------- *)
+fun qe_get_terms th = let val (_$(Const("op =",Type ("fun",[Type ("bool", []),_])) $ A $ B )) = prop_of th in (A,B) end;
+
+
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+
+(* ------------------------------------------------------------------------- *)
+(* Modified version of the simple version with minimal amount of checking and postprocessing*)
+(* ------------------------------------------------------------------------- *)
+
+fun simple_prove_goal_cterm2 G tacs =
+  let
+    fun check None = error "prove_goal: tactic failed"
+      | check (Some (thm, _)) = (case nprems_of thm of
+            0 => thm
+          | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
+  in check (Seq.pull (EVERY tacs (trivial G))) end;
+
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+
+fun cert_Trueprop sg t = cterm_of sg (HOLogic.mk_Trueprop t);
+
+(* ------------------------------------------------------------------------- *)
+(*This function proove elementar will be used to generate proofs at runtime*)
+(*It is is based on the isabelle function proove_goalw_cterm and is thought to *)
+(*prove properties such as a dvd b (essentially) that are only to make at
+runtime.*)
+(* ------------------------------------------------------------------------- *)
+fun prove_elementar sg s fm2 = case s of 
+  (*"ss" like simplification with simpset*)
+  "ss" =>
+    let
+      val ss = presburger_ss addsimps
+        [zdvd_iff_zmod_eq_0,unity_coeff_ex]
+      val ct =  cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]
+    end
+
+  (*"bl" like blast tactic*)
+  (* Is only used in the harrisons like proof procedure *)
+  | "bl" =>
+     let val ct = cert_Trueprop sg fm2
+     in
+       simple_prove_goal_cterm2 ct [blast_tac HOL_cs 1]
+     end
+
+  (*"ed" like Existence disjunctions ...*)
+  (* Is only used in the harrisons like proof procedure *)
+  | "ed" =>
+    let
+      val ex_disj_tacs =
+        let
+          val tac1 = EVERY[REPEAT(resolve_tac [disjI1,disjI2] 1), etac exI 1]
+          val tac2 = EVERY[etac exE 1, rtac exI 1,
+            REPEAT(resolve_tac [disjI1,disjI2] 1), assumption 1]
+	in [rtac iffI 1,
+          etac exE 1, REPEAT(EVERY[etac disjE 1, tac1]), tac1,
+          REPEAT(EVERY[etac disjE 1, tac2]), tac2]
+        end
+
+      val ct = cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct ex_disj_tacs
+    end
+
+  | "fa" =>
+    let val ct = cert_Trueprop sg fm2
+    in simple_prove_goal_cterm2 ct [simple_arith_tac 1]
+    end
+
+  | "sa" =>
+    let
+      val ss = presburger_ss addsimps zadd_ac
+      val ct = cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]
+    end
+
+  | "ac" =>
+    let
+      val ss = HOL_basic_ss addsimps zadd_ac
+      val ct = cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct [simp_tac ss 1]
+    end
+
+  | "lf" =>
+    let
+      val ss = presburger_ss addsimps zadd_ac
+      val ct = cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]
+    end;
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function return an Isabelle proof, of the adjustcoffeq result.*)
+(* The proofs are in Presburger.thy and are generally based on the arithmetic *)
+(* ------------------------------------------------------------------------- *)
+fun proof_of_adjustcoeffeq sg (prt,rs) = case prt of
+   ACfm fm => instantiate' [Some cboolT]
+    [Some (cterm_of sg fm)] refl
+ | ACAt (k,at as (Const(p,_) $a $( Const ("op +", _)$(Const ("op *",_) $ 
+      c $ x ) $t ))) => 
+   let
+     val ck = cterm_of sg (mk_numeral k)
+     val cc = cterm_of sg c
+     val ct = cterm_of sg t
+     val cx = cterm_of sg x
+     val ca = cterm_of sg a
+   in case p of
+     "op <" => let val pre = prove_elementar sg "ss" 
+	                  (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral k)))
+	           val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_lt_eq)))
+		      in [th1,(prove_elementar sg "ss" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans
+                   end
+    |"op =" =>let val pre = prove_elementar sg "ss" 
+	    (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
+	          in let val th1 = (pre RS(instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_eq_eq)))
+	             in [th1,(prove_elementar sg "ss" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans
+                      end
+                  end
+    |"Divides.op dvd" =>let val pre = prove_elementar sg "ss" 
+	   (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
+	                 val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct]) (ac_dvd_eq))
+                         in [th1,(prove_elementar sg "ss" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans
+                        
+                          end
+  end
+ |ACPI(k,at as (Const("Not",_)$(Const("op <",_) $a $( Const ("op +", _)$(Const ("op *",_) $ c $ x ) $t )))) => 
+   let
+     val ck = cterm_of sg (mk_numeral k)
+     val cc = cterm_of sg c
+     val ct = cterm_of sg t
+     val cx = cterm_of sg x
+     val pre = prove_elementar sg "ss" 
+       (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral k)))
+       val th1 = (pre RS (instantiate' [] [Some ck,Some cc, Some cx, Some ct] (ac_pi_eq)))
+
+         in [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans
+   end
+ |ACNeg(pr) => let val (Const("Not",_)$nrs) = rs
+               in (proof_of_adjustcoeffeq sg (pr,nrs)) RS (qe_Not) 
+               end
+ |ACConst(s,pr1,pr2) =>
+   let val (Const(_,_)$rs1$rs2) = rs
+       val th1 = proof_of_adjustcoeffeq sg (pr1,rs1)
+       val th2 = proof_of_adjustcoeffeq sg (pr2,rs2)
+       in case s of 
+	 "CJ" => [th1,th2] MRS (qe_conjI)
+         |"DJ" => [th1,th2] MRS (qe_disjI)
+         |"IM" => [th1,th2] MRS (qe_impI)
+         |"EQ" => [th1,th2] MRS (qe_eqI)
+   end;
+
+
+
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function return an Isabelle proof, of some properties on the atoms*)
+(* The proofs are in Presburger.thy and are generally based on the arithmetic *)
+(* This function doese only instantiate the the theorems in the theory *)
+(* ------------------------------------------------------------------------- *)
+fun atomar_minf_proof_of sg dlcm (Modd_minf (x,fm1)) =
+  let
+    (*Some certified Terms*)
+    
+   val ctrue = cterm_of sg HOLogic.true_const
+   val cfalse = cterm_of sg HOLogic.false_const
+   val fm = norm_zero_one fm1
+  in  case fm1 of 
+      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+         if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
+           else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+
+      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	   if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one) 
+	   then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf))
+	 	 else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf)) 
+
+      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+           if (y=x) andalso (c1 = zero) then 
+            if (pm1 = one) then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf)) else
+	     (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
+	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+  
+      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
+			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
+	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf)))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+      |(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
+      c $ y ) $ z))) => 
+         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
+			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
+	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf)))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+		
+    
+   |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf)
+   end	
+
+ |atomar_minf_proof_of sg dlcm (Eq_minf (x,fm1)) =  let
+       (*Some certified types*)
+   val fm = norm_zero_one fm1
+    in  case fm1 of 
+      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+         if  (x=y) andalso (c1=zero) andalso (c2=one) 
+	   then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_minf))
+           else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+
+      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	   if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
+	     then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_minf))
+	     else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf)) 
+
+      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+           if (y=x) andalso (c1 =zero) then 
+            if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else
+	     (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_minf))
+	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
+	 		  val cz = cterm_of sg (norm_zero_one z)
+	 	      in(instantiate' [] [Some cd,  Some cz] (not_dvd_eq_minf)) 
+		      end
+
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+		
+      |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
+	 		  val cz = cterm_of sg (norm_zero_one z)
+	 	      in(instantiate' [] [Some cd, Some cz ] (dvd_eq_minf))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+
+      		
+    |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ end;
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function combines proofs of some special form already synthetised from the subtrees to make*)
+(* a new proof of the same form. The combination occures whith isabelle theorems which have been already prooved *)
+(*these Theorems are in Presburger.thy and mostly do not relay on the arithmetic.*)
+(* These are Theorems for the Property of P_{-infty}*)
+(* ------------------------------------------------------------------------- *)
+fun combine_minf_proof s pr1 pr2 = case s of
+    "ECJ" => [pr1 , pr2] MRS (eq_minf_conjI)
+
+   |"EDJ" => [pr1 , pr2] MRS (eq_minf_disjI)
+   
+   |"MCJ" => [pr1 , pr2] MRS (modd_minf_conjI)
+
+   |"MDJ" => [pr1 , pr2] MRS (modd_minf_disjI);
+
+(* ------------------------------------------------------------------------- *)
+(*This function return an isabelle Proof for the minusinfinity theorem*)
+(* It interpretates the protool and gives the protokoles property of P_{...} as a theorem*)
+(* ------------------------------------------------------------------------- *)
+fun minf_proof_ofh sg dlcm prl = case prl of 
+
+    Eq_minf (_) => atomar_minf_proof_of sg dlcm prl
+    
+   |Modd_minf (_) => atomar_minf_proof_of sg dlcm prl
+   
+   |Eq_minf_conjI (prl1,prl2) => let val pr1 = minf_proof_ofh sg dlcm prl1
+   				    val pr2 = minf_proof_ofh sg dlcm prl2
+				 in (combine_minf_proof "ECJ" pr1 pr2)
+				 end
+				 
+   |Eq_minf_disjI (prl1,prl2) => let val pr1 = minf_proof_ofh sg dlcm prl1
+   				    val pr2 = minf_proof_ofh sg dlcm prl2
+				 in (combine_minf_proof "EDJ" pr1 pr2)
+				 end
+				 
+   |Modd_minf_conjI (prl1,prl2) => let val pr1 = minf_proof_ofh sg dlcm prl1
+   				    val pr2 = minf_proof_ofh sg dlcm prl2
+				 in (combine_minf_proof "MCJ" pr1 pr2)
+				 end
+				 
+   |Modd_minf_disjI (prl1,prl2) => let val pr1 = minf_proof_ofh sg dlcm prl1
+   				    val pr2 = minf_proof_ofh sg dlcm prl2
+				 in (combine_minf_proof "MDJ" pr1 pr2)
+				 end;
+(* ------------------------------------------------------------------------- *)
+(* Main function For the rest both properies of P_{..} are needed and here both theorems are returned.*)				 
+(* ------------------------------------------------------------------------- *)
+fun  minf_proof_of sg dlcm (Minusinf (prl1,prl2))  = 
+  let val pr1 = minf_proof_ofh sg dlcm prl1
+      val pr2 = minf_proof_ofh sg dlcm prl2
+  in (pr1, pr2)
+end;
+				 
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function return an Isabelle proof, of some properties on the atoms*)
+(* The proofs are in Presburger.thy and are generally based on the arithmetic *)
+(* This function doese only instantiate the the theorems in the theory *)
+(* ------------------------------------------------------------------------- *)
+fun atomar_pinf_proof_of sg dlcm (Modd_minf (x,fm1)) =
+ let
+    (*Some certified Terms*)
+    
+  val ctrue = cterm_of sg HOLogic.true_const
+  val cfalse = cterm_of sg HOLogic.false_const
+  val fm = norm_zero_one fm1
+ in  case fm1 of 
+      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+         if ((x=y) andalso (c1= zero) andalso (c2= one))
+	 then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf))
+         else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+
+      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	if ((is_arith_rel fm) andalso (x = y) andalso (c1 = zero)  andalso (c2 = one)) 
+	then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
+	else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+
+      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+        if ((y=x) andalso (c1 = zero)) then 
+          if (pm1 = one) 
+	  then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf)) 
+	  else (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
+	else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+  
+      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
+			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
+	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf)))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+      |(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
+      c $ y ) $ z))) => 
+         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
+			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
+	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf)))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+		
+    
+   |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf)
+   end	
+
+ |atomar_pinf_proof_of sg dlcm (Eq_minf (x,fm1)) =  let
+					val fm = norm_zero_one fm1
+    in  case fm1 of 
+      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+         if  (x=y) andalso (c1=zero) andalso (c2=one) 
+	   then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_pinf))
+           else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+
+      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	   if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
+	     then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_pinf))
+	     else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf)) 
+
+      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+           if (y=x) andalso (c1 =zero) then 
+            if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else
+	     (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_pinf))
+	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
+	 		  val cz = cterm_of sg (norm_zero_one z)
+	 	      in(instantiate' [] [Some cd,  Some cz] (not_dvd_eq_pinf)) 
+		      end
+
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+		
+      |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
+	 		  val cz = cterm_of sg (norm_zero_one z)
+	 	      in(instantiate' [] [Some cd, Some cz ] (dvd_eq_pinf))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+
+      		
+    |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ end;
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function combines proofs of some special form already synthetised from the subtrees to make*)
+(* a new proof of the same form. The combination occures whith isabelle theorems which have been already prooved *)
+(*these Theorems are in Presburger.thy and mostly do not relay on the arithmetic.*)
+(* These are Theorems for the Property of P_{+infty}*)
+(* ------------------------------------------------------------------------- *)
+fun combine_pinf_proof s pr1 pr2 = case s of
+    "ECJ" => [pr1 , pr2] MRS (eq_pinf_conjI)
+
+   |"EDJ" => [pr1 , pr2] MRS (eq_pinf_disjI)
+   
+   |"MCJ" => [pr1 , pr2] MRS (modd_pinf_conjI)
+
+   |"MDJ" => [pr1 , pr2] MRS (modd_pinf_disjI);
+
+(* ------------------------------------------------------------------------- *)
+(*This function return an isabelle Proof for the minusinfinity theorem*)
+(* It interpretates the protool and gives the protokoles property of P_{...} as a theorem*)
+(* ------------------------------------------------------------------------- *)
+fun pinf_proof_ofh sg dlcm prl = case prl of 
+
+    Eq_minf (_) => atomar_pinf_proof_of sg dlcm prl
+    
+   |Modd_minf (_) => atomar_pinf_proof_of sg dlcm prl
+   
+   |Eq_minf_conjI (prl1,prl2) => let val pr1 = pinf_proof_ofh sg dlcm prl1
+   				    val pr2 = pinf_proof_ofh sg dlcm prl2
+				 in (combine_pinf_proof "ECJ" pr1 pr2)
+				 end
+				 
+   |Eq_minf_disjI (prl1,prl2) => let val pr1 = pinf_proof_ofh sg dlcm prl1
+   				    val pr2 = pinf_proof_ofh sg dlcm prl2
+				 in (combine_pinf_proof "EDJ" pr1 pr2)
+				 end
+				 
+   |Modd_minf_conjI (prl1,prl2) => let val pr1 = pinf_proof_ofh sg dlcm prl1
+   				    val pr2 = pinf_proof_ofh sg dlcm prl2
+				 in (combine_pinf_proof "MCJ" pr1 pr2)
+				 end
+				 
+   |Modd_minf_disjI (prl1,prl2) => let val pr1 = pinf_proof_ofh sg dlcm prl1
+   				    val pr2 = pinf_proof_ofh sg dlcm prl2
+				 in (combine_pinf_proof "MDJ" pr1 pr2)
+				 end;
+(* ------------------------------------------------------------------------- *)
+(* Main function For the rest both properies of P_{..} are needed and here both theorems are returned.*)				 
+(* ------------------------------------------------------------------------- *)
+fun pinf_proof_of sg dlcm (Minusinf (prl1,prl2))  = 
+  let val pr1 = pinf_proof_ofh sg dlcm prl1
+      val pr2 = pinf_proof_ofh sg dlcm prl2
+  in (pr1, pr2)
+end;
+				 
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* Here we generate the theorem for the Bset Property in the simple direction*)
+(* It is just an instantiation*)
+(* ------------------------------------------------------------------------- *)
+fun bsetproof_of sg (Bset(x as Free(xn,xT),fm,bs,dlcm))   = 
+  let
+    val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
+    val cdlcm = cterm_of sg dlcm
+    val cB = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one bs))
+  in instantiate' [] [Some cdlcm,Some cB, Some cp] (bst_thm)
+    end;
+
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* Here we generate the theorem for the Bset Property in the simple direction*)
+(* It is just an instantiation*)
+(* ------------------------------------------------------------------------- *)
+fun asetproof_of sg (Aset(x as Free(xn,xT),fm,ast,dlcm))   = 
+  let
+    val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
+    val cdlcm = cterm_of sg dlcm
+    val cA = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one ast))
+  in instantiate' [] [Some cdlcm,Some cA, Some cp] (ast_thm)
+end;
+
+
+
+
+(* ------------------------------------------------------------------------- *)    
+(* Protokol interpretation function for the backwards direction for cooper's Theorem*)
+
+(* For the generation of atomic Theorems*)
+(* Prove the premisses on runtime and then make RS*)
+(* ------------------------------------------------------------------------- *)
+fun generate_atomic_not_bst_p sg (x as Free(xn,xT)) fm dlcm B at = 
+  let
+    val cdlcm = cterm_of sg dlcm
+    val cB = cterm_of sg B
+    val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
+    val cat = cterm_of sg (norm_zero_one at)
+  in
+  case at of 
+   (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+      if  (x=y) andalso (c1=zero) andalso (c2=one) 
+	 then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
+	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
+		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	 in  (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_bst_p_ne)))
+	 end
+         else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+
+   |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
+     if (is_arith_rel at) andalso (x=y)
+	then let val bst_z = norm_zero_one (linear_neg (linear_add [] z (mk_numeral 1)))
+	         in let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ bst_z $ B)
+	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq (bst_z,Const("op -",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
+		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	 in  (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_bst_p_eq)))
+	 end
+       end
+         else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+
+   |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+        if (y=x) andalso (c1 =zero) then 
+        if pm1 = one then 
+	  let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
+              val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
+	  in  (instantiate' [] [Some cfma,  Some cdlcm]([th1,th2] MRS (not_bst_p_gt)))
+	    end
+	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	      in (instantiate' [] [Some cfma, Some cB,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt)))
+	      end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+
+   |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+      if y=x then  
+           let val cz = cterm_of sg (norm_zero_one z)
+	       val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
+ 	     in (instantiate' []  [Some cfma, Some cB,Some cz] (th1 RS (not_bst_p_ndvd)))
+	     end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+
+   |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+       if y=x then  
+	 let val cz = cterm_of sg (norm_zero_one z)
+	     val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
+ 	    in (instantiate' []  [Some cfma,Some cB,Some cz] (th1 RS (not_bst_p_dvd)))
+	  end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+      		
+   |_ => (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+      		
+    end;
+    
+(* ------------------------------------------------------------------------- *)    
+(* Main interpretation function for this backwards dirction*)
+(* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*)
+(*Help Function*)
+(* ------------------------------------------------------------------------- *)
+fun not_bst_p_proof_of_h sg x fm dlcm B prt = case prt of 
+	(Not_bst_p_atomic(fm2)) => (generate_atomic_not_bst_p sg x fm dlcm B fm2)
+	
+	|(Not_bst_p_conjI(pr1,pr2)) => 
+			let val th1 = (not_bst_p_proof_of_h sg x fm dlcm B pr1)
+			    val th2 = (not_bst_p_proof_of_h sg x fm dlcm B pr2)
+			    in ([th1,th2] MRS (not_bst_p_conjI))
+			    end
+
+	|(Not_bst_p_disjI(pr1,pr2)) => 
+			let val th1 = (not_bst_p_proof_of_h sg x fm dlcm B pr1)
+			    val th2 = (not_bst_p_proof_of_h sg x fm dlcm B pr2)
+			    in ([th1,th2] MRS not_bst_p_disjI)
+			    end;
+(* Main function*)
+fun not_bst_p_proof_of sg (Not_bst_p(x as Free(xn,xT),fm,dlcm,B,prl)) =
+  let val th =  not_bst_p_proof_of_h sg x fm dlcm B prl
+      val fma = absfree (xn,xT, norm_zero_one fm)
+  in let val th1 =  prove_elementar sg "ss"  (HOLogic.mk_eq (fma,fma))
+     in [th,th1] MRS (not_bst_p_Q_elim)
+     end
+  end;
+
+
+(* ------------------------------------------------------------------------- *)    
+(* Protokol interpretation function for the backwards direction for cooper's Theorem*)
+
+(* For the generation of atomic Theorems*)
+(* Prove the premisses on runtime and then make RS*)
+(* ------------------------------------------------------------------------- *)
+fun generate_atomic_not_ast_p sg (x as Free(xn,xT)) fm dlcm A at = 
+  let
+    val cdlcm = cterm_of sg dlcm
+    val cA = cterm_of sg A
+    val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
+    val cat = cterm_of sg (norm_zero_one at)
+  in
+  case at of 
+   (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+      if  (x=y) andalso (c1=zero) andalso (c2=one) 
+	 then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ A)
+	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
+		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	 in  (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_ast_p_ne)))
+	 end
+         else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+
+   |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
+     if (is_arith_rel at) andalso (x=y)
+	then let val ast_z = norm_zero_one (linear_sub [] one z )
+	         val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ ast_z $ A)
+	         val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq (ast_z,Const("op +",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
+		 val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	 in  (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_ast_p_eq)))
+       end
+         else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+
+   |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+        if (y=x) andalso (c1 =zero) then 
+        if pm1 = (mk_numeral ~1) then 
+	  let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one z) $ A)
+              val th2 =  prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm))
+	  in  (instantiate' [] [Some cfma]([th2,th1] MRS (not_ast_p_lt)))
+	    end
+	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	      in (instantiate' [] [Some cfma, Some cA,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt)))
+	      end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+
+   |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+      if y=x then  
+           let val cz = cterm_of sg (norm_zero_one z)
+	       val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
+ 	     in (instantiate' []  [Some cfma, Some cA,Some cz] (th1 RS (not_ast_p_ndvd)))
+	     end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+
+   |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+       if y=x then  
+	 let val cz = cterm_of sg (norm_zero_one z)
+	     val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
+ 	    in (instantiate' []  [Some cfma,Some cA,Some cz] (th1 RS (not_ast_p_dvd)))
+	  end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+      		
+   |_ => (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+      		
+    end;
+    
+(* ------------------------------------------------------------------------- *)    
+(* Main interpretation function for this backwards dirction*)
+(* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*)
+(*Help Function*)
+(* ------------------------------------------------------------------------- *)
+fun not_ast_p_proof_of_h sg x fm dlcm A prt = case prt of 
+	(Not_ast_p_atomic(fm2)) => (generate_atomic_not_ast_p sg x fm dlcm A fm2)
+	
+	|(Not_ast_p_conjI(pr1,pr2)) => 
+			let val th1 = (not_ast_p_proof_of_h sg x fm dlcm A pr1)
+			    val th2 = (not_ast_p_proof_of_h sg x fm dlcm A pr2)
+			    in ([th1,th2] MRS (not_ast_p_conjI))
+			    end
+
+	|(Not_ast_p_disjI(pr1,pr2)) => 
+			let val th1 = (not_ast_p_proof_of_h sg x fm dlcm A pr1)
+			    val th2 = (not_ast_p_proof_of_h sg x fm dlcm A pr2)
+			    in ([th1,th2] MRS (not_ast_p_disjI))
+			    end;
+(* Main function*)
+fun not_ast_p_proof_of sg (Not_ast_p(x as Free(xn,xT),fm,dlcm,A,prl)) =
+  let val th =  not_ast_p_proof_of_h sg x fm dlcm A prl
+      val fma = absfree (xn,xT, norm_zero_one fm)
+      val th1 =  prove_elementar sg "ss"  (HOLogic.mk_eq (fma,fma))
+  in [th,th1] MRS (not_ast_p_Q_elim)
+end;
+
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* Interpretaion of Protocols of the cooper procedure : minusinfinity version*)
+(* ------------------------------------------------------------------------- *)
+
+
+fun coopermi_proof_of sg x (Cooper (dlcm,Simp(fm,miprt),bsprt,nbst_p_prt)) =
+  (* Get the Bset thm*)
+  let val bst = bsetproof_of sg bsprt
+      val (mit1,mit2) = minf_proof_of sg dlcm miprt
+      val fm1 = norm_zero_one (simpl fm) 
+      val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm));
+      val nbstpthm = not_bst_p_proof_of sg nbst_p_prt
+    (* Return the four theorems needed to proove the whole Cooper Theorem*)
+  in (dpos,mit2,bst,nbstpthm,mit1)
+end;
+
+
+(* ------------------------------------------------------------------------- *)
+(* Interpretaion of Protocols of the cooper procedure : plusinfinity version *)
+(* ------------------------------------------------------------------------- *)
+
+
+fun cooperpi_proof_of sg x (Cooper (dlcm,Simp(fm,miprt),bsprt,nast_p_prt)) =
+  let val ast = asetproof_of sg bsprt
+      val (mit1,mit2) = pinf_proof_of sg dlcm miprt
+      val fm1 = norm_zero_one (simpl fm) 
+      val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm));
+      val nastpthm = not_ast_p_proof_of sg nast_p_prt
+  in (dpos,mit2,ast,nastpthm,mit1)
+end;
+
+
+(* ------------------------------------------------------------------------- *)
+(* Interpretaion of Protocols of the cooper procedure : full version*)
+(* ------------------------------------------------------------------------- *)
+
+
+
+fun cooper_thm sg s (x as Free(xn,xT)) vars cfm = case s of
+  "pi" => let val (rs,prt) = cooperpi_wp (xn::vars) (HOLogic.mk_exists(xn,xT,cfm))
+	      val (dpsthm,th1,th2,nbpth,th3) = cooperpi_proof_of sg x prt
+		   in [dpsthm,th1,th2,nbpth,th3] MRS (cppi_eq)
+           end
+  |"mi" => let val (rs,prt) = coopermi_wp (xn::vars) (HOLogic.mk_exists(xn,xT,cfm))
+	       val (dpsthm,th1,th2,nbpth,th3) = coopermi_proof_of sg x prt
+		   in [dpsthm,th1,th2,nbpth,th3] MRS (cpmi_eq)
+                end
+ |_ => error "parameter error";
+
+(* ------------------------------------------------------------------------- *)
+(* This function should evoluate to the end prove Procedure for one quantifier elimination for Presburger arithmetic*)
+(* It shoud be plugged in the qfnp argument of the quantifier elimination proof function*)
+(* ------------------------------------------------------------------------- *)
+
+fun cooper_prv sg (x as Free(xn,xT)) efm vars = let 
+   val l = formlcm x efm
+   val ac_thm = proof_of_adjustcoeffeq sg (adjustcoeffeq_wp  x l efm)
+   val fm = snd (qe_get_terms ac_thm)
+   val  cfm = unitycoeff x fm
+   val afm = adjustcoeff x l fm
+   val P = absfree(xn,xT,afm)
+   val ss = presburger_ss addsimps
+     [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
+   val uth = instantiate' [] [Some (cterm_of sg P) , Some (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
+   val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
+   val cms = if ((length (aset x cfm)) < (length (bset x cfm))) then "pi" else "mi"
+   val cp_thm = cooper_thm sg cms x vars cfm
+   val exp_cp_thm = refl RS (simplify ss (cp_thm RSN (2,trans)))
+   val (lsuth,rsuth) = qe_get_terms (uth)
+   val (lseacth,rseacth) = qe_get_terms(e_ac_thm)
+   val (lscth,rscth) = qe_get_terms (exp_cp_thm)
+   val  u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans
+ in  ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans)
+   end
+|cooper_prv _ _ _ _ = error "Parameters format";
+
+
+(*====================================================*)
+(*Interpretation function for the evaluation protokol *)
+(*====================================================*)
+
+fun proof_of_evalc sg fm =
+let
+fun proof_of_evalch prt = case prt of
+  EvalAt(at) => prove_elementar sg "ss" at
+ |Evalfm(fm) => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl
+ |EvalConst(s,pr1,pr2) => 
+   let val th1 = proof_of_evalch pr1
+       val th2 = proof_of_evalch pr2
+   in case s of
+     "CJ" =>[th1,th2] MRS (qe_conjI)
+    |"DJ" =>[th1,th2] MRS (qe_disjI)
+    |"IM" =>[th1,th2] MRS (qe_impI)
+    |"EQ" =>[th1,th2] MRS (qe_eqI)
+    end
+in proof_of_evalch (evalc_wp fm)
+end;
+
+(*============================================================*)
+(*Interpretation function for the NNF-Transformation protokol *)
+(*============================================================*)
+
+fun proof_of_cnnf sg fm pf = 
+let fun proof_of_cnnfh prt pat = case prt of
+  NNFAt(at) => pat at
+ |NNFSimp (pr) => let val th1 = proof_of_cnnfh pr pat
+                  in let val fm2 = snd (qe_get_terms th1) 
+		     in [th1,prove_elementar sg "ss" (HOLogic.mk_eq(fm2 ,simpl fm2))] MRS trans
+                     end
+                  end
+ |NNFNN (pr) => (proof_of_cnnfh pr pat) RS (nnf_nn)
+ |NNFConst (s,pr1,pr2) =>
+   let val th1 = proof_of_cnnfh pr1 pat
+       val th2 = proof_of_cnnfh pr2 pat
+   in case s of
+     "CJ" => [th1,th2] MRS (qe_conjI)
+    |"DJ" => [th1,th2] MRS (qe_disjI)
+    |"IM" => [th1,th2] MRS (nnf_im)
+    |"EQ" => [th1,th2] MRS (nnf_eq)
+    |"SDJ" => let val (Const("op &",_)$A$_) = fst (qe_get_terms th1)
+	          val (Const("op &",_)$C$_) = fst (qe_get_terms th2)
+	      in [th1,th2,prove_elementar sg "ss" (HOLogic.mk_eq (A,HOLogic.Not $ C))] MRS (nnf_sdj)
+	      end
+    |"NCJ" => [th1,th2] MRS (nnf_ncj)
+    |"NIM" => [th1,th2] MRS (nnf_nim)
+    |"NEQ" => [th1,th2] MRS (nnf_neq)
+    |"NDJ" => [th1,th2] MRS (nnf_ndj)
+   end
+in proof_of_cnnfh (cnnf_wp fm) pf
+end;
+
+
+
+
+(*====================================================*)
+(* Interpretation function for the linform protokol   *)
+(*====================================================*)
+
+
+fun proof_of_linform sg vars f = 
+  let fun proof_of_linformh prt = 
+  case prt of
+    (LfAt (at)) =>  prove_elementar sg "lf" (HOLogic.mk_eq (at, linform vars at))
+   |(LfAtdvd (Const("Divides.op dvd",_)$d$t)) => (prove_elementar sg "lf" (HOLogic.mk_eq (t, lint vars t))) RS (instantiate' [] [None , None, Some (cterm_of sg d)](linearize_dvd))
+   |(Lffm (fm)) => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl)
+   |(LfConst (s,pr1,pr2)) =>
+     let val th1 = proof_of_linformh pr1
+	 val th2 = proof_of_linformh pr2
+     in case s of
+       "CJ" => [th1,th2] MRS (qe_conjI)
+      |"DJ" =>[th1,th2] MRS (qe_disjI)
+      |"IM" =>[th1,th2] MRS (qe_impI)
+      |"EQ" =>[th1,th2] MRS (qe_eqI)
+     end
+   |(LfNot(pr)) => 
+     let val th = proof_of_linformh pr
+     in (th RS (qe_Not))
+     end
+   |(LfQ(s,xn,xT,pr)) => 
+     let val th = forall_intr (cterm_of sg (Free(xn,xT)))(proof_of_linformh pr)
+     in if s = "Ex" 
+        then (th COMP(qe_exI) )
+        else (th COMP(qe_ALLI) )
+     end
+in
+ proof_of_linformh (linform_wp f)
+end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Integ/presburger.ML	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,231 @@
+(*  Title:      HOL/Integ/presburger.ML
+    ID:         $Id$
+    Author:     Amine Chaieb and Stefan Berghofer, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+Tactic for solving arithmetical Goals in Presburger Arithmetic
+*)
+
+signature PRESBURGER = 
+sig
+ val presburger_tac : bool -> int -> tactic
+ val presburger_method : bool -> int -> Proof.method
+ val setup : (theory -> theory) list
+ val trace : bool ref
+end;
+
+structure Presburger: PRESBURGER =
+struct
+
+val trace = ref false;
+fun trace_msg s = if !trace then tracing s else ();
+
+(*-----------------------------------------------------------------*)
+(*cooper_pp: provefunction for the one-exstance quantifier elimination*)
+(* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*)
+(*-----------------------------------------------------------------*)
+
+val presburger_ss = simpset_of (theory "Presburger");
+
+fun cooper_pp sg vrl (fm as e$Abs(xn,xT,p)) = 
+  let val (xn1,p1) = variant_abs (xn,xT,p)
+  in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1 vrl) end;
+
+fun mnnf_pp sg fm = CooperProof.proof_of_cnnf sg fm
+  (CooperProof.proof_of_evalc sg);
+
+fun mproof_of_int_qelim sg fm =
+  Qelim.proof_of_mlift_qelim sg CooperDec.is_arith_rel
+    (CooperProof.proof_of_linform sg) (mnnf_pp sg) (cooper_pp sg) fm;
+
+(* Theorems to be used in this tactic*)
+
+val zdvd_int = thm "zdvd_int";
+val zdiff_int_split = thm "zdiff_int_split";
+val all_nat = thm "all_nat";
+val ex_nat = thm "ex_nat";
+val number_of1 = thm "number_of1";
+val number_of2 = thm "number_of2";
+val split_zdiv = thm "split_zdiv";
+val split_zmod = thm "split_zmod";
+val mod_div_equality' = thm "mod_div_equality'";
+val split_div' = thm "split_div'";
+val Suc_plus1 = thm "Suc_plus1";
+val imp_le_cong = thm "imp_le_cong";
+val conj_le_cong = thm "conj_le_cong";
+
+(* extract all the constants in a term*)
+fun add_term_typed_consts (Const (c, T), cs) = (c,T) ins cs
+  | add_term_typed_consts (t $ u, cs) =
+      add_term_typed_consts (t, add_term_typed_consts (u, cs))
+  | add_term_typed_consts (Abs (_, _, t), cs) = add_term_typed_consts (t, cs)
+  | add_term_typed_consts (_, cs) = cs;
+
+fun term_typed_consts t = add_term_typed_consts(t,[]);
+
+(* put a term into eta long beta normal form *)
+fun eta_long Ts (Abs (s, T, t)) = Abs (s, T, eta_long (T :: Ts) t)
+  | eta_long Ts t = (case strip_comb t of
+      (Abs _, _) => eta_long Ts (Envir.beta_norm t)
+    | (u, ts) => 
+      let val Us = binder_types (fastype_of1 (Ts, t))
+      in list_abs (map (pair "x") Us, Unify.combound
+        (list_comb (u, map (eta_long Ts) ts), 0, length Us))
+      end);
+
+(* Some Types*)
+val bT = HOLogic.boolT;
+val iT = HOLogic.intT;
+val binT = HOLogic.binT;
+val nT = HOLogic.natT;
+
+(* Allowed Consts in formulae for presburger tactic*)
+
+val allowed_consts =
+  [("All", (iT --> bT) --> bT),
+   ("Ex", (iT --> bT) --> bT),
+   ("All", (nT --> bT) --> bT),
+   ("Ex", (nT --> bT) --> bT),
+
+   ("op &", bT --> bT --> bT),
+   ("op |", bT --> bT --> bT),
+   ("op -->", bT --> bT --> bT),
+   ("op =", bT --> bT --> bT),
+   ("Not", bT --> bT),
+
+   ("op <=", iT --> iT --> bT),
+   ("op =", iT --> iT --> bT),
+   ("op <", iT --> iT --> bT),
+   ("Divides.op dvd", iT --> iT --> bT),
+   ("Divides.op div", iT --> iT --> iT),
+   ("Divides.op mod", iT --> iT --> iT),
+   ("op +", iT --> iT --> iT),
+   ("op -", iT --> iT --> iT),
+   ("op *", iT --> iT --> iT), 
+   ("HOL.abs", iT --> iT),
+   ("uminus", iT --> iT),
+
+   ("op <=", nT --> nT --> bT),
+   ("op =", nT --> nT --> bT),
+   ("op <", nT --> nT --> bT),
+   ("Divides.op dvd", nT --> nT --> bT),
+   ("Divides.op div", nT --> nT --> nT),
+   ("Divides.op mod", nT --> nT --> nT),
+   ("op +", nT --> nT --> nT),
+   ("op -", nT --> nT --> nT),
+   ("op *", nT --> nT --> nT), 
+   ("Suc", nT --> nT),
+
+   ("Numeral.bin.Bit", binT --> bT --> binT),
+   ("Numeral.bin.Pls", binT),
+   ("Numeral.bin.Min", binT),
+   ("Numeral.number_of", binT --> iT),
+   ("Numeral.number_of", binT --> nT),
+   ("0", nT),
+   ("0", iT),
+   ("1", nT),
+   ("1", iT),
+
+   ("False", bT),
+   ("True", bT)];
+
+(*returns true if the formula is relevant for presburger arithmetic tactic*)
+fun relevant t = (term_typed_consts t) subset allowed_consts;
+
+(* Preparation of the formula to be sent to the Integer quantifier *)
+(* elimination procedure                                           *)
+(* Transforms meta implications and meta quantifiers to object     *)
+(* implications and object quantifiers                             *)
+
+fun prepare_for_presburger q fm = 
+  let
+    val ps = Logic.strip_params fm
+    val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm)
+    val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm)
+    val _ = if relevant c then () else raise CooperDec.COOPER
+    fun mk_all ((s, T), (P,n)) =
+      if 0 mem loose_bnos P then
+        (HOLogic.all_const T $ Abs (s, T, P), n)
+      else (incr_boundvars ~1 P, n-1)
+    fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t;
+    val (rhs,irhs) = partition relevant hs
+    val np = length ps
+    val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
+      (ps,(foldr HOLogic.mk_imp (rhs, c), np))
+    val (vs, _) = partition (fn t => q orelse (type_of t) = nT)
+      (term_frees fm' @ term_vars fm');
+    val fm2 = foldr mk_all2 (vs, fm')
+  in (fm2, np + length vs, length rhs) end;
+
+(*Object quantifier to meta --*)
+fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ;
+
+(* object implication to meta---*)
+fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp;
+
+(* the presburger tactic*)
+fun presburger_tac q i st =
+  let
+    val g = BasisLibrary.List.nth (prems_of st, i - 1);
+    val sg = sign_of_thm st;
+    (* Transform the term*)
+    val (t,np,nh) = prepare_for_presburger q g
+    (* Some simpsets for dealing with mod div abs and nat*)
+
+    val simpset0 = HOL_basic_ss
+      addsimps [mod_div_equality', Suc_plus1]
+      addsplits [split_zdiv, split_zmod, split_div']
+    (* Simp rules for changing (n::int) to int n *)
+    val simpset1 = HOL_basic_ss
+      addsimps [nat_number_of_def, zdvd_int] @ map (fn r => r RS sym)
+        [int_int_eq, zle_int, zless_int, zadd_int, zmult_int]
+      addsplits [zdiff_int_split]
+    (*simp rules for elimination of int n*)
+
+    val simpset2 = HOL_basic_ss
+      addsimps [nat_0_le, all_nat, ex_nat, number_of1, number_of2, int_0, int_1]
+      addcongs [conj_le_cong, imp_le_cong]
+    (* simp rules for elimination of abs *)
+
+    val simpset3 = HOL_basic_ss addsplits [zabs_split]
+	      
+    val ct = cterm_of sg (HOLogic.mk_Trueprop t)
+
+    (* Theorem for the nat --> int transformation *)
+    val pre_thm = Seq.hd (EVERY
+      [simp_tac simpset0 i,
+       TRY (simp_tac simpset1 i), TRY (simp_tac simpset2 i),
+       TRY (simp_tac simpset3 i), TRY (simp_tac presburger_ss i)]
+      (trivial ct))
+
+    fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i);
+
+    (* The result of the quantifier elimination *)
+    val (th, tac) = case (prop_of pre_thm) of
+        Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
+          (trace_msg ("calling procedure with term:\n" ^
+             Sign.string_of_term sg t1);
+           ((mproof_of_int_qelim sg (eta_long [] t1) RS iffD2) RS pre_thm,
+            assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i)))
+      | _ => (pre_thm, assm_tac i)
+  in (rtac (((mp_step nh) o (spec_step np)) th) i THEN tac) st
+  end handle Subscript => no_tac st | CooperDec.COOPER => no_tac st;
+
+fun presburger_args meth =
+  Method.simple_args (Scan.optional (Args.$$$ "no_quantify" >> K false) true)
+    (fn q => fn _ => meth q 1);
+
+fun presburger_method q i = Method.METHOD (fn facts =>
+  Method.insert_tac facts 1 THEN presburger_tac q i)
+
+val setup =
+  [Method.add_method ("presburger",
+     presburger_args presburger_method,
+     "decision procedure for Presburger arithmetic"),
+   ArithTheoryData.map (fn {splits, inj_consts, discrete, presburger} =>
+     {splits = splits, inj_consts = inj_consts, discrete = discrete,
+      presburger = Some (presburger_tac true)})];
+
+end;
+
+val presburger_tac = Presburger.presburger_tac true;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Integ/qelim.ML	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,176 @@
+(*  Title:      HOL/Integ/qelim.ML
+    ID:         $Id$
+    Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+File containing the implementation of the proof protocoling
+and proof generation for multiple quantified formulae.
+*)
+
+signature QELIM =
+sig
+ val proof_of_mlift_qelim: Sign.sg -> (term -> bool) ->
+   (string list -> term -> thm) -> (term -> thm) ->
+   (string list -> term -> thm) -> term -> thm
+end;
+
+structure Qelim : QELIM =
+struct
+open CooperDec;
+open CooperProof;
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---         Protocoling part                                  ---*)
+(*---                                                           ---*)
+(*---           includes the protocolling datastructure         ---*)
+(*---                                                           ---*)
+(*---          and the protocolling fuctions                    ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+
+val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT;
+
+(* List of the theorems to be used in the following*)
+
+val qe_ex_conj = thm "qe_ex_conj";
+val qe_ex_nconj = thm "qe_ex_nconj";
+val qe_ALL = thm "qe_ALL";
+
+
+(*Datatype declaration for the protocoling procedure.*)
+
+
+datatype QeLog = AFN of term*(string list)
+                |QFN of term*(string list)
+                |ExConj of term*QeLog
+                |ExDisj of (string*typ*term)*term*QeLog*QeLog
+                |QeConst of string*QeLog*QeLog
+                |QeNot of QeLog
+                |QeAll of QeLog
+                |Lift_Qelim of term*QeLog
+		|QeUnk of term;
+
+datatype mQeLog = mQeProof of (string list)*mQeLog
+		|mAFN of term
+                |mNFN of mQeLog
+                |mQeConst of string*mQeLog*mQeLog
+                |mQeNot of mQeLog
+		|mQelim of term*(string list)*mQeLog
+		|mQeAll of mQeLog
+		|mQefm of term;
+
+(* This is the protokoling my function for the quantifier elimination*)
+fun mlift_qelim_wp isat vars = 
+  let fun mqelift_wp vars fm = if (isat fm) then mAFN(fm)
+    else  
+    (case fm of 
+     ( Const ("Not",_) $ p) => mQeNot(mqelift_wp vars p)
+    |( Const ("op &",_) $ p $q) => mQeConst("CJ", mqelift_wp vars p,mqelift_wp vars q) 
+
+    |( Const ("op |",_) $ p $q) => mQeConst("DJ", mqelift_wp vars p,mqelift_wp vars q) 
+
+    |( Const ("op -->",_) $ p $q) => mQeConst("IM", mqelift_wp vars p,mqelift_wp vars q) 
+
+    |( Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $q) =>mQeConst("EQ", mqelift_wp vars p,mqelift_wp vars q) 
+
+    				
+    |( Const ("All",QT) $ Abs(x,T,p)) =>mQeAll (mqelift_wp vars (Const("Ex",QT) $ Abs(x,T,(HOLogic.Not $ p))))
+
+    |(Const ("Ex",_) $ Abs (x,T,p))  => 
+      let val (x1,p1) = variant_abs (x,T,p)
+          val prt = mqelift_wp (x1::vars) p1
+      in mQelim(Free(x1,T),vars,mNFN(prt))
+      end
+    | _ => mQefm(fm)
+   ) 
+     
+  in (fn fm => mQeProof(vars,mNFN(mqelift_wp vars fm )))
+  end;  
+
+
+
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---      Interpretation and Proofgeneration Part              ---*)
+(*---                                                           ---*)
+(*---      Protocole interpretation functions                   ---*)
+(*---                                                           ---*)
+(*---      and proofgeneration functions                        ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*function that interpretates the protokol generated by the _wp function*)
+
+
+(* proof_of_lift_qelim interpretates a protokol for the quantifier elimination one some quantified formula. It uses the functions afnp nfnp and qfnp as proof functions to generate a prove for the hole quantifier elimination.*)
+(* afnp : must retun a proof for the transformations on the atomic formalae*)
+(* nfnp : must return a proof for the post one-quatifiers elimination process*)
+(* qfnp mus return a proof for the one quantifier elimination (existential) *)
+(* All these function are independent of the theory on whitch we are trying to prove quantifier elimination*)
+(* But the following invariants mus be respected : *)
+(* afnp : term -> string list -> thm*)
+(*   nfnp : term -> thm*)
+(*   qfnp : term -> string list -> thm*)
+(*For all theorms generated by these function must hold :*)
+(*    All of them are logical equivalences.*)
+(*    on left side of the equivalence must appear the term exactely as ist was given as a parameter (or eventually modulo Gamma, where Gamma are the rules whitch are allowed to be used during unification ex. beta reduction.....)*)
+(* qfnp must take as an argument for the term an existential quantified formula*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+fun proof_of_mlift_qelim sg isat afnp nfnp qfnp =
+ let fun h_proof_of_mlift_qelim isat afnp nfnp qfnp prtkl vrl = 
+   (case prtkl of 
+   mAFN (fm) => afnp vrl fm 
+   |mNFN (prt) => let val th1 = h_proof_of_mlift_qelim isat  afnp nfnp qfnp prt vrl 
+                  val th2 = nfnp (snd (qe_get_terms th1)) 
+                    in [th1,th2] MRS trans 
+                 end 
+   |mQeConst (s,pr1,pr2) =>  
+     let val th1 =  h_proof_of_mlift_qelim isat afnp nfnp qfnp pr1 vrl  
+         val th2 =  h_proof_of_mlift_qelim isat afnp nfnp qfnp pr2 vrl  
+     in (case s of 
+        "CJ" => [th1,th2] MRS (qe_conjI)  
+       |"DJ" => [th1,th2] MRS (qe_disjI)  
+       |"IM" => [th1,th2] MRS (qe_impI)  
+       |"EQ" => [th1,th2] MRS (qe_eqI)  
+       )  
+    end  
+   |mQeNot (pr) =>(h_proof_of_mlift_qelim isat afnp nfnp qfnp pr vrl ) RS (qe_Not)  
+   |mQeAll(pr) => (h_proof_of_mlift_qelim isat afnp nfnp qfnp pr vrl ) RS (qe_ALL) 
+   |mQelim (x as (Free(xn,xT)),vl,pr) => 
+     let val th_1 = h_proof_of_mlift_qelim isat afnp nfnp qfnp pr vl
+         val mQeProof(l2,pr2) = mlift_qelim_wp isat (xn::vrl) (snd(qe_get_terms th_1))
+         val th_2 = [th_1,(h_proof_of_mlift_qelim isat afnp nfnp qfnp pr2 l2)] MRS trans
+         val th1 = (forall_intr (cterm_of sg x) th_2) COMP (qe_exI)
+	 val th2 = qfnp vl (snd (qe_get_terms th1)) 
+       in [th1,th2] MRS trans 
+       end
+   |mQefm (fm) => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl
+)
+in (fn fm => let val mQeProof(vars,prt) = (mlift_qelim_wp isat (fv fm) fm)
+                 in (h_proof_of_mlift_qelim isat afnp nfnp qfnp prt vars)
+                 end)
+end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Presburger.thy	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,1002 @@
+(*  Title:      HOL/Integ/Presburger.thy
+    ID:         $Id$
+    Author:     Amine Chaieb, Tobias Nipkow and Stefan Berghofer, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+File containing necessary theorems for the proof
+generation for Cooper Algorithm  
+*)
+
+theory Presburger = NatSimprocs
+files
+  ("cooper_dec.ML")
+  ("cooper_proof.ML")
+  ("qelim.ML")
+  ("presburger.ML"):
+
+(* Theorem for unitifying the coeffitients of x in an existential formula*)
+
+theorem unity_coeff_ex: "(\<exists>x::int. P (l * x)) = (\<exists>x. l dvd (1*x+0) \<and> P x)"
+  apply (rule iffI)
+  apply (erule exE)
+  apply (rule_tac x = "l * x" in exI)
+  apply simp
+  apply (erule exE)
+  apply (erule conjE)
+  apply (erule dvdE)
+  apply (rule_tac x = k in exI)
+  apply simp
+  done
+
+lemma uminus_dvd_conv: "(d dvd (t::int)) = (-d dvd t)"
+apply(unfold dvd_def)
+apply(rule iffI)
+apply(clarsimp)
+apply(rename_tac k)
+apply(rule_tac x = "-k" in exI)
+apply simp
+apply(clarsimp)
+apply(rename_tac k)
+apply(rule_tac x = "-k" in exI)
+apply simp
+done
+
+lemma uminus_dvd_conv': "(d dvd (t::int)) = (d dvd -t)"
+apply(unfold dvd_def)
+apply(rule iffI)
+apply(clarsimp)
+apply(rule_tac x = "-k" in exI)
+apply simp
+apply(clarsimp)
+apply(rule_tac x = "-k" in exI)
+apply simp
+done
+
+
+
+(*Theorems for the combination of proofs of the equality of P and P_m for integers x less than some integer z.*)
+
+theorem eq_minf_conjI: "\<exists>z1::int. \<forall>x. x < z1 \<longrightarrow> (A1 x = A2 x) \<Longrightarrow>
+  \<exists>z2::int. \<forall>x. x < z2 \<longrightarrow> (B1 x = B2 x) \<Longrightarrow>
+  \<exists>z::int. \<forall>x. x < z \<longrightarrow> ((A1 x \<and> B1 x) = (A2 x \<and> B2 x))"
+  apply (erule exE)+
+  apply (rule_tac x = "min z1 z2" in exI)
+  apply simp
+  done
+
+
+theorem eq_minf_disjI: "\<exists>z1::int. \<forall>x. x < z1 \<longrightarrow> (A1 x = A2 x) \<Longrightarrow>
+  \<exists>z2::int. \<forall>x. x < z2 \<longrightarrow> (B1 x = B2 x) \<Longrightarrow>
+  \<exists>z::int. \<forall>x. x < z \<longrightarrow> ((A1 x \<or> B1 x) = (A2 x \<or> B2 x))"
+
+  apply (erule exE)+
+  apply (rule_tac x = "min z1 z2" in exI)
+  apply simp
+  done
+
+
+(*Theorems for the combination of proofs of the equality of P and P_m for integers x greather than some integer z.*)
+
+theorem eq_pinf_conjI: "\<exists>z1::int. \<forall>x. z1 < x \<longrightarrow> (A1 x = A2 x) \<Longrightarrow>
+  \<exists>z2::int. \<forall>x. z2 < x \<longrightarrow> (B1 x = B2 x) \<Longrightarrow>
+  \<exists>z::int. \<forall>x. z < x \<longrightarrow> ((A1 x \<and> B1 x) = (A2 x \<and> B2 x))"
+  apply (erule exE)+
+  apply (rule_tac x = "max z1 z2" in exI)
+  apply simp
+  done
+
+
+theorem eq_pinf_disjI: "\<exists>z1::int. \<forall>x. z1 < x \<longrightarrow> (A1 x = A2 x) \<Longrightarrow>
+  \<exists>z2::int. \<forall>x. z2 < x \<longrightarrow> (B1 x = B2 x) \<Longrightarrow>
+  \<exists>z::int. \<forall>x. z < x  \<longrightarrow> ((A1 x \<or> B1 x) = (A2 x \<or> B2 x))"
+  apply (erule exE)+
+  apply (rule_tac x = "max z1 z2" in exI)
+  apply simp
+  done
+(*=============================================================================*)
+(*Theorems for the combination of proofs of the modulo D property for P
+pluusinfinity*)
+(* FIXME : This is THE SAME theorem as for the minusinf version, but with +k.. instead of -k.. In the future replace these both with only one*)
+
+theorem modd_pinf_conjI: "\<forall>(x::int) k. A x = A (x+k*d) \<Longrightarrow>
+  \<forall>(x::int) k. B x = B (x+k*d) \<Longrightarrow>
+  \<forall>(x::int) (k::int). (A x \<and> B x) = (A (x+k*d) \<and> B (x+k*d))"
+  by simp
+
+
+theorem modd_pinf_disjI: "\<forall>(x::int) k. A x = A (x+k*d) \<Longrightarrow>
+  \<forall>(x::int) k. B x = B (x+k*d) \<Longrightarrow>
+  \<forall>(x::int) (k::int). (A x \<or> B x) = (A (x+k*d) \<or> B (x+k*d))"
+  by simp
+
+(*=============================================================================*)
+(*This is one of the cases where the simplifed formula is prooved to habe some property
+(in relation to P_m) but we need to proove the property for the original formula (P_m)*)
+(*FIXME : This is exaclty the same thm as for minusinf.*)
+lemma pinf_simp_eq: "ALL x. P(x) = Q(x) ==> (EX (x::int). P(x)) --> (EX (x::int). F(x))  ==> (EX (x::int). Q(x)) --> (EX (x::int). F(x)) "
+by blast
+
+
+
+(*=============================================================================*)
+(*Theorems for the combination of proofs of the modulo D property for P
+minusinfinity*)
+
+theorem modd_minf_conjI: "\<forall>(x::int) k. A x = A (x-k*d) \<Longrightarrow>
+  \<forall>(x::int) k. B x = B (x-k*d) \<Longrightarrow>
+  \<forall>(x::int) (k::int). (A x \<and> B x) = (A (x-k*d) \<and> B (x-k*d))"
+  by simp
+
+
+theorem modd_minf_disjI: "\<forall>(x::int) k. A x = A (x-k*d) \<Longrightarrow>
+  \<forall>(x::int) k. B x = B (x-k*d) \<Longrightarrow>
+  \<forall>(x::int) (k::int). (A x \<or> B x) = (A (x-k*d) \<or> B (x-k*d))"
+  by simp
+
+(*=============================================================================*)
+(*This is one of the cases where the simplifed formula is prooved to habe some property
+(in relation to P_m) but we need to proove the property for the original formula (P_m)*)
+
+lemma minf_simp_eq: "ALL x. P(x) = Q(x) ==> (EX (x::int). P(x)) --> (EX (x::int). F(x))  ==> (EX (x::int). Q(x)) --> (EX (x::int). F(x)) "
+by blast
+
+(*=============================================================================*)
+
+(*theorem needed for prooving at runtime divide properties using the arithmetic tatic
+(who knows only about modulo = 0)*)
+
+lemma zdvd_iff_zmod_eq_0: "(m dvd n) = (n mod m = (0::int))"
+by(simp add:dvd_def zmod_eq_0_iff)
+
+(*=============================================================================*)
+
+
+
+(*Theorems used for the combination of proof for the backwards direction of cooper's
+theorem. they rely exclusively on Predicate calculus.*)
+
+lemma not_ast_p_disjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P1(x) --> P1(x + d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P2(x) --> P2(x + d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->(P1(x) \<or> P2(x)) --> (P1(x + d) \<or> P2(x + d))) "
+by blast
+
+
+
+lemma not_ast_p_conjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a- j)) --> P1(x) --> P1(x + d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P2(x) --> P2(x + d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->(P1(x) \<and> P2(x)) --> (P1(x + d)
+\<and> P2(x + d))) "
+by blast
+
+lemma not_ast_p_Q_elim: "
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->P(x) --> P(x + d))
+==> ( P = Q )
+==> (ALL x. ~(EX (j::int) : {1..d}. EX (a::int) : A. P(a - j)) -->P(x) --> P(x + d))"
+by blast
+(*=============================================================================*)
+
+
+(*Theorems used for the combination of proof for the backwards direction of cooper's
+theorem. they rely exclusively on Predicate calculus.*)
+
+lemma not_bst_p_disjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P1(x) --> P1(x - d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P2(x) --> P2(x - d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->(P1(x) \<or> P2(x)) --> (P1(x - d)
+\<or> P2(x-d))) "
+by blast
+
+
+
+lemma not_bst_p_conjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P1(x) --> P1(x - d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P2(x) --> P2(x - d))
+==>
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->(P1(x) \<and> P2(x)) --> (P1(x - d)
+\<and> P2(x-d))) "
+by blast
+
+lemma not_bst_p_Q_elim: "
+(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->P(x) --> P(x - d)) 
+==> ( P = Q )
+==> (ALL x. ~(EX (j::int) : {1..d}. EX (b::int) : B. P(b+j)) -->P(x) --> P(x - d))"
+by blast
+(*=============================================================================*)
+(*The Theorem for the second proof step- about bset. it is trivial too. *)
+lemma bst_thm: " (EX (j::int) : {1..d}. EX (b::int) : B. P (b+j) )--> (EX x::int. P (x)) "
+by blast
+
+(*The Theorem for the second proof step- about aset. it is trivial too. *)
+lemma ast_thm: " (EX (j::int) : {1..d}. EX (a::int) : A. P (a - j) )--> (EX x::int. P (x)) "
+by blast
+
+
+(*=============================================================================*)
+(*This is the first direction of cooper's theorem*)
+lemma cooper_thm: "(R --> (EX x::int. P x))  ==> (Q -->(EX x::int.  P x )) ==> ((R|Q) --> (EX x::int. P x )) "
+by blast
+
+(*=============================================================================*)
+(*The full cooper's theoorem in its equivalence Form- Given the premisses it is trivial
+too, it relies exclusively on prediacte calculus.*)
+lemma cooper_eq_thm: "(R --> (EX x::int. P x))  ==> (Q -->(EX x::int.  P x )) ==> ((~Q)
+--> (EX x::int. P x ) --> R) ==> (EX x::int. P x) = R|Q "
+by blast
+
+(*=============================================================================*)
+(*Some of the atomic theorems generated each time the atom does not depend on x, they
+are trivial.*)
+
+lemma  fm_eq_minf: "EX z::int. ALL x. x < z --> (P = P) "
+by blast
+
+lemma  fm_modd_minf: "ALL (x::int). ALL (k::int). (P = P)"
+by blast
+
+lemma not_bst_p_fm: "ALL (x::int). Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> fm --> fm"
+by blast
+
+
+
+lemma  fm_eq_pinf: "EX z::int. ALL x. z < x --> (P = P) "
+by blast
+
+(* The next 2 thms are the same as the minusinf version*)
+lemma  fm_modd_pinf: "ALL (x::int). ALL (k::int). (P = P)"
+by blast
+
+lemma not_ast_p_fm: "ALL (x::int). Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> fm --> fm"
+by blast
+
+
+(* Theorems to be deleted from simpset when proving simplified formulaes*)
+lemma P_eqtrue: "(P=True) = P"
+  by rules
+
+lemma P_eqfalse: "(P=False) = (~P)"
+  by rules
+
+(*=============================================================================*)
+
+(*Theorems for the generation of the bachwards direction of cooper's theorem*)
+(*These are the 6 interesting atomic cases which have to be proved relying on the
+properties of B-set ant the arithmetic and contradiction proofs*)
+
+lemma not_bst_p_lt: "0 < (d::int) ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ( 0 < -x + a) --> (0 < -(x - d) + a )"
+by arith
+
+lemma not_bst_p_gt: "\<lbrakk> (g::int) \<in> B; g = -a \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> (0 < (x) + a) --> ( 0 < (x - d) + a)"
+apply clarsimp
+apply(rule ccontr)
+apply(drule_tac x = "x+a" in bspec)
+apply(simp add:atLeastAtMost_iff)
+apply(drule_tac x = "-a" in bspec)
+apply assumption
+apply(simp)
+done
+
+lemma not_bst_p_eq: "\<lbrakk> 0 < d; (g::int) \<in> B; g = -a - 1 \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> (0 = x + a) --> (0 = (x - d) + a )"
+apply clarsimp
+apply(subgoal_tac "x = -a")
+ prefer 2 apply arith
+apply(drule_tac x = "1" in bspec)
+apply(simp add:atLeastAtMost_iff)
+apply(drule_tac x = "-a- 1" in bspec)
+apply assumption
+apply(simp)
+done
+
+
+lemma not_bst_p_ne: "\<lbrakk> 0 < d; (g::int) \<in> B; g = -a \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ~(0 = x + a) --> ~(0 = (x - d) + a)"
+apply clarsimp
+apply(subgoal_tac "x = -a+d")
+ prefer 2 apply arith
+apply(drule_tac x = "d" in bspec)
+apply(simp add:atLeastAtMost_iff)
+apply(drule_tac x = "-a" in bspec)
+apply assumption
+apply(simp)
+done
+
+
+lemma not_bst_p_dvd: "(d1::int) dvd d ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> d1 dvd (x + a) --> d1 dvd ((x - d) + a )"
+apply(clarsimp simp add:dvd_def)
+apply(rename_tac m)
+apply(rule_tac x = "m - k" in exI)
+apply(simp add:int_distrib)
+done
+
+lemma not_bst_p_ndvd: "(d1::int) dvd d ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ~(d1 dvd (x + a)) --> ~(d1 dvd ((x - d) + a ))"
+apply(clarsimp simp add:dvd_def)
+apply(rename_tac m)
+apply(erule_tac x = "m + k" in allE)
+apply(simp add:int_distrib)
+done
+
+
+
+(*Theorems for the generation of the bachwards direction of cooper's theorem*)
+(*These are the 6 interesting atomic cases which have to be proved relying on the
+properties of A-set ant the arithmetic and contradiction proofs*)
+
+lemma not_ast_p_gt: "0 < (d::int) ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ( 0 < x + t) --> (0 < (x + d) + t )"
+by arith
+
+
+lemma not_ast_p_lt: "\<lbrakk>0 < d ;(t::int) \<in> A \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> (0 < -x + t) --> ( 0 < -(x + d) + t)"
+  apply clarsimp
+  apply (rule ccontr)
+  apply (drule_tac x = "t-x" in bspec)
+  apply simp
+  apply (drule_tac x = "t" in bspec)
+  apply assumption
+  apply simp
+  done
+
+lemma not_ast_p_eq: "\<lbrakk> 0 < d; (g::int) \<in> A; g = -t + 1 \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> (0 = x + t) --> (0 = (x + d) + t )"
+  apply clarsimp
+  apply (drule_tac x="1" in bspec)
+  apply simp
+  apply (drule_tac x="- t + 1" in bspec)
+  apply assumption
+  apply(subgoal_tac "x = -t")
+  prefer 2 apply arith
+  apply simp
+  done
+
+lemma not_ast_p_ne: "\<lbrakk> 0 < d; (g::int) \<in> A; g = -t \<rbrakk> \<Longrightarrow>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ~(0 = x + t) --> ~(0 = (x + d) + t)"
+  apply clarsimp
+  apply (subgoal_tac "x = -t-d")
+  prefer 2 apply arith
+  apply (drule_tac x = "d" in bspec)
+  apply simp
+  apply (drule_tac x = "-t" in bspec)
+  apply assumption
+  apply simp
+  done
+
+lemma not_ast_p_dvd: "(d1::int) dvd d ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> d1 dvd (x + t) --> d1 dvd ((x + d) + t )"
+  apply(clarsimp simp add:dvd_def)
+  apply(rename_tac m)
+  apply(rule_tac x = "m + k" in exI)
+  apply(simp add:int_distrib)
+  done
+
+lemma not_ast_p_ndvd: "(d1::int) dvd d ==>
+ ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ~(d1 dvd (x + t)) --> ~(d1 dvd ((x + d) + t ))"
+  apply(clarsimp simp add:dvd_def)
+  apply(rename_tac m)
+  apply(erule_tac x = "m - k" in allE)
+  apply(simp add:int_distrib)
+  done
+
+
+
+(*=============================================================================*)
+(*These are the atomic cases for the proof generation for the modulo D property for P
+plusinfinity*)
+(*They are fully based on arithmetics*)
+
+lemma  dvd_modd_pinf: "((d::int) dvd d1) ==>
+ (ALL (x::int). ALL (k::int). (((d::int) dvd (x + t)) = (d dvd (x+k*d1 + t))))"
+  apply(clarsimp simp add:dvd_def)
+  apply(rule iffI)
+  apply(clarsimp)
+  apply(rename_tac n m)
+  apply(rule_tac x = "m + n*k" in exI)
+  apply(simp add:int_distrib)
+  apply(clarsimp)
+  apply(rename_tac n m)
+  apply(rule_tac x = "m - n*k" in exI)
+  apply(simp add:int_distrib zmult_ac)
+  done
+
+lemma  not_dvd_modd_pinf: "((d::int) dvd d1) ==>
+ (ALL (x::int). ALL k. (~((d::int) dvd (x + t))) = (~(d dvd (x+k*d1 + t))))"
+  apply(clarsimp simp add:dvd_def)
+  apply(rule iffI)
+  apply(clarsimp)
+  apply(rename_tac n m)
+  apply(erule_tac x = "m - n*k" in allE)
+  apply(simp add:int_distrib zmult_ac)
+  apply(clarsimp)
+  apply(rename_tac n m)
+  apply(erule_tac x = "m + n*k" in allE)
+  apply(simp add:int_distrib zmult_ac)
+  done
+
+(*=============================================================================*)
+(*These are the atomic cases for the proof generation for the equivalence of P and P
+plusinfinity for integers x greather than some integer z.*)
+(*They are fully based on arithmetics*)
+
+lemma  eq_eq_pinf: "EX z::int. ALL x. z < x --> (( 0 = x +t ) = False )"
+  apply(rule_tac x = "-t" in exI)
+  apply simp
+  done
+
+lemma  neq_eq_pinf: "EX z::int. ALL x.  z < x --> ((~( 0 = x +t )) = True )"
+  apply(rule_tac x = "-t" in exI)
+  apply simp
+  done
+
+lemma  le_eq_pinf: "EX z::int. ALL x.  z < x --> ( 0 < x +t  = True )"
+  apply(rule_tac x = "-t" in exI)
+  apply simp
+  done
+
+lemma  len_eq_pinf: "EX z::int. ALL x. z < x  --> (0 < -x +t  = False )"
+  apply(rule_tac x = "t" in exI)
+  apply simp
+  done
+
+lemma  dvd_eq_pinf: "EX z::int. ALL x.  z < x --> ((d dvd (x + t)) = (d dvd (x + t))) "
+by simp
+
+lemma  not_dvd_eq_pinf: "EX z::int. ALL x. z < x  --> ((~(d dvd (x + t))) = (~(d dvd (x + t)))) "
+by simp
+
+
+
+
+(*=============================================================================*)
+(*These are the atomic cases for the proof generation for the modulo D property for P
+minusinfinity*)
+(*They are fully based on arithmetics*)
+
+lemma  dvd_modd_minf: "((d::int) dvd d1) ==>
+ (ALL (x::int). ALL (k::int). (((d::int) dvd (x + t)) = (d dvd (x-k*d1 + t))))"
+apply(clarsimp simp add:dvd_def)
+apply(rule iffI)
+apply(clarsimp)
+apply(rename_tac n m)
+apply(rule_tac x = "m - n*k" in exI)
+apply(simp add:int_distrib)
+apply(clarsimp)
+apply(rename_tac n m)
+apply(rule_tac x = "m + n*k" in exI)
+apply(simp add:int_distrib zmult_ac)
+done
+
+
+lemma  not_dvd_modd_minf: "((d::int) dvd d1) ==>
+ (ALL (x::int). ALL k. (~((d::int) dvd (x + t))) = (~(d dvd (x-k*d1 + t))))"
+apply(clarsimp simp add:dvd_def)
+apply(rule iffI)
+apply(clarsimp)
+apply(rename_tac n m)
+apply(erule_tac x = "m + n*k" in allE)
+apply(simp add:int_distrib zmult_ac)
+apply(clarsimp)
+apply(rename_tac n m)
+apply(erule_tac x = "m - n*k" in allE)
+apply(simp add:int_distrib zmult_ac)
+done
+
+
+(*=============================================================================*)
+(*These are the atomic cases for the proof generation for the equivalence of P and P
+minusinfinity for integers x less than some integer z.*)
+(*They are fully based on arithmetics*)
+
+lemma  eq_eq_minf: "EX z::int. ALL x. x < z --> (( 0 = x +t ) = False )"
+apply(rule_tac x = "-t" in exI)
+apply simp
+done
+
+lemma  neq_eq_minf: "EX z::int. ALL x. x < z --> ((~( 0 = x +t )) = True )"
+apply(rule_tac x = "-t" in exI)
+apply simp
+done
+
+lemma  le_eq_minf: "EX z::int. ALL x. x < z --> ( 0 < x +t  = False )"
+apply(rule_tac x = "-t" in exI)
+apply simp
+done
+
+
+lemma  len_eq_minf: "EX z::int. ALL x. x < z --> (0 < -x +t  = True )"
+apply(rule_tac x = "t" in exI)
+apply simp
+done
+
+lemma  dvd_eq_minf: "EX z::int. ALL x. x < z --> ((d dvd (x + t)) = (d dvd (x + t))) "
+by simp
+
+lemma  not_dvd_eq_minf: "EX z::int. ALL x. x < z --> ((~(d dvd (x + t))) = (~(d dvd (x + t)))) "
+by simp
+
+
+(*=============================================================================*)
+(*This Theorem combines whithnesses about P minusinfinity to schow one component of the
+equivalence proof for cooper's theorem*)
+
+(* FIXME: remove once they are part of the distribution *)
+theorem int_ge_induct[consumes 1,case_names base step]:
+  assumes ge: "k \<le> (i::int)" and
+        base: "P(k)" and
+        step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
+  shows "P i"
+proof -
+  { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k <= i \<Longrightarrow> P i"
+    proof (induct n)
+      case 0
+      hence "i = k" by arith
+      thus "P i" using base by simp
+    next
+      case (Suc n)
+      hence "n = nat((i - 1) - k)" by arith
+      moreover
+      have ki1: "k \<le> i - 1" using Suc.prems by arith
+      ultimately
+      have "P(i - 1)" by(rule Suc.hyps)
+      from step[OF ki1 this] show ?case by simp
+    qed
+  }
+  from this ge show ?thesis by fast
+qed
+
+theorem int_gr_induct[consumes 1,case_names base step]:
+  assumes gr: "k < (i::int)" and
+        base: "P(k+1)" and
+        step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)"
+  shows "P i"
+apply(rule int_ge_induct[of "k + 1"])
+  using gr apply arith
+ apply(rule base)
+apply(rule step)
+ apply simp+
+done
+
+lemma decr_lemma: "0 < (d::int) \<Longrightarrow> x - (abs(x-z)+1) * d < z"
+apply(induct rule: int_gr_induct)
+ apply simp
+ apply arith
+apply (simp add:int_distrib)
+apply arith
+done
+
+lemma incr_lemma: "0 < (d::int) \<Longrightarrow> z < x + (abs(x-z)+1) * d"
+apply(induct rule: int_gr_induct)
+ apply simp
+ apply arith
+apply (simp add:int_distrib)
+apply arith
+done
+
+lemma  minusinfinity:
+  assumes "0 < d" and
+    P1eqP1: "ALL x k. P1 x = P1(x - k*d)" and
+    ePeqP1: "EX z::int. ALL x. x < z \<longrightarrow> (P x = P1 x)"
+  shows "(EX x. P1 x) \<longrightarrow> (EX x. P x)"
+proof
+  assume eP1: "EX x. P1 x"
+  then obtain x where P1: "P1 x" ..
+  from ePeqP1 obtain z where P1eqP: "ALL x. x < z \<longrightarrow> (P x = P1 x)" ..
+  let ?w = "x - (abs(x-z)+1) * d"
+  show "EX x. P x"
+  proof
+    have w: "?w < z" by(rule decr_lemma)
+    have "P1 x = P1 ?w" using P1eqP1 by blast
+    also have "\<dots> = P(?w)" using w P1eqP by blast
+    finally show "P ?w" using P1 by blast
+  qed
+qed
+
+(*=============================================================================*)
+(*This Theorem combines whithnesses about P minusinfinity to schow one component of the
+equivalence proof for cooper's theorem*)
+
+lemma plusinfinity:
+  assumes "0 < d" and
+    P1eqP1: "ALL (x::int) (k::int). P1 x = P1 (x + k * d)" and
+    ePeqP1: "EX z::int. ALL x. z < x  --> (P x = P1 x)"
+  shows "(EX x::int. P1 x) --> (EX x::int. P x)"
+proof
+  assume eP1: "EX x. P1 x"
+  then obtain x where P1: "P1 x" ..
+  from ePeqP1 obtain z where P1eqP: "ALL x. z < x \<longrightarrow> (P x = P1 x)" ..
+  let ?w = "x + (abs(x-z)+1) * d"
+  show "EX x. P x"
+  proof
+    have w: "z < ?w" by(rule incr_lemma)
+    have "P1 x = P1 ?w" using P1eqP1 by blast
+    also have "\<dots> = P(?w)" using w P1eqP by blast
+    finally show "P ?w" using P1 by blast
+  qed
+qed
+ 
+
+
+(*=============================================================================*)
+(*Theorem for periodic function on discrete sets*)
+
+lemma minf_vee:
+  assumes dpos: "(0::int) < d" and modd: "ALL x k. P x = P(x - k*d)"
+  shows "(EX x. P x) = (EX j : {1..d}. P j)"
+  (is "?LHS = ?RHS")
+proof
+  assume ?LHS
+  then obtain x where P: "P x" ..
+  have "x mod d = x - (x div d)*d"
+    by(simp add:zmod_zdiv_equality zmult_ac eq_zdiff_eq)
+  hence Pmod: "P x = P(x mod d)" using modd by simp
+  show ?RHS
+  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 ?RHS ..
+  next
+    assume not0: "x mod d \<noteq> 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 \<le> x mod d" by(rule pos_mod_sign)
+      moreover have "x mod d < d" by(rule pos_mod_bound)
+      ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff)
+    qed
+    ultimately show ?RHS ..
+  qed
+next
+  assume ?RHS thus ?LHS by blast
+qed
+
+(*=============================================================================*)
+(*Theorem for periodic function on discrete sets*)
+lemma pinf_vee:
+  assumes dpos: "0 < (d::int)" and modd: "ALL (x::int) (k::int). P x = P (x+k*d)"
+  shows "(EX x::int. P x) = (EX (j::int) : {1..d} . P j)"
+  (is "?LHS = ?RHS")
+proof
+  assume ?LHS
+  then obtain x where P: "P x" ..
+  have "x mod d = x + (-(x div d))*d"
+    by(simp add:zmod_zdiv_equality zmult_ac eq_zdiff_eq)
+  hence Pmod: "P x = P(x mod d)" using modd by (simp only:)
+  show ?RHS
+  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 ?RHS ..
+  next
+    assume not0: "x mod d \<noteq> 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 \<le> x mod d" by(rule pos_mod_sign)
+      moreover have "x mod d < d" by(rule pos_mod_bound)
+      ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff)
+    qed
+    ultimately show ?RHS ..
+  qed
+next
+  assume ?RHS thus ?LHS by blast
+qed
+
+lemma decr_mult_lemma:
+  assumes dpos: "(0::int) < d" and
+          minus: "ALL x::int. P x \<longrightarrow> P(x - d)" and
+          knneg: "0 <= k"
+  shows "ALL x. P x \<longrightarrow> P(x - k*d)"
+using knneg
+proof (induct rule:int_ge_induct)
+  case base thus ?case by simp
+next
+  case (step i)
+  show ?case
+  proof
+    fix x
+    have "P x \<longrightarrow> P (x - i * d)" using step.hyps by blast
+    also have "\<dots> \<longrightarrow> P(x - (i + 1) * d)"
+      using minus[THEN spec, of "x - i * d"]
+      by (simp add:int_distrib zdiff_zdiff_eq[symmetric])
+    ultimately show "P x \<longrightarrow> P(x - (i + 1) * d)" by blast
+  qed
+qed
+
+lemma incr_mult_lemma:
+  assumes dpos: "(0::int) < d" and
+          plus: "ALL x::int. P x \<longrightarrow> P(x + d)" and
+          knneg: "0 <= k"
+  shows "ALL x. P x \<longrightarrow> P(x + k*d)"
+using knneg
+proof (induct rule:int_ge_induct)
+  case base thus ?case by simp
+next
+  case (step i)
+  show ?case
+  proof
+    fix x
+    have "P x \<longrightarrow> P (x + i * d)" using step.hyps by blast
+    also have "\<dots> \<longrightarrow> P(x + (i + 1) * d)"
+      using plus[THEN spec, of "x + i * d"]
+      by (simp add:int_distrib zadd_ac)
+    ultimately show "P x \<longrightarrow> P(x + (i + 1) * d)" by blast
+  qed
+qed
+
+lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x))
+==> (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)) --> (EX (x::int). P x) 
+==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) 
+==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D))))
+==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))"
+apply(rule iffI)
+prefer 2
+apply(drule minusinfinity)
+apply assumption+
+apply(fastsimp)
+apply clarsimp
+apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)")
+apply(frule_tac x = x and z=z in decr_lemma)
+apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)")
+prefer 2
+apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
+prefer 2 apply arith
+ apply fastsimp
+apply(drule (1) minf_vee)
+apply blast
+apply(blast dest:decr_mult_lemma)
+done
+
+(* Cooper Thm `, plus infinity version*)
+lemma cppi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. z < x --> (P x = P1 x))
+==> (EX (j::int) : {1..D}. EX (a::int) : A. P (a - j)) --> (EX (x::int). P x) 
+==> ALL x.~(EX (j::int) : {1..D}. EX (a::int) : A. P(a - j)) --> P (x) --> P (x + D) 
+==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x+k*D))))
+==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (a::int) : A. P (a - j)))"
+  apply(rule iffI)
+  prefer 2
+  apply(drule plusinfinity)
+  apply assumption+
+  apply(fastsimp)
+  apply clarsimp
+  apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x + k*D)")
+  apply(frule_tac x = x and z=z in incr_lemma)
+  apply(subgoal_tac "P1(x + (\<bar>x - z\<bar> + 1) * D)")
+  prefer 2
+  apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
+  prefer 2 apply arith
+  apply fastsimp
+  apply(drule (1) pinf_vee)
+  apply blast
+  apply(blast dest:incr_mult_lemma)
+  done
+
+
+(*=============================================================================*)
+
+(*Theorems for the quantifier elminination Functions.*)
+
+lemma qe_ex_conj: "(EX (x::int). A x) = R
+		==> (EX (x::int). P x) = (Q & (EX x::int. A x))
+		==> (EX (x::int). P x) = (Q & R)"
+by blast
+
+lemma qe_ex_nconj: "(EX (x::int). P x) = (True & Q)
+		==> (EX (x::int). P x) = Q"
+by blast
+
+lemma qe_conjI: "P1 = P2 ==> Q1 = Q2 ==> (P1 & Q1) = (P2 & Q2)"
+by blast
+
+lemma qe_disjI: "P1 = P2 ==> Q1 = Q2 ==> (P1 | Q1) = (P2 | Q2)"
+by blast
+
+lemma qe_impI: "P1 = P2 ==> Q1 = Q2 ==> (P1 --> Q1) = (P2 --> Q2)"
+by blast
+
+lemma qe_eqI: "P1 = P2 ==> Q1 = Q2 ==> (P1 = Q1) = (P2 = Q2)"
+by blast
+
+lemma qe_Not: "P = Q ==> (~P) = (~Q)"
+by blast
+
+lemma qe_ALL: "(EX x. ~P x) = R ==> (ALL x. P x) = (~R)"
+by blast
+
+(* Theorems for proving NNF *)
+
+lemma nnf_im: "((~P) = P1) ==> (Q=Q1) ==> ((P --> Q) = (P1 | Q1))"
+by blast
+
+lemma nnf_eq: "((P & Q) = (P1 & Q1)) ==> (((~P) & (~Q)) = (P2 & Q2)) ==> ((P = Q) = ((P1 & Q1)|(P2 & Q2)))"
+by blast
+
+lemma nnf_nn: "(P = Q) ==> ((~~P) = Q)"
+  by blast
+lemma nnf_ncj: "((~P) = P1) ==> ((~Q) = Q1) ==> ((~(P & Q)) = (P1 | Q1))"
+by blast
+
+lemma nnf_ndj: "((~P) = P1) ==> ((~Q) = Q1) ==> ((~(P | Q)) = (P1 & Q1))"
+by blast
+lemma nnf_nim: "(P = P1) ==> ((~Q) = Q1) ==> ((~(P --> Q)) = (P1 & Q1))"
+by blast
+lemma nnf_neq: "((P & (~Q)) = (P1 & Q1)) ==> (((~P) & Q) = (P2 & Q2)) ==> ((~(P = Q)) = ((P1 & Q1)|(P2 & Q2)))"
+by blast
+lemma nnf_sdj: "((A & (~B)) = (A1 & B1)) ==> ((C & (~D)) = (C1 & D1)) ==> (A = (~C)) ==> ((~((A & B) | (C & D))) = ((A1 & B1) | (C1 & D1)))"
+by blast
+
+
+lemma qe_exI2: "A = B ==> (EX (x::int). A(x)) = (EX (x::int). B(x))"
+  by simp
+
+lemma qe_exI: "(!!x::int. A x = B x) ==> (EX (x::int). A(x)) = (EX (x::int). B(x))"
+  by rules
+
+lemma qe_ALLI: "(!!x::int. A x = B x) ==> (ALL (x::int). A(x)) = (ALL (x::int). B(x))"
+  by rules
+
+lemma cp_expand: "(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (b::int) : B. (P1 (j) | P(b+j)))
+==>(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (b::int) : B. (P1 (j) | P(b+j))) "
+by blast
+
+lemma cppi_expand: "(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (a::int) : A. (P1 (j) | P(a - j)))
+==>(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (a::int) : A. (P1 (j) | P(a - j))) "
+by blast
+
+
+lemma simp_from_to: "{i..j::int} = (if j < i then {} else insert i {i+1..j})"
+apply(simp add:atLeastAtMost_def atLeast_def atMost_def)
+apply(fastsimp)
+done
+
+(* Theorems required for the adjustcoeffitienteq*)
+
+lemma ac_dvd_eq: assumes not0: "0 ~= (k::int)"
+shows "((m::int) dvd (c*n+t)) = (k*m dvd ((k*c)*n+(k*t)))" (is "?P = ?Q")
+proof
+  assume ?P
+  thus ?Q
+    apply(simp add:dvd_def)
+    apply clarify
+    apply(rename_tac d)
+    apply(drule_tac f = "op * k" in arg_cong)
+    apply(simp only:int_distrib)
+    apply(rule_tac x = "d" in exI)
+    apply(simp only:zmult_ac)
+    done
+next
+  assume ?Q
+  then obtain d where "k * c * n + k * t = (k*m)*d" by(fastsimp simp:dvd_def)
+  hence "(c * n + t) * k = (m*d) * k" by(simp add:int_distrib zmult_ac)
+  hence "((c * n + t) * k) div k = ((m*d) * k) div k" by(rule arg_cong[of _ _ "%t. t div k"])
+  hence "c*n+t = m*d" by(simp add: zdiv_zmult_self1[OF not0[symmetric]])
+  thus ?P by(simp add:dvd_def)
+qed
+
+lemma ac_lt_eq: assumes gr0: "0 < (k::int)"
+shows "((m::int) < (c*n+t)) = (k*m <((k*c)*n+(k*t)))" (is "?P = ?Q")
+proof
+  assume P: ?P
+  show ?Q using zmult_zless_mono2[OF P gr0] by(simp add: int_distrib zmult_ac)
+next
+  assume ?Q
+  hence "0 < k*(c*n + t - m)" by(simp add: int_distrib zmult_ac)
+  with gr0 have "0 < (c*n + t - m)" by(simp add:int_0_less_mult_iff)
+  thus ?P by(simp)
+qed
+
+lemma ac_eq_eq : assumes not0: "0 ~= (k::int)" shows "((m::int) = (c*n+t)) = (k*m =((k*c)*n+(k*t)) )" (is "?P = ?Q")
+proof
+  assume ?P
+  thus ?Q
+    apply(drule_tac f = "op * k" in arg_cong)
+    apply(simp only:int_distrib)
+    done
+next
+  assume ?Q
+  hence "m * k = (c*n + t) * k" by(simp add:int_distrib zmult_ac)
+  hence "((m) * k) div k = ((c*n + t) * k) div k" by(rule arg_cong[of _ _ "%t. t div k"])
+  thus ?P by(simp add: zdiv_zmult_self1[OF not0[symmetric]])
+qed
+
+lemma ac_pi_eq: assumes gr0: "0 < (k::int)" shows "(~((0::int) < (c*n + t))) = (0 < ((-k)*c)*n + ((-k)*t + k))"
+proof -
+  have "(~ (0::int) < (c*n + t)) = (0<1-(c*n + t))" by arith
+  also have  "(1-(c*n + t)) = (-1*c)*n + (-t+1)" by(simp add: int_distrib zmult_ac)
+  also have "0<(-1*c)*n + (-t+1) = (0 < (k*(-1*c)*n) + (k*(-t+1)))" by(rule ac_lt_eq[of _ 0,OF gr0,simplified])
+  also have "(k*(-1*c)*n) + (k*(-t+1)) = ((-k)*c)*n + ((-k)*t + k)" by(simp add: int_distrib zmult_ac)
+  finally show ?thesis .
+qed
+
+lemma binminus_uminus_conv: "(a::int) - b = a + (-b)"
+by arith
+
+lemma  linearize_dvd: "(t::int) = t1 ==> (d dvd t) = (d dvd t1)"
+by simp
+
+lemma lf_lt: "(l::int) = ll ==> (r::int) = lr ==> (l < r) =(ll < lr)"
+by simp
+
+lemma lf_eq: "(l::int) = ll ==> (r::int) = lr ==> (l = r) =(ll = lr)"
+by simp
+
+lemma lf_dvd: "(l::int) = ll ==> (r::int) = lr ==> (l dvd r) =(ll dvd lr)"
+by simp
+
+(* Theorems for transforming predicates on nat to predicates on int*)
+
+theorem all_nat: "(\<forall>x::nat. P x) = (\<forall>x::int. 0 <= x \<longrightarrow> P (nat x))"
+  by (simp split add: split_nat)
+
+theorem ex_nat: "(\<exists>x::nat. P x) = (\<exists>x::int. 0 <= x \<and> P (nat x))"
+  apply (simp split add: split_nat)
+  apply (rule iffI)
+  apply (erule exE)
+  apply (rule_tac x = "int x" in exI)
+  apply simp
+  apply (erule exE)
+  apply (rule_tac x = "nat x" in exI)
+  apply (erule conjE)
+  apply (erule_tac x = "nat x" in allE)
+  apply simp
+  done
+
+theorem zdiff_int_split: "P (int (x - y)) =
+  ((y \<le> x \<longrightarrow> P (int x - int y)) \<and> (x < y \<longrightarrow> P 0))"
+  apply (case_tac "y \<le> x")
+  apply (simp_all add: zdiff_int)
+  done
+
+theorem zdvd_int: "(x dvd y) = (int x dvd int y)"
+  apply (simp only: dvd_def ex_nat int_int_eq [symmetric] zmult_int [symmetric]
+    nat_0_le cong add: conj_cong)
+  apply (rule iffI)
+  apply rules
+  apply (erule exE)
+  apply (case_tac "x=0")
+  apply (rule_tac x=0 in exI)
+  apply simp
+  apply (case_tac "0 \<le> k")
+  apply rules
+  apply (simp add: linorder_not_le)
+  apply (drule zmult_zless_mono2_neg [OF iffD2 [OF zero_less_int_conv]])
+  apply assumption
+  apply (simp add: zmult_ac)
+  done
+
+theorem number_of1: "(0::int) <= number_of n \<Longrightarrow> (0::int) <= number_of (n BIT b)"
+  by simp
+
+theorem number_of2: "(0::int) <= number_of bin.Pls" by simp
+
+theorem Suc_plus1: "Suc n = n + 1" by simp
+
+(* specific instances of congruence rules, to prevent simplifier from looping *)
+
+theorem imp_le_cong: "(0 <= x \<Longrightarrow> P = P') \<Longrightarrow> (0 <= (x::nat) \<longrightarrow> P) = (0 <= x \<longrightarrow> P')"
+  by simp
+
+theorem conj_le_cong: "(0 <= x \<Longrightarrow> P = P') \<Longrightarrow> (0 <= (x::nat) \<and> P) = (0 <= x \<and> P')"
+  by simp
+
+use "cooper_dec.ML"
+use "cooper_proof.ML"
+use "qelim.ML"
+use "presburger.ML"
+
+setup "Presburger.setup"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Presburger/cooper_dec.ML	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,773 @@
+(*  Title:      HOL/Integ/cooper_dec.ML
+    ID:         $Id$
+    Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+File containing the implementation of Cooper Algorithm
+decision procedure (intensively inspired from J.Harrison)
+*)
+ 
+signature COOPER_DEC = 
+sig
+  exception COOPER
+  val is_arith_rel : term -> bool
+  val mk_numeral : int -> term
+  val dest_numeral : term -> int
+  val zero : term
+  val one : term
+  val linear_cmul : int -> term -> term
+  val linear_add : string list -> term -> term -> term 
+  val linear_sub : string list -> term -> term -> term 
+  val linear_neg : term -> term
+  val lint : string list -> term -> term
+  val linform : string list -> term -> term
+  val formlcm : term -> term -> int
+  val adjustcoeff : term -> int -> term -> term
+  val unitycoeff : term -> term -> term
+  val divlcm : term -> term -> int
+  val bset : term -> term -> term list
+  val aset : term -> term -> term list
+  val linrep : string list -> term -> term -> term -> term
+  val list_disj : term list -> term
+  val simpl : term -> term
+  val fv : term -> string list
+  val negate : term -> term
+  val operations : (string * (int * int -> bool)) list
+end;
+
+structure  CooperDec : COOPER_DEC =
+struct
+
+(* ========================================================================= *) 
+(* Cooper's algorithm for Presburger arithmetic.                             *) 
+(* ========================================================================= *) 
+exception COOPER;
+
+(* ------------------------------------------------------------------------- *) 
+(* Lift operations up to numerals.                                           *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+(*Assumption : The construction of atomar formulas in linearl arithmetic is based on 
+relation operations of Type : [int,int]---> bool *) 
+ 
+(* ------------------------------------------------------------------------- *) 
+ 
+ 
+(*Function is_arith_rel returns true if and only if the term is an atomar presburger 
+formula *) 
+fun is_arith_rel tm = case tm of 
+	 Const(p,Type ("fun",[Type ("Numeral.bin", []),Type ("fun",[Type ("Numeral.bin", 
+	 []),Type ("bool",[])] )])) $ _ $_ => true 
+	|Const(p,Type ("fun",[Type ("IntDef.int", []),Type ("fun",[Type ("IntDef.int", 
+	 []),Type ("bool",[])] )])) $ _ $_ => true 
+	|_ => false; 
+ 
+(*Function is_arith_rel returns true if and only if the term is an operation of the 
+form [int,int]---> int*) 
+ 
+(*Transform a natural number to a term*) 
+ 
+fun mk_numeral 0 = Const("0",HOLogic.intT)
+   |mk_numeral 1 = Const("1",HOLogic.intT)
+   |mk_numeral n = (HOLogic.number_of_const HOLogic.intT) $ (HOLogic.mk_bin n); 
+ 
+(*Transform an Term to an natural number*)	  
+	  
+fun dest_numeral (Const("0",Type ("IntDef.int", []))) = 0
+   |dest_numeral (Const("1",Type ("IntDef.int", []))) = 1
+   |dest_numeral (Const ("Numeral.number_of",_) $ n)= HOLogic.dest_binum n; 
+(*Some terms often used for pattern matching*) 
+ 
+val zero = mk_numeral 0; 
+val one = mk_numeral 1; 
+ 
+(*Tests if a Term is representing a number*) 
+ 
+fun is_numeral t = (t = zero) orelse (t = one) orelse (can dest_numeral t); 
+ 
+(*maps a unary natural function on a term containing an natural number*) 
+ 
+fun numeral1 f n = mk_numeral (f(dest_numeral n)); 
+ 
+(*maps a binary natural function on 2 term containing  natural numbers*) 
+ 
+fun numeral2 f m n = mk_numeral(f(dest_numeral m) (dest_numeral n)); 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Operations on canonical linear terms c1 * x1 + ... + cn * xn + k          *) 
+(*                                                                           *) 
+(* Note that we're quite strict: the ci must be present even if 1            *) 
+(* (but if 0 we expect the monomial to be omitted) and k must be there       *) 
+(* even if it's zero. Thus, it's a constant iff not an addition term.        *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+ 
+fun linear_cmul n tm =  if n = 0 then zero else let fun times n k = n*k in  
+  ( case tm of  
+     (Const("op +",T)  $  (Const ("op *",T1 ) $c1 $  x1) $ rest) => 
+       Const("op +",T) $ ((Const("op *",T1) $ (numeral1 (times n) c1) $ x1)) $ (linear_cmul n rest) 
+    |_ =>  numeral1 (times n) tm) 
+    end ; 
+ 
+ 
+ 
+ 
+(* Whether the first of two items comes earlier in the list  *) 
+fun earlier [] x y = false 
+	|earlier (h::t) x y =if h = y then false 
+              else if h = x then true 
+              	else earlier t x y ; 
+ 
+fun earlierv vars (Bound i) (Bound j) = i < j 
+   |earlierv vars (Bound _) _ = true 
+   |earlierv vars _ (Bound _)  = false 
+   |earlierv vars (Free (x,_)) (Free (y,_)) = earlier vars x y; 
+ 
+ 
+fun linear_add vars tm1 tm2 = 
+  let fun addwith x y = x + y in
+ (case (tm1,tm2) of 
+	((Const ("op +",T1) $ ( Const("op *",T2) $ c1 $  x1) $ rest1),(Const 
+	("op +",T3)$( Const("op *",T4) $ c2 $  x2) $ rest2)) => 
+         if x1 = x2 then 
+              let val c = (numeral2 (addwith) c1 c2) 
+	      in 
+              if c = zero then (linear_add vars rest1  rest2)  
+	      else (Const("op +",T1) $ (Const("op *",T2) $ c $ x1) $ (linear_add vars  rest1 rest2)) 
+              end 
+	   else 
+		if earlierv vars x1 x2 then (Const("op +",T1) $  
+		(Const("op *",T2)$ c1 $ x1) $ (linear_add vars rest1 tm2)) 
+    	       else (Const("op +",T1) $ (Const("op *",T2) $ c2 $ x2) $ (linear_add vars tm1 rest2)) 
+   	|((Const("op +",T1) $ (Const("op *",T2) $ c1 $ x1) $ rest1) ,_) => 
+    	  (Const("op +",T1)$ (Const("op *",T2) $ c1 $ x1) $ (linear_add vars 
+	  rest1 tm2)) 
+   	|(_, (Const("op +",T1) $(Const("op *",T2) $ c2 $ x2) $ rest2)) => 
+      	  (Const("op +",T1) $ (Const("op *",T2) $ c2 $ x2) $ (linear_add vars tm1 
+	  rest2)) 
+   	| (_,_) => numeral2 (addwith) tm1 tm2) 
+	 
+	end; 
+ 
+(*To obtain the unary - applyed on a formula*) 
+ 
+fun linear_neg tm = linear_cmul (0 - 1) tm; 
+ 
+(*Substraction of two terms *) 
+ 
+fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2); 
+ 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Linearize a term.                                                         *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+(* linearises a term from the point of view of Variable Free (x,T). 
+After this fuction the all expressions containig ths variable will have the form  
+ c*Free(x,T) + t where c is a constant ant t is a Term which is not containing 
+ Free(x,T)*) 
+  
+fun lint vars tm = if is_numeral tm then tm else case tm of 
+   (Free (x,T)) =>  (HOLogic.mk_binop "op +" ((HOLogic.mk_binop "op *" ((mk_numeral 1),Free (x,T))), zero)) 
+  |(Bound i) =>  (Const("op +",HOLogic.intT -->HOLogic.intT -->HOLogic.intT) $ 
+  (Const("op *",HOLogic.intT -->HOLogic.intT -->HOLogic.intT) $ (mk_numeral 1) $ (Bound i)) $ zero) 
+  |(Const("uminus",_) $ t ) => (linear_neg (lint vars t)) 
+  |(Const("op +",_) $ s $ t) => (linear_add vars (lint vars s) (lint vars t)) 
+  |(Const("op -",_) $ s $ t) => (linear_sub vars (lint vars s) (lint vars t)) 
+  |(Const ("op *",_) $ s $ t) => 
+        let val s' = lint vars s  
+            val t' = lint vars t  
+        in 
+        if is_numeral s' then (linear_cmul (dest_numeral s') t') 
+        else if is_numeral t' then (linear_cmul (dest_numeral t') s') 
+ 
+         else (warning "lint: apparent nonlinearity"; raise COOPER)
+         end 
+  |_ =>   error "lint: unknown term"; 
+   
+ 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Linearize the atoms in a formula, and eliminate non-strict inequalities.  *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun mkatom vars p t = Const(p,HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ zero $ (lint vars t); 
+ 
+fun linform vars (Const ("Divides.op dvd",_) $ c $ t) =  
+      let val c' = (mk_numeral(abs(dest_numeral c)))  
+      in (HOLogic.mk_binrel "Divides.op dvd" (c,lint vars t)) 
+      end 
+  |linform vars  (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ s $ t ) = (mkatom vars "op =" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ t $ s) ) 
+  |linform vars  (Const("op <",_)$ s $t ) = (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ t $ s))
+  |linform vars  (Const("op >",_) $ s $ t ) = (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ s $ t)) 
+  |linform vars  (Const("op <=",_)$ s $ t ) = 
+        (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $ (Const("op +",HOLogic.intT --> HOLogic.intT --> HOLogic.intT) $t $(mk_numeral 1)) $ s)) 
+  |linform vars  (Const("op >=",_)$ s $ t ) = 
+        (mkatom vars "op <" (Const ("op -",HOLogic.intT --> HOLogic.intT --> 
+	HOLogic.intT) $ (Const("op +",HOLogic.intT --> HOLogic.intT --> 
+	HOLogic.intT) $s $(mk_numeral 1)) $ t)) 
+ 
+   |linform vars  fm =  fm; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Post-NNF transformation eliminating negated inequalities.                 *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun posineq fm = case fm of  
+ (Const ("Not",_)$(Const("op <",_)$ c $ t)) =>
+   (HOLogic.mk_binrel "op <"  (zero , (linear_sub [] (mk_numeral 1) (linear_add [] c t ) ))) 
+  | ( Const ("op &",_) $ p $ q)  => HOLogic.mk_conj (posineq p,posineq q)
+  | ( Const ("op |",_) $ p $ q ) => HOLogic.mk_disj (posineq p,posineq q)
+  | _ => fm; 
+  
+
+(* ------------------------------------------------------------------------- *) 
+(* Find the LCM of the coefficients of x.                                    *) 
+(* ------------------------------------------------------------------------- *) 
+(*gcd calculates gcd (a,b) and helps lcm_num calculating lcm (a,b)*) 
+ 
+fun gcd a b = if a=0 then b else gcd (b mod a) a ; 
+fun lcm_num a b = (abs a*b) div (gcd (abs a) (abs b)); 
+ 
+fun formlcm x fm = case fm of 
+    (Const (p,_)$ _ $(Const ("op +", _)$(Const ("op *",_)$ c $ y ) $z ) ) =>  if 
+    (is_arith_rel fm) andalso (x = y) then abs(dest_numeral c) else 1 
+  | ( Const ("Not", _) $p) => formlcm x p 
+  | ( Const ("op &",_) $ p $ q) => lcm_num (formlcm x p) (formlcm x q) 
+  | ( Const ("op |",_) $ p $ q )=> lcm_num (formlcm x p) (formlcm x q) 
+  |  _ => 1; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Adjust all coefficients of x in formula; fold in reduction to +/- 1.      *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun adjustcoeff x l fm = 
+     case fm of  
+      (Const(p,_) $d $( Const ("op +", _)$(Const ("op *",_) $ 
+      c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then  
+        let val m = l div (dest_numeral c) 
+            val n = (if p = "op <" then abs(m) else m) 
+            val xtm = HOLogic.mk_binop "op *" ((mk_numeral (m div n)), x) 
+	in
+        (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
+	end 
+	else fm 
+  |( Const ("Not", _) $ p) => HOLogic.Not $ (adjustcoeff x l p) 
+  |( Const ("op &",_) $ p $ q) => HOLogic.conj$(adjustcoeff x l p) $(adjustcoeff x l q) 
+  |( Const ("op |",_) $ p $ q) => HOLogic.disj $(adjustcoeff x l p)$ (adjustcoeff x l q) 
+  |_ => fm; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Hence make coefficient of x one in existential formula.                   *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun unitycoeff x fm = 
+  let val l = formlcm x fm 
+      val fm' = adjustcoeff x l fm in
+     if l = 1 then fm' else 
+     let val xp = (HOLogic.mk_binop "op +"  
+     		((HOLogic.mk_binop "op *" ((mk_numeral 1), x )), zero)) in 
+      HOLogic.conj $(HOLogic.mk_binrel "Divides.op dvd" ((mk_numeral l) , xp )) $ (adjustcoeff x l fm) 
+      end 
+  end; 
+ 
+(* adjustcoeffeq l fm adjusts the coeffitients c_i of x  overall in fm to l*)
+(* Here l must be a multiple of all c_i otherwise the obtained formula is not equivalent*)
+(*
+fun adjustcoeffeq x l fm = 
+    case fm of  
+      (Const(p,_) $d $( Const ("op +", _)$(Const ("op *",_) $ 
+      c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then  
+        let val m = l div (dest_numeral c) 
+            val n = (if p = "op <" then abs(m) else m)  
+            val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x))
+            in (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
+	    end 
+	else fm 
+  |( Const ("Not", _) $ p) => HOLogic.Not $ (adjustcoeffeq x l p) 
+  |( Const ("op &",_) $ p $ q) => HOLogic.conj$(adjustcoeffeq x l p) $(adjustcoeffeq x l q) 
+  |( Const ("op |",_) $ p $ q) => HOLogic.disj $(adjustcoeffeq x l p)$ (adjustcoeffeq x l q) 
+  |_ => fm;
+ 
+
+*)
+
+(* ------------------------------------------------------------------------- *) 
+(* The "minus infinity" version.                                             *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun minusinf x fm = case fm of  
+    (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ (c1 ) $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) => 
+  	 if (is_arith_rel fm) andalso (x=y) andalso (c2 = one) andalso (c1 =zero) then HOLogic.false_const  
+	 				 else fm 
+ 
+  |(Const("op <",_) $ c $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z 
+  )) => 
+        if (x =y) andalso (pm1 = one) andalso (c = zero) then HOLogic.false_const else HOLogic.true_const 
+	 
+  |(Const ("Not", _) $ p) => HOLogic.Not $ (minusinf x p) 
+  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (minusinf x p) $ (minusinf x q) 
+  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (minusinf x p) $ (minusinf x q) 
+  |_ => fm; 
+
+(* ------------------------------------------------------------------------- *)
+(* The "Plus infinity" version.                                             *)
+(* ------------------------------------------------------------------------- *)
+
+fun plusinf x fm = case fm of
+    (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ (c1 ) $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	 if (is_arith_rel fm) andalso (x=y) andalso (c2 = one) andalso (c1 =zero) then HOLogic.false_const
+	 				 else fm
+
+  |(Const("op <",_) $ c $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z
+  )) =>
+        if (x =y) andalso (pm1 = one) andalso (c = zero) then HOLogic.true_const else HOLogic.false_const
+
+  |(Const ("Not", _) $ p) => HOLogic.Not $ (plusinf x p)
+  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (plusinf x p) $ (plusinf x q)
+  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (plusinf x p) $ (plusinf x q)
+  |_ => fm;
+ 
+(* ------------------------------------------------------------------------- *) 
+(* The LCM of all the divisors that involve x.                               *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun divlcm x (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z ) ) =  
+        if x = y then abs(dest_numeral d) else 1 
+  |divlcm x ( Const ("Not", _) $ p) = divlcm x p 
+  |divlcm x ( Const ("op &",_) $ p $ q) = lcm_num (divlcm x p) (divlcm x q) 
+  |divlcm x ( Const ("op |",_) $ p $ q ) = lcm_num (divlcm x p) (divlcm x q) 
+  |divlcm x  _ = 1; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Construct the B-set.                                                      *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun bset x fm = case fm of 
+   (Const ("Not", _) $ p) => if (is_arith_rel p) then  
+          (case p of  
+	      (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $c2 $y) $a ) )  
+	             => if (is_arith_rel p) andalso (x=	y) andalso (c2 = one) andalso (c1 = zero)  
+	                then [linear_neg a] 
+			else  bset x p 
+   	  |_ =>[]) 
+			 
+			else bset x p 
+  |(Const ("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +",_) $ (Const ("op *",_) $c2 $ x) $ a)) =>  if (c1 =zero) andalso (c2 = one) then [linear_neg(linear_add [] a (mk_numeral 1))]  else [] 
+  |(Const ("op <",_) $ c1$ (Const ("op +",_) $(Const ("op *",_)$ c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = one) then [linear_neg a] else [] 
+  |(Const ("op &",_) $ p $ q) => (bset x p) union (bset x q) 
+  |(Const ("op |",_) $ p $ q) => (bset x p) union (bset x q) 
+  |_ => []; 
+ 
+(* ------------------------------------------------------------------------- *)
+(* Construct the A-set.                                                      *)
+(* ------------------------------------------------------------------------- *)
+
+fun aset x fm = case fm of
+   (Const ("Not", _) $ p) => if (is_arith_rel p) then
+          (case p of
+	      (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $c2 $y) $a ) )
+	             => if (x=	y) andalso (c2 = one) andalso (c1 = zero)
+	                then [linear_neg a]
+			else  []
+   	  |_ =>[])
+
+			else aset x p
+  |(Const ("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +",_) $ (Const ("op *",_) $c2 $ x) $ a)) =>  if (c1 =zero) andalso (c2 = one) then [linear_sub [] (mk_numeral 1) a]  else []
+  |(Const ("op <",_) $ c1$ (Const ("op +",_) $(Const ("op *",_)$ c2 $ x) $ a)) => if (c1 =zero) andalso (c2 = (mk_numeral (~1))) then [a] else []
+  |(Const ("op &",_) $ p $ q) => (aset x p) union (aset x q)
+  |(Const ("op |",_) $ p $ q) => (aset x p) union (aset x q)
+  |_ => [];
+
+
+(* ------------------------------------------------------------------------- *) 
+(* Replace top variable with another linear form, retaining canonicality.    *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun linrep vars x t fm = case fm of  
+   ((Const(p,_)$ d $ (Const("op +",_)$(Const("op *",_)$ c $ y) $ z))) => 
+      if (x = y) andalso (is_arith_rel fm)  
+      then  
+        let val ct = linear_cmul (dest_numeral c) t  
+	in (HOLogic.mk_binrel p (d, linear_add vars ct z)) 
+	end 
+	else fm 
+  |(Const ("Not", _) $ p) => HOLogic.Not $ (linrep vars x t p) 
+  |(Const ("op &",_) $ p $ q) => HOLogic.conj $ (linrep vars x t p) $ (linrep vars x t q) 
+  |(Const ("op |",_) $ p $ q) => HOLogic.disj $ (linrep vars x t p) $ (linrep vars x t q) 
+  |_ => fm; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Evaluation of constant expressions.                                       *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+val operations = 
+  [("op =",op=), ("op <",op<), ("op >",op>), ("op <=",op<=) , ("op >=",op>=), 
+   ("Divides.op dvd",fn (x,y) =>((y mod x) = 0))]; 
+ 
+fun applyoperation (Some f) (a,b) = f (a, b) 
+    |applyoperation _ (_, _) = false; 
+ 
+(*Evaluation of constant atomic formulas*) 
+ 
+fun evalc_atom at = case at of  
+      (Const (p,_) $ s $ t) =>(  
+         case assoc (operations,p) of 
+             Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then HOLogic.true_const else HOLogic.false_const)  
+              handle _ => at) 
+             | _ =>  at) 
+     |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
+         case assoc (operations,p) of 
+             Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then 
+	     HOLogic.false_const else HOLogic.true_const)  
+              handle _ => at) 
+             | _ =>  at) 
+     | _ =>  at; 
+ 
+(*Function onatoms apllys function f on the atomic formulas involved in a.*) 
+ 
+fun onatoms f a = if (is_arith_rel a) then f a else case a of 
+ 
+  	(Const ("Not",_) $ p) => if is_arith_rel p then HOLogic.Not $ (f p) 
+				 
+				else HOLogic.Not $ (onatoms f p) 
+  	|(Const ("op &",_) $ p $ q) => HOLogic.conj $ (onatoms f p) $ (onatoms f q) 
+  	|(Const ("op |",_) $ p $ q) => HOLogic.disj $ (onatoms f p) $ (onatoms f q) 
+  	|(Const ("op -->",_) $ p $ q) => HOLogic.imp $ (onatoms f p) $ (onatoms f q) 
+  	|((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q) => (Const ("op =", [HOLogic.boolT, HOLogic.boolT] ---> HOLogic.boolT)) $ (onatoms f p) $ (onatoms f q) 
+  	|(Const("All",_) $ Abs(x,T,p)) => Const("All", [HOLogic.intT --> 
+	HOLogic.boolT] ---> HOLogic.boolT)$ Abs (x ,T, (onatoms f p)) 
+  	|(Const("Ex",_) $ Abs(x,T,p)) => Const("Ex", [HOLogic.intT --> HOLogic.boolT]---> HOLogic.boolT) $ Abs( x ,T, (onatoms f p)) 
+  	|_ => a; 
+ 
+val evalc = onatoms evalc_atom; 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Hence overall quantifier elimination.                                     *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+(*Applyes a function iteratively on the list*) 
+ 
+fun end_itlist f []     = error "end_itlist" 
+   |end_itlist f [x]    = x 
+   |end_itlist f (h::t) = f h (end_itlist f t); 
+ 
+ 
+(*list_disj[conj] makes a disj[conj] of a given list. used with conjucts or disjuncts 
+it liearises iterated conj[disj]unctions. *) 
+ 
+fun disj_help p q = HOLogic.disj $ p $ q ; 
+ 
+fun list_disj l = 
+  if l = [] then HOLogic.false_const else end_itlist disj_help l; 
+   
+fun conj_help p q = HOLogic.conj $ p $ q ; 
+ 
+fun list_conj l = 
+  if l = [] then HOLogic.true_const else end_itlist conj_help l; 
+   
+(*Simplification of Formulas *) 
+ 
+(*Function q_bnd_chk checks if a quantified Formula makes sens : Means if in 
+the body of the existential quantifier there are bound variables to the 
+existential quantifier.*) 
+ 
+fun has_bound fm =let fun has_boundh fm i = case fm of 
+		 Bound n => (i = n) 
+		 |Abs (_,_,p) => has_boundh p (i+1) 
+		 |t1 $ t2 => (has_boundh t1 i) orelse (has_boundh t2 i) 
+		 |_ =>false
+
+in  case fm of 
+	Bound _ => true 
+       |Abs (_,_,p) => has_boundh p 0 
+       |t1 $ t2 => (has_bound t1 ) orelse (has_bound t2 ) 
+       |_ =>false
+end;
+ 
+(*has_sub_abs checks if in a given Formula there are subformulas which are quantifed 
+too. Is no used no more.*) 
+ 
+fun has_sub_abs fm = case fm of  
+		 Abs (_,_,_) => true 
+		 |t1 $ t2 => (has_bound t1 ) orelse (has_bound t2 ) 
+		 |_ =>false ; 
+		  
+(*update_bounds called with i=0 udates the numeration of bounded variables because the 
+formula will not be quantified any more.*) 
+ 
+fun update_bounds fm i = case fm of 
+		 Bound n => if n >= i then Bound (n-1) else fm 
+		 |Abs (x,T,p) => Abs(x,T,(update_bounds p (i+1))) 
+		 |t1 $ t2 => (update_bounds t1 i) $ (update_bounds t2 i) 
+		 |_ => fm ; 
+ 
+(*psimpl : Simplification of propositions (general purpose)*) 
+fun psimpl1 fm = case fm of 
+    Const("Not",_) $ Const ("False",_) => HOLogic.true_const 
+  | Const("Not",_) $ Const ("True",_) => HOLogic.false_const 
+  | Const("op &",_) $ Const ("False",_) $ q => HOLogic.false_const 
+  | Const("op &",_) $ p $ Const ("False",_)  => HOLogic.false_const 
+  | Const("op &",_) $ Const ("True",_) $ q => q 
+  | Const("op &",_) $ p $ Const ("True",_) => p 
+  | Const("op |",_) $ Const ("False",_) $ q => q 
+  | Const("op |",_) $ p $ Const ("False",_)  => p 
+  | Const("op |",_) $ Const ("True",_) $ q => HOLogic.true_const 
+  | Const("op |",_) $ p $ Const ("True",_)  => HOLogic.true_const 
+  | Const("op -->",_) $ Const ("False",_) $ q => HOLogic.true_const 
+  | Const("op -->",_) $ Const ("True",_) $  q => q 
+  | Const("op -->",_) $ p $ Const ("True",_)  => HOLogic.true_const 
+  | Const("op -->",_) $ p $ Const ("False",_)  => HOLogic.Not $  p 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ Const ("True",_) $ q => q 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ Const ("True",_) => p 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ Const ("False",_) $ q => HOLogic.Not $  q 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ Const ("False",_)  => HOLogic.Not $  p 
+  | _ => fm; 
+ 
+fun psimpl fm = case fm of 
+   Const ("Not",_) $ p => psimpl1 (HOLogic.Not $ (psimpl p)) 
+  | Const("op &",_) $ p $ q => psimpl1 (HOLogic.mk_conj (psimpl p,psimpl q)) 
+  | Const("op |",_) $ p $ q => psimpl1 (HOLogic.mk_disj (psimpl p,psimpl q)) 
+  | Const("op -->",_) $ p $ q => psimpl1 (HOLogic.mk_imp(psimpl p,psimpl q)) 
+  | Const("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q => psimpl1 (HOLogic.mk_eq(psimpl p,psimpl q)) 
+  | _ => fm; 
+ 
+ 
+(*simpl : Simplification of Terms involving quantifiers too. 
+ This function is able to drop out some quantified expressions where there are no 
+ bound varaibles.*) 
+  
+fun simpl1 fm  = 
+  case fm of 
+    Const("All",_) $Abs(x,_,p) => if (has_bound fm ) then fm  
+    				else (update_bounds p 0) 
+  | Const("Ex",_) $ Abs (x,_,p) => if has_bound fm then fm  
+    				else (update_bounds p 0) 
+  | _ => psimpl1 fm; 
+ 
+fun simpl fm = case fm of 
+    Const ("Not",_) $ p => simpl1 (HOLogic.Not $(simpl p))  
+  | Const ("op &",_) $ p $ q => simpl1 (HOLogic.mk_conj (simpl p ,simpl q))  
+  | Const ("op |",_) $ p $ q => simpl1 (HOLogic.mk_disj (simpl p ,simpl q ))  
+  | Const ("op -->",_) $ p $ q => simpl1 (HOLogic.mk_imp(simpl p ,simpl q ))  
+  | Const("op =", Type ("fun",[Type ("bool", []),_]))$ p $ q => simpl1 
+  (HOLogic.mk_eq(simpl p ,simpl q ))  
+  | Const ("All",Ta) $ Abs(Vn,VT,p) => simpl1(Const("All",Ta) $ 
+  Abs(Vn,VT,simpl p ))  
+  | Const ("Ex",Ta)  $ Abs(Vn,VT,p) => simpl1(Const("Ex",Ta)  $ 
+  Abs(Vn,VT,simpl p ))  
+  | _ => fm; 
+ 
+(* ------------------------------------------------------------------------- *) 
+ 
+(* Puts fm into NNF*) 
+ 
+fun  nnf fm = if (is_arith_rel fm) then fm  
+else (case fm of 
+  ( Const ("op &",_) $ p $ q)  => HOLogic.conj $ (nnf p) $(nnf q) 
+  | (Const("op |",_) $ p $q) => HOLogic.disj $ (nnf p)$(nnf q) 
+  | (Const ("op -->",_)  $ p $ q) => HOLogic.disj $ (nnf (HOLogic.Not $ p)) $ (nnf q) 
+  | ((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q) =>(HOLogic.disj $ (HOLogic.conj $ (nnf p) $ (nnf q)) $ (HOLogic.conj $ (nnf (HOLogic.Not $ p) ) $ (nnf(HOLogic.Not $ q)))) 
+  | (Const ("Not",_)) $ ((Const ("Not",_)) $ p) => (nnf p) 
+  | (Const ("Not",_)) $ (( Const ("op &",_)) $ p $ q) =>HOLogic.disj $ (nnf(HOLogic.Not $ p)) $ (nnf(HOLogic.Not $q)) 
+  | (Const ("Not",_)) $ (( Const ("op |",_)) $ p $ q) =>HOLogic.conj $ (nnf(HOLogic.Not $ p)) $ (nnf(HOLogic.Not $ q)) 
+  | (Const ("Not",_)) $ (( Const ("op -->",_)) $ p $ q ) =>HOLogic.conj $ (nnf p) $(nnf(HOLogic.Not $ q)) 
+  | (Const ("Not",_)) $ ((Const ("op =", Type ("fun",[Type ("bool", []),_]))) $ p $ q ) =>(HOLogic.disj $ (HOLogic.conj $(nnf p) $ (nnf(HOLogic.Not $ q))) $ (HOLogic.conj $(nnf(HOLogic.Not $ p)) $ (nnf q))) 
+  | _ => fm); 
+ 
+ 
+(* Function remred to remove redundancy in a list while keeping the order of appearance of the 
+elements. but VERY INEFFICIENT!! *) 
+ 
+fun remred1 el [] = [] 
+    |remred1 el (h::t) = if el=h then (remred1 el t) else h::(remred1 el t); 
+     
+fun remred [] = [] 
+    |remred (x::l) =  x::(remred1 x (remred l)); 
+ 
+(*Makes sure that all free Variables are of the type integer but this function is only 
+used temporarily, this job must be done by the parser later on.*) 
+ 
+fun mk_uni_vars T  (node $ rest) = (case node of 
+    Free (name,_) => Free (name,T) $ (mk_uni_vars T rest) 
+    |_=> (mk_uni_vars T node) $ (mk_uni_vars T rest )  ) 
+    |mk_uni_vars T (Free (v,_)) = Free (v,T) 
+    |mk_uni_vars T tm = tm; 
+ 
+fun mk_uni_int T (Const ("0",T2)) = if T = T2 then (mk_numeral 0) else (Const ("0",T2)) 
+    |mk_uni_int T (Const ("1",T2)) = if T = T2 then (mk_numeral 1) else (Const ("1",T2)) 
+    |mk_uni_int T (node $ rest) = (mk_uni_int T node) $ (mk_uni_int T rest )  
+    |mk_uni_int T (Abs(AV,AT,p)) = Abs(AV,AT,mk_uni_int T p) 
+    |mk_uni_int T tm = tm; 
+ 
+
+(* Minusinfinity Version*) 
+fun coopermi vars1 fm = 
+  case fm of 
+   Const ("Ex",_) $ Abs(x0,T,p0) => let 
+    val (xn,p1) = variant_abs (x0,T,p0) 
+    val x = Free (xn,T)  
+    val vars = (xn::vars1) 
+    val p = unitycoeff x  (posineq (simpl p1))
+    val p_inf = simpl (minusinf x p) 
+    val bset = bset x p 
+    val js = 1 upto divlcm x p  
+    fun p_element j b = linrep vars x (linear_add vars b (mk_numeral j)) p  
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) bset)  
+   in (list_disj (map stage js))
+    end 
+  | _ => error "cooper: not an existential formula"; 
+ 
+
+
+(* The plusinfinity version of cooper*)
+fun cooperpi vars1 fm =
+  case fm of
+   Const ("Ex",_) $ Abs(x0,T,p0) => let 
+    val (xn,p1) = variant_abs (x0,T,p0)
+    val x = Free (xn,T)
+    val vars = (xn::vars1)
+    val p = unitycoeff x  (posineq (simpl p1))
+    val p_inf = simpl (plusinf x p)
+    val aset = aset x p
+    val js = 1 upto divlcm x p
+    fun p_element j a = linrep vars x (linear_sub vars a (mk_numeral j)) p
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) aset)
+   in (list_disj (map stage js))
+   end
+  | _ => error "cooper: not an existential formula";
+  
+
+
+(*Cooper main procedure*) 
+  
+fun cooper vars1 fm =
+  case fm of
+   Const ("Ex",_) $ Abs(x0,T,p0) => let 
+    val (xn,p1) = variant_abs (x0,T,p0)
+    val x = Free (xn,T)
+    val vars = (xn::vars1)
+    val p = unitycoeff x  (posineq (simpl p1))
+    val ast = aset x p
+    val bst = bset x p
+    val js = 1 upto divlcm x p
+    val (p_inf,f,S ) = 
+    if (length bst) < (length ast) 
+     then (minusinf x p,linear_add,bst)
+     else (plusinf x p, linear_sub,ast)
+    fun p_element j a = linrep vars x (f vars a (mk_numeral j)) p
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) S)
+   in (list_disj (map stage js))
+   end
+  | _ => error "cooper: not an existential formula";
+
+
+
+ 
+(*Function itlist applys a double parametred function f : 'a->'b->b iteratively to a List l : 'a 
+list With End condition b. ict calculates f(e1,f(f(e2,f(e3,...(...f(en,b))..))))) 
+ assuming l = [e1,e2,...,en]*) 
+ 
+fun itlist f l b = case l of 
+    [] => b 
+  | (h::t) => f h (itlist f t b); 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Free variables in terms and formulas.	                             *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+fun fvt tml = case tml of 
+    [] => [] 
+  | Free(x,_)::r => x::(fvt r) 
+ 
+fun fv fm = fvt (term_frees fm); 
+ 
+ 
+(* ========================================================================= *) 
+(* Quantifier elimination.                                                   *) 
+(* ========================================================================= *) 
+(*conj[/disj]uncts lists iterated conj[disj]unctions*) 
+ 
+fun disjuncts fm = case fm of 
+    Const ("op |",_) $ p $ q => (disjuncts p) @ (disjuncts q) 
+  | _ => [fm]; 
+ 
+fun conjuncts fm = case fm of 
+    Const ("op &",_) $p $ q => (conjuncts p) @ (conjuncts q) 
+  | _ => [fm]; 
+ 
+ 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Lift procedure given literal modifier, formula normalizer & basic quelim. *) 
+(* ------------------------------------------------------------------------- *) 
+   
+fun lift_qelim afn nfn qfn isat = 
+ let   fun qelim x vars p = 
+  let val cjs = conjuncts p 
+      val (ycjs,ncjs) = partition (has_bound) cjs in 
+      (if ycjs = [] then p else 
+                          let val q = (qfn vars ((HOLogic.exists_const HOLogic.intT 
+			  ) $ Abs(x,HOLogic.intT,(list_conj ycjs)))) in 
+                          (itlist conj_help ncjs q)  
+			  end) 
+       end 
+    
+  fun qelift vars fm = if (isat fm) then afn vars fm 
+    else  
+    case fm of 
+      Const ("Not",_) $ p => HOLogic.Not $ (qelift vars p) 
+    | Const ("op &",_) $ p $q => HOLogic.conj $ (qelift vars p) $ (qelift vars q) 
+    | Const ("op |",_) $ p $ q => HOLogic.disj $ (qelift vars p) $ (qelift vars q) 
+    | Const ("op -->",_) $ p $ q => HOLogic.imp $ (qelift vars p) $ (qelift vars q) 
+    | Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q => HOLogic.mk_eq ((qelift vars p),(qelift vars q)) 
+    | Const ("All",QT) $ Abs(x,T,p) => HOLogic.Not $(qelift vars (Const ("Ex",QT) $ Abs(x,T,(HOLogic.Not $ p)))) 
+    | Const ("Ex",_) $ Abs (x,T,p)  => let  val djs = disjuncts(nfn(qelift (x::vars) p)) in 
+    			list_disj(map (qelim x vars) djs) end 
+    | _ => fm 
+ 
+  in (fn fm => simpl(qelift (fv fm) fm)) 
+  end; 
+ 
+ 
+(* ------------------------------------------------------------------------- *) 
+(* Cleverer (proposisional) NNF with conditional and literal modification.   *) 
+(* ------------------------------------------------------------------------- *) 
+ 
+(*Function Negate used by cnnf, negates a formula p*) 
+ 
+fun negate (Const ("Not",_) $ p) = p 
+    |negate p = (HOLogic.Not $ p); 
+ 
+fun cnnf lfn = 
+  let fun cnnfh fm = case  fm of 
+      (Const ("op &",_) $ p $ q) => HOLogic.mk_conj(cnnfh p,cnnfh q) 
+    | (Const ("op |",_) $ p $ q) => HOLogic.mk_disj(cnnfh p,cnnfh q) 
+    | (Const ("op -->",_) $ p $q) => HOLogic.mk_disj(cnnfh(HOLogic.Not $ p),cnnfh q) 
+    | (Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q) => HOLogic.mk_disj( 
+    		HOLogic.mk_conj(cnnfh p,cnnfh q), 
+		HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $q))) 
+
+    | (Const ("Not",_) $ (Const("Not",_) $ p)) => cnnfh p 
+    | (Const ("Not",_) $ (Const ("op &",_) $ p $ q)) => HOLogic.mk_disj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $ q)) 
+    | (Const ("Not",_) $(Const ("op |",_) $ (Const ("op &",_) $ p $ q) $  
+    			(Const ("op &",_) $ p1 $ r))) => if p1 = negate p then 
+		         HOLogic.mk_disj(  
+			   cnnfh (HOLogic.mk_conj(p,cnnfh(HOLogic.Not $ q))), 
+			   cnnfh (HOLogic.mk_conj(p1,cnnfh(HOLogic.Not $ r)))) 
+			 else  HOLogic.mk_conj(
+			  cnnfh (HOLogic.mk_disj(cnnfh (HOLogic.Not $ p),cnnfh(HOLogic.Not $ q))), 
+			   cnnfh (HOLogic.mk_disj(cnnfh (HOLogic.Not $ p1),cnnfh(HOLogic.Not $ r)))
+			 ) 
+    | (Const ("Not",_) $ (Const ("op |",_) $ p $ q)) => HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh(HOLogic.Not $ q)) 
+    | (Const ("Not",_) $ (Const ("op -->",_) $ p $q)) => HOLogic.mk_conj(cnnfh p,cnnfh(HOLogic.Not $ q)) 
+    | (Const ("Not",_) $ (Const ("op =",Type ("fun",[Type ("bool", []),_]))  $ p $ q)) => HOLogic.mk_disj(HOLogic.mk_conj(cnnfh p,cnnfh(HOLogic.Not $ q)),HOLogic.mk_conj(cnnfh(HOLogic.Not $ p),cnnfh q)) 
+    | _ => lfn fm  
+  in cnnfh o simpl
+  end; 
+ 
+(*End- function the quantifierelimination an decion procedure of presburger formulas.*)   
+val integer_qelim = simpl o evalc o (lift_qelim linform (simpl o (cnnf posineq o evalc)) cooper is_arith_rel) ; 
+
+end;
+ 
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Presburger/cooper_proof.ML	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,1488 @@
+(*  Title:      HOL/Integ/cooper_proof.ML
+    ID:         $Id$
+    Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+File containing the implementation of the proof
+generation for Cooper Algorithm
+*)
+
+signature COOPER_PROOF =
+sig
+  val qe_Not : thm
+  val qe_conjI : thm
+  val qe_disjI : thm
+  val qe_impI : thm
+  val qe_eqI : thm
+  val qe_exI : thm
+  val qe_get_terms : thm -> term * term
+  val cooper_prv : Sign.sg -> term -> term -> string list -> thm
+  val proof_of_evalc : Sign.sg -> term -> thm
+  val proof_of_cnnf : Sign.sg -> term -> (term -> thm) -> thm
+  val proof_of_linform : Sign.sg -> string list -> term -> thm
+end;
+
+structure CooperProof : COOPER_PROOF =
+struct
+
+open CooperDec;
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---         Protocoling part                                  ---*)
+(*---                                                           ---*)
+(*---           includes the protocolling datastructure         ---*)
+(*---                                                           ---*)
+(*---          and the protocolling fuctions                    ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+val presburger_ss = simpset_of (theory "Presburger")
+  addsimps [zdiff_def] delsimps [symmetric zdiff_def];
+val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT;
+
+(*Theorems that will be used later for the proofgeneration*)
+
+val zdvd_iff_zmod_eq_0 = thm "zdvd_iff_zmod_eq_0";
+val unity_coeff_ex = thm "unity_coeff_ex";
+
+(* Thorems for proving the adjustment of the coeffitients*)
+
+val ac_lt_eq =  thm "ac_lt_eq";
+val ac_eq_eq = thm "ac_eq_eq";
+val ac_dvd_eq = thm "ac_dvd_eq";
+val ac_pi_eq = thm "ac_pi_eq";
+
+(* The logical compination of the sythetised properties*)
+val qe_Not = thm "qe_Not";
+val qe_conjI = thm "qe_conjI";
+val qe_disjI = thm "qe_disjI";
+val qe_impI = thm "qe_impI";
+val qe_eqI = thm "qe_eqI";
+val qe_exI = thm "qe_exI";
+val qe_ALLI = thm "qe_ALLI";
+
+(*Modulo D property for Pminusinf an Plusinf *)
+val fm_modd_minf = thm "fm_modd_minf";
+val not_dvd_modd_minf = thm "not_dvd_modd_minf";
+val dvd_modd_minf = thm "dvd_modd_minf";
+
+val fm_modd_pinf = thm "fm_modd_pinf";
+val not_dvd_modd_pinf = thm "not_dvd_modd_pinf";
+val dvd_modd_pinf = thm "dvd_modd_pinf";
+
+(* the minusinfinity proprty*)
+
+val fm_eq_minf = thm "fm_eq_minf";
+val neq_eq_minf = thm "neq_eq_minf";
+val eq_eq_minf = thm "eq_eq_minf";
+val le_eq_minf = thm "le_eq_minf";
+val len_eq_minf = thm "len_eq_minf";
+val not_dvd_eq_minf = thm "not_dvd_eq_minf";
+val dvd_eq_minf = thm "dvd_eq_minf";
+
+(* the Plusinfinity proprty*)
+
+val fm_eq_pinf = thm "fm_eq_pinf";
+val neq_eq_pinf = thm "neq_eq_pinf";
+val eq_eq_pinf = thm "eq_eq_pinf";
+val le_eq_pinf = thm "le_eq_pinf";
+val len_eq_pinf = thm "len_eq_pinf";
+val not_dvd_eq_pinf = thm "not_dvd_eq_pinf";
+val dvd_eq_pinf = thm "dvd_eq_pinf";
+
+(*Logical construction of the Property*)
+val eq_minf_conjI = thm "eq_minf_conjI";
+val eq_minf_disjI = thm "eq_minf_disjI";
+val modd_minf_disjI = thm "modd_minf_disjI";
+val modd_minf_conjI = thm "modd_minf_conjI";
+
+val eq_pinf_conjI = thm "eq_pinf_conjI";
+val eq_pinf_disjI = thm "eq_pinf_disjI";
+val modd_pinf_disjI = thm "modd_pinf_disjI";
+val modd_pinf_conjI = thm "modd_pinf_conjI";
+
+(*A/B - set Theorem *)
+
+val bst_thm = thm "bst_thm";
+val ast_thm = thm "ast_thm";
+
+(*Cooper Backwards...*)
+(*Bset*)
+val not_bst_p_fm = thm "not_bst_p_fm";
+val not_bst_p_ne = thm "not_bst_p_ne";
+val not_bst_p_eq = thm "not_bst_p_eq";
+val not_bst_p_gt = thm "not_bst_p_gt";
+val not_bst_p_lt = thm "not_bst_p_lt";
+val not_bst_p_ndvd = thm "not_bst_p_ndvd";
+val not_bst_p_dvd = thm "not_bst_p_dvd";
+
+(*Aset*)
+val not_ast_p_fm = thm "not_ast_p_fm";
+val not_ast_p_ne = thm "not_ast_p_ne";
+val not_ast_p_eq = thm "not_ast_p_eq";
+val not_ast_p_gt = thm "not_ast_p_gt";
+val not_ast_p_lt = thm "not_ast_p_lt";
+val not_ast_p_ndvd = thm "not_ast_p_ndvd";
+val not_ast_p_dvd = thm "not_ast_p_dvd";
+
+(*Logical construction of the prop*)
+(*Bset*)
+val not_bst_p_conjI = thm "not_bst_p_conjI";
+val not_bst_p_disjI = thm "not_bst_p_disjI";
+val not_bst_p_Q_elim = thm "not_bst_p_Q_elim";
+
+(*Aset*)
+val not_ast_p_conjI = thm "not_ast_p_conjI";
+val not_ast_p_disjI = thm "not_ast_p_disjI";
+val not_ast_p_Q_elim = thm "not_ast_p_Q_elim";
+
+(*Cooper*)
+val cppi_eq = thm "cppi_eq";
+val cpmi_eq = thm "cpmi_eq";
+
+(*Others*)
+val simp_from_to = thm "simp_from_to";
+val P_eqtrue = thm "P_eqtrue";
+val P_eqfalse = thm "P_eqfalse";
+
+(*For Proving NNF*)
+
+val nnf_nn = thm "nnf_nn";
+val nnf_im = thm "nnf_im";
+val nnf_eq = thm "nnf_eq";
+val nnf_sdj = thm "nnf_sdj";
+val nnf_ncj = thm "nnf_ncj";
+val nnf_nim = thm "nnf_nim";
+val nnf_neq = thm "nnf_neq";
+val nnf_ndj = thm "nnf_ndj";
+
+(*For Proving term linearizition*)
+val linearize_dvd = thm "linearize_dvd";
+val lf_lt = thm "lf_lt";
+val lf_eq = thm "lf_eq";
+val lf_dvd = thm "lf_dvd";
+
+
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the cooperprocedure.*)
+(* ------------------------------------------------------------------------- *)
+
+
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the adjustcoeff step.*)
+(* ------------------------------------------------------------------------- *)
+datatype CpLog = No
+                |Simp of term*CpLog
+		|Blast of CpLog*CpLog
+		|Aset of (term*term*(term list)*term)
+		|Bset of (term*term*(term list)*term)
+		|Minusinf of CpLog*CpLog
+		|Cooper of term*CpLog*CpLog*CpLog
+		|Eq_minf of term*term
+		|Modd_minf of term*term
+		|Eq_minf_conjI of CpLog*CpLog
+		|Modd_minf_conjI of CpLog*CpLog	
+		|Modd_minf_disjI of CpLog*CpLog
+		|Eq_minf_disjI of CpLog*CpLog	
+		|Not_bst_p of term*term*term*term*CpLog
+		|Not_bst_p_atomic of term
+		|Not_bst_p_conjI of CpLog*CpLog
+		|Not_bst_p_disjI of CpLog*CpLog
+		|Not_ast_p of term*term*term*term*CpLog
+		|Not_ast_p_atomic of term
+		|Not_ast_p_conjI of CpLog*CpLog
+		|Not_ast_p_disjI of CpLog*CpLog
+		|CpLogError;
+
+
+
+datatype ACLog = ACAt of int*term
+                |ACPI of int*term
+                |ACfm of term
+                |ACNeg of ACLog
+		|ACConst of string*ACLog*ACLog;
+
+
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the CNNF step.*)
+(* ------------------------------------------------------------------------- *)
+
+
+datatype NNFLog = NNFAt of term
+                |NNFSimp of NNFLog
+                |NNFNN of NNFLog
+		|NNFConst of string*NNFLog*NNFLog;
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the linform  step.*)
+(* ------------------------------------------------------------------------- *)
+
+
+datatype LfLog = LfAt of term
+                |LfAtdvd of term
+                |Lffm of term
+                |LfConst of string*LfLog*LfLog
+		|LfNot of LfLog
+		|LfQ of string*string*typ*LfLog;
+
+
+(* ------------------------------------------------------------------------- *)
+(*Datatatype declarations for Proofprotocol for the evaluation- evalc-  step.*)
+(* ------------------------------------------------------------------------- *)
+
+
+datatype EvalLog = EvalAt of term
+                |Evalfm of term
+		|EvalConst of string*EvalLog*EvalLog;
+
+(* ------------------------------------------------------------------------- *)
+(*This function norm_zero_one  replaces the occurences of Numeral1 and Numeral0*)
+(*Respectively by their abstract representation Const("1",..) and COnst("0",..)*)
+(*this is necessary because the theorems use this representation.*)
+(* This function should be elminated in next versions...*)
+(* ------------------------------------------------------------------------- *)
+
+fun norm_zero_one fm = case fm of
+  (Const ("op *",_) $ c $ t) => 
+    if c = one then (norm_zero_one t)
+    else if (dest_numeral c = ~1) 
+         then (Const("uminus",HOLogic.intT --> HOLogic.intT) $ (norm_zero_one t))
+         else (HOLogic.mk_binop "op *" (norm_zero_one c,norm_zero_one t))
+  |(node $ rest) => ((norm_zero_one node)$(norm_zero_one rest))
+  |(Abs(x,T,p)) => (Abs(x,T,(norm_zero_one p)))
+  |_ => fm;
+
+
+(* ------------------------------------------------------------------------- *)
+(* Intended to tell that here we changed the structure of the formula with respect to the posineq theorem : ~(0 < t) = 0 < 1-t*)
+(* ------------------------------------------------------------------------- *)
+fun adjustcoeffeq_wp  x l fm = 
+    case fm of  
+  (Const("Not",_)$(Const("op <",_) $(Const("0",_)) $(rt as (Const ("op +", _)$(Const ("op *",_) $    c $ y ) $z )))) => 
+  if (x = y) 
+  then let  
+       val m = l div (dest_numeral c) 
+       val n = abs (m)
+       val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x)) 
+       val rs = (HOLogic.mk_binrel "op <" (zero,linear_sub [] one (HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
+       in (ACPI(n,fm),rs)
+       end
+  else  let val rs = (HOLogic.mk_binrel "op <" (zero,linear_sub [] one rt )) 
+        in (ACPI(1,fm),rs)
+        end
+
+  |(Const(p,_) $d $( Const ("op +", _)$(Const ("op *",_) $ 
+      c $ y ) $z )) => if (is_arith_rel fm) andalso (x = y) then  
+        let val m = l div (dest_numeral c) 
+           val n = (if p = "op <" then abs(m) else m)  
+           val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x))
+           val rs = (HOLogic.mk_binrel p ((linear_cmul n d),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
+	   in (ACAt(n,fm),rs)
+	   end
+        else (ACfm(fm),fm) 
+  |( Const ("Not", _) $ p) => let val (rsp,rsr) = adjustcoeffeq_wp x l p 
+                              in (ACNeg(rsp),HOLogic.Not $ rsr) 
+                              end
+  |( Const ("op &",_) $ p $ q) =>let val (rspp,rspr) = adjustcoeffeq_wp x l p
+                                     val (rsqp,rsqr) = adjustcoeffeq_wp x l q
+
+                                  in (ACConst ("CJ",rspp,rsqp), HOLogic.mk_conj (rspr,rsqr)) 
+                                  end 
+  |( Const ("op |",_) $ p $ q) =>let val (rspp,rspr) = adjustcoeffeq_wp x l p
+                                     val (rsqp,rsqr) = adjustcoeffeq_wp x l q
+
+                                  in (ACConst ("DJ",rspp,rsqp), HOLogic.mk_disj (rspr,rsqr)) 
+                                  end
+
+  |_ => (ACfm(fm),fm);
+
+
+(*_________________________________________*)
+(*-----------------------------------------*)
+(* Protocol generation for the liform step *)
+(*_________________________________________*)
+(*-----------------------------------------*)
+
+
+fun linform_wp fm = 
+  let fun at_linform_wp at =
+    case at of
+      (Const("op <=",_)$s$t) => LfAt(at)
+      |(Const("op <",_)$s$t) => LfAt(at)
+      |(Const("op =",_)$s$t) => LfAt(at)
+      |(Const("Divides.op dvd",_)$s$t) => LfAtdvd(at)
+  in
+  if is_arith_rel fm 
+  then at_linform_wp fm 
+  else case fm of
+    (Const("Not",_) $ A) => LfNot(linform_wp A)
+   |(Const("op &",_)$ A $ B) => LfConst("CJ",linform_wp A, linform_wp B)
+   |(Const("op |",_)$ A $ B) => LfConst("DJ",linform_wp A, linform_wp B)
+   |(Const("op -->",_)$ A $ B) => LfConst("IM",linform_wp A, linform_wp B)
+   |(Const("op =",Type ("fun",[Type ("bool", []),_]))$ A $ B) => LfConst("EQ",linform_wp A, linform_wp B)
+   |Const("Ex",_)$Abs(x,T,p) => 
+     let val (xn,p1) = variant_abs(x,T,p)
+     in LfQ("Ex",xn,T,linform_wp p1)
+     end 
+   |Const("All",_)$Abs(x,T,p) => 
+     let val (xn,p1) = variant_abs(x,T,p)
+     in LfQ("All",xn,T,linform_wp p1)
+     end 
+end;
+
+
+(* ------------------------------------------------------------------------- *)
+(*For simlified formulas we just notice the original formula, for whitch we habe been
+intendes to make the proof.*)
+(* ------------------------------------------------------------------------- *)
+fun simpl_wp (fm,pr) = let val fm2 = simpl fm
+				in (fm2,Simp(fm,pr))
+				end;
+
+	
+(* ------------------------------------------------------------------------- *)
+(*Help function for the generation of the proof EX.P_{minus \infty} --> EX. P(x) *)
+(* ------------------------------------------------------------------------- *)
+fun minusinf_wph x fm = let fun mk_atomar_minusinf_proof x fm = (Modd_minf(x,fm),Eq_minf(x,fm))
+  
+	      fun combine_minusinf_proofs opr (ppr1,ppr2) (qpr1,qpr2) = case opr of 
+		 "CJ" => (Modd_minf_conjI(ppr1,qpr1),Eq_minf_conjI(ppr2,qpr2))
+		|"DJ" => (Modd_minf_disjI(ppr1,qpr1),Eq_minf_disjI(ppr2,qpr2))
+	in 
+ 
+ case fm of 
+ (Const ("Not", _) $  (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+     if (x=y) andalso (c1= zero) andalso (c2= one) then (HOLogic.true_const ,(mk_atomar_minusinf_proof x fm))
+        else (fm ,(mk_atomar_minusinf_proof x fm))
+ |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	 if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one)
+	 then (HOLogic.false_const ,(mk_atomar_minusinf_proof x fm))
+	 				 else (fm,(mk_atomar_minusinf_proof x fm)) 
+ |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y ) $ z )) =>
+       if (y=x) andalso (c1 = zero) then 
+        if c2 = one then (HOLogic.false_const,(mk_atomar_minusinf_proof x fm)) else
+	(HOLogic.true_const,(mk_atomar_minusinf_proof x fm))
+	else (fm,(mk_atomar_minusinf_proof x fm))
+  
+  |(Const("Not",_)$(Const ("Divides.op dvd",_) $_ )) => (fm,mk_atomar_minusinf_proof x fm)
+  
+  |(Const ("Divides.op dvd",_) $_ ) => (fm,mk_atomar_minusinf_proof x fm)
+  
+  |(Const ("op &",_) $ p $ q) => let val (pfm,ppr) = minusinf_wph x p
+  				    val (qfm,qpr) = minusinf_wph x q
+				    val pr = (combine_minusinf_proofs "CJ" ppr qpr)
+				     in 
+				     (HOLogic.conj $ pfm $qfm , pr)
+				     end 
+  |(Const ("op |",_) $ p $ q) => let val (pfm,ppr) = minusinf_wph x p
+  				     val (qfm,qpr) = minusinf_wph x q
+				     val pr = (combine_minusinf_proofs "DJ" ppr qpr)
+				     in 
+				     (HOLogic.disj $ pfm $qfm , pr)
+				     end 
+
+  |_ => (fm,(mk_atomar_minusinf_proof x fm))
+  
+  end;					 
+(* ------------------------------------------------------------------------- *)	    (* Protokol for the Proof of the property of the minusinfinity formula*)
+(* Just combines the to protokols *)
+(* ------------------------------------------------------------------------- *)
+fun minusinf_wp x fm  = let val (fm2,pr) = (minusinf_wph x fm)
+                       in (fm2,Minusinf(pr))
+                        end;
+
+(* ------------------------------------------------------------------------- *)
+(*Help function for the generation of the proof EX.P_{plus \infty} --> EX. P(x) *)
+(* ------------------------------------------------------------------------- *)
+
+fun plusinf_wph x fm = let fun mk_atomar_plusinf_proof x fm = (Modd_minf(x,fm),Eq_minf(x,fm))
+  
+	      fun combine_plusinf_proofs opr (ppr1,ppr2) (qpr1,qpr2) = case opr of 
+		 "CJ" => (Modd_minf_conjI(ppr1,qpr1),Eq_minf_conjI(ppr2,qpr2))
+		|"DJ" => (Modd_minf_disjI(ppr1,qpr1),Eq_minf_disjI(ppr2,qpr2))
+	in 
+ 
+ case fm of 
+ (Const ("Not", _) $  (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+     if (x=y) andalso (c1= zero) andalso (c2= one) then (HOLogic.true_const ,(mk_atomar_plusinf_proof x fm))
+        else (fm ,(mk_atomar_plusinf_proof x fm))
+ |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	 if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one)
+	 then (HOLogic.false_const ,(mk_atomar_plusinf_proof x fm))
+	 				 else (fm,(mk_atomar_plusinf_proof x fm)) 
+ |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y ) $ z )) =>
+       if (y=x) andalso (c1 = zero) then 
+        if c2 = one then (HOLogic.true_const,(mk_atomar_plusinf_proof x fm)) else
+	(HOLogic.false_const,(mk_atomar_plusinf_proof x fm))
+	else (fm,(mk_atomar_plusinf_proof x fm))
+  
+  |(Const("Not",_)$(Const ("Divides.op dvd",_) $_ )) => (fm,mk_atomar_plusinf_proof x fm)
+  
+  |(Const ("Divides.op dvd",_) $_ ) => (fm,mk_atomar_plusinf_proof x fm)
+  
+  |(Const ("op &",_) $ p $ q) => let val (pfm,ppr) = plusinf_wph x p
+  				    val (qfm,qpr) = plusinf_wph x q
+				    val pr = (combine_plusinf_proofs "CJ" ppr qpr)
+				     in 
+				     (HOLogic.conj $ pfm $qfm , pr)
+				     end 
+  |(Const ("op |",_) $ p $ q) => let val (pfm,ppr) = plusinf_wph x p
+  				     val (qfm,qpr) = plusinf_wph x q
+				     val pr = (combine_plusinf_proofs "DJ" ppr qpr)
+				     in 
+				     (HOLogic.disj $ pfm $qfm , pr)
+				     end 
+
+  |_ => (fm,(mk_atomar_plusinf_proof x fm))
+  
+  end;					 
+(* ------------------------------------------------------------------------- *)	    (* Protokol for the Proof of the property of the minusinfinity formula*)
+(* Just combines the to protokols *)
+(* ------------------------------------------------------------------------- *)
+fun plusinf_wp x fm  = let val (fm2,pr) = (plusinf_wph x fm)
+                       in (fm2,Minusinf(pr))
+                        end;
+
+
+(* ------------------------------------------------------------------------- *)
+(*Protocol that we here uses Bset.*)
+(* ------------------------------------------------------------------------- *)
+fun bset_wp x fm = let val bs = bset x fm in
+				(bs,Bset(x,fm,bs,mk_numeral (divlcm x fm)))
+				end;
+
+(* ------------------------------------------------------------------------- *)
+(*Protocol that we here uses Aset.*)
+(* ------------------------------------------------------------------------- *)
+fun aset_wp x fm = let val ast = aset x fm in
+				(ast,Aset(x,fm,ast,mk_numeral (divlcm x fm)))
+				end;
+ 
+
+
+(* ------------------------------------------------------------------------- *)
+(*function list to Set, constructs a set containing all elements of a given list.*)
+(* ------------------------------------------------------------------------- *)
+fun list_to_set T1 l = let val T = (HOLogic.mk_setT T1) in 
+	case l of 
+		[] => Const ("{}",T)
+		|(h::t) => Const("insert", T1 --> (T --> T)) $ h $(list_to_set T1 t)
+		end;
+		
+
+(*====================================================================*)
+(* ------------------------------------------------------------------------- *)
+(* ------------------------------------------------------------------------- *)
+(*Protocol for the proof of the backward direction of the cooper theorem.*)
+(* Helpfunction - Protokols evereything about the proof reconstruction*)
+(* ------------------------------------------------------------------------- *)
+fun not_bst_p_wph fm = case fm of
+	Const("Not",_) $ R => if (is_arith_rel R) then (Not_bst_p_atomic (fm)) else CpLogError
+	|Const("op &",_) $ ls $ rs => Not_bst_p_conjI((not_bst_p_wph ls),(not_bst_p_wph rs))
+	|Const("op |",_) $ ls $ rs => Not_bst_p_disjI((not_bst_p_wph ls),(not_bst_p_wph rs))
+	|_ => Not_bst_p_atomic (fm);
+(* ------------------------------------------------------------------------- *)	
+(* Main protocoling function for the backward direction gives the Bset and the divlcm and the Formula herself. Needed as inherited attributes for the proof reconstruction*)
+(* ------------------------------------------------------------------------- *)
+fun not_bst_p_wp x fm = let val prt = not_bst_p_wph fm
+			    val D = mk_numeral (divlcm x fm)
+			    val B = map norm_zero_one (bset x fm)
+			in (Not_bst_p (x,fm,D,(list_to_set HOLogic.intT B) , prt))
+			end;
+(*====================================================================*)
+(* ------------------------------------------------------------------------- *)
+(* ------------------------------------------------------------------------- *)
+(*Protocol for the proof of the backward direction of the cooper theorem.*)
+(* Helpfunction - Protokols evereything about the proof reconstruction*)
+(* ------------------------------------------------------------------------- *)
+fun not_ast_p_wph fm = case fm of
+	Const("Not",_) $ R => if (is_arith_rel R) then (Not_ast_p_atomic (fm)) else CpLogError
+	|Const("op &",_) $ ls $ rs => Not_ast_p_conjI((not_ast_p_wph ls),(not_ast_p_wph rs))
+	|Const("op |",_) $ ls $ rs => Not_ast_p_disjI((not_ast_p_wph ls),(not_ast_p_wph rs))
+	|_ => Not_ast_p_atomic (fm);
+(* ------------------------------------------------------------------------- *)	
+(* Main protocoling function for the backward direction gives the Bset and the divlcm and the Formula herself. Needed as inherited attributes for the proof reconstruction*)
+(* ------------------------------------------------------------------------- *)
+fun not_ast_p_wp x fm = let val prt = not_ast_p_wph fm
+			    val D = mk_numeral (divlcm x fm)
+			    val B = map norm_zero_one (aset x fm)
+			in (Not_ast_p (x,fm,D,(list_to_set HOLogic.intT B) , prt))
+			end;
+
+(*======================================================*)
+(* Protokolgeneration for the formula evaluation process*)
+(*======================================================*)
+
+fun evalc_wp fm = 
+  let fun evalc_atom_wp at =case at of  
+    (Const (p,_) $ s $ t) =>(  
+    case assoc (operations,p) of 
+        Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then EvalAt(HOLogic.mk_eq(at,HOLogic.true_const)) else EvalAt(HOLogic.mk_eq(at, HOLogic.false_const)))  
+		   handle _ => Evalfm(at)) 
+        | _ =>  Evalfm(at)) 
+     |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
+       case assoc (operations,p) of 
+         Some f => ((if (f ((dest_numeral s),(dest_numeral t))) then 
+	  EvalAt(HOLogic.mk_eq(at, HOLogic.false_const))  else EvalAt(HOLogic.mk_eq(at,HOLogic.true_const)))  
+		      handle _ => Evalfm(at)) 
+         | _ => Evalfm(at)) 
+     | _ => Evalfm(at)  
+ 
+  in
+   case fm of
+    (Const("op &",_)$A$B) => EvalConst("CJ",evalc_wp A,evalc_wp B)
+   |(Const("op |",_)$A$B) => EvalConst("DJ",evalc_wp A,evalc_wp B) 
+   |(Const("op -->",_)$A$B) => EvalConst("IM",evalc_wp A,evalc_wp B) 
+   |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => EvalConst("EQ",evalc_wp A,evalc_wp B) 
+   |_ => evalc_atom_wp fm
+  end;
+
+
+
+(*======================================================*)
+(* Protokolgeneration for the NNF Transformation        *)
+(*======================================================*)
+
+fun cnnf_wp f = 
+  let fun hcnnf_wp fm =
+    case fm of
+    (Const ("op &",_) $ p $ q) => NNFConst("CJ",hcnnf_wp p,hcnnf_wp q) 
+    | (Const ("op |",_) $ p $ q) =>  NNFConst("DJ",hcnnf_wp p,hcnnf_wp q)
+    | (Const ("op -->",_) $ p $q) => NNFConst("IM",hcnnf_wp (HOLogic.Not $ p),hcnnf_wp q)
+    | (Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $ q) => NNFConst("EQ",hcnnf_wp (HOLogic.mk_conj(p,q)),hcnnf_wp (HOLogic.mk_conj((HOLogic.Not $ p), (HOLogic.Not $ q)))) 
+
+    | (Const ("Not",_) $ (Const("Not",_) $ p)) => NNFNN(hcnnf_wp p) 
+    | (Const ("Not",_) $ (Const ("op &",_) $ p $ q)) => NNFConst ("NCJ",(hcnnf_wp(HOLogic.Not $ p)),(hcnnf_wp(HOLogic.Not $ q))) 
+    | (Const ("Not",_) $(Const ("op |",_) $ (A as (Const ("op &",_) $ p $ q)) $  
+    			(B as (Const ("op &",_) $ p1 $ r)))) => if p1 = negate p then 
+		         NNFConst("SDJ",  
+			   NNFConst("CJ",hcnnf_wp p,hcnnf_wp(HOLogic.Not $ q)),
+			   NNFConst("CJ",hcnnf_wp p1,hcnnf_wp(HOLogic.Not $ r)))
+			 else  NNFConst ("NDJ",(hcnnf_wp(HOLogic.Not $ A)),(hcnnf_wp(HOLogic.Not $ B))) 
+
+    | (Const ("Not",_) $ (Const ("op |",_) $ p $ q)) => NNFConst ("NDJ",(hcnnf_wp(HOLogic.Not $ p)),(hcnnf_wp(HOLogic.Not $ q))) 
+    | (Const ("Not",_) $ (Const ("op -->",_) $ p $q)) =>  NNFConst ("NIM",(hcnnf_wp(p)),(hcnnf_wp(HOLogic.Not $ q))) 
+    | (Const ("Not",_) $ (Const ("op =",Type ("fun",[Type ("bool", []),_]))  $ p $ q)) =>NNFConst ("NEQ",(NNFConst("CJ",hcnnf_wp p,hcnnf_wp(HOLogic.Not $ q))),(NNFConst("CJ",hcnnf_wp(HOLogic.Not $ p),hcnnf_wp q))) 
+    | _ => NNFAt(fm)  
+  in NNFSimp(hcnnf_wp f)
+end; 
+   
+
+
+
+
+
+(* ------------------------------------------------------------------------- *)
+(*Cooper decision Procedure with proof protocoling*)
+(* ------------------------------------------------------------------------- *)
+
+fun coopermi_wp vars fm =
+  case fm of
+   Const ("Ex",_) $ Abs(xo,T,po) => let 
+    val (xn,np) = variant_abs(xo,T,po) 
+    val x = (Free(xn , T))
+    val p = np     (* Is this a legal proof for the P=NP Problem??*)
+    val (p_inf,miprt) = simpl_wp (minusinf_wp x p)
+    val (bset,bsprt) = bset_wp x p
+    val nbst_p_prt = not_bst_p_wp x p
+    val dlcm = divlcm x p 
+    val js = 1 upto dlcm 
+    fun p_element j b = linrep vars x (linear_add vars b (mk_numeral j)) p 
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) bset) 
+   in (list_disj (map stage js),Cooper(mk_numeral dlcm,miprt,bsprt,nbst_p_prt))
+   end
+   
+  | _ => (error "cooper: not an existential formula",No);
+				
+fun cooperpi_wp vars fm =
+  case fm of
+   Const ("Ex",_) $ Abs(xo,T,po) => let 
+    val (xn,np) = variant_abs(xo,T,po) 
+    val x = (Free(xn , T))
+    val p = np     (* Is this a legal proof for the P=NP Problem??*)
+    val (p_inf,piprt) = simpl_wp (plusinf_wp x p)
+    val (aset,asprt) = aset_wp x p
+    val nast_p_prt = not_ast_p_wp x p
+    val dlcm = divlcm x p 
+    val js = 1 upto dlcm 
+    fun p_element j a = linrep vars x (linear_sub vars a (mk_numeral j)) p 
+    fun stage j = list_disj (linrep vars x (mk_numeral j) p_inf :: map (p_element j) aset) 
+   in (list_disj (map stage js),Cooper(mk_numeral dlcm,piprt,asprt,nast_p_prt))
+   end
+  | _ => (error "cooper: not an existential formula",No);
+				
+
+
+
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---      Interpretation and Proofgeneration Part              ---*)
+(*---                                                           ---*)
+(*---      Protocole interpretation functions                   ---*)
+(*---                                                           ---*)
+(*---      and proofgeneration functions                        ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+(* ------------------------------------------------------------------------- *)
+(* Returns both sides of an equvalence in the theorem*)
+(* ------------------------------------------------------------------------- *)
+fun qe_get_terms th = let val (_$(Const("op =",Type ("fun",[Type ("bool", []),_])) $ A $ B )) = prop_of th in (A,B) end;
+
+
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+
+(* ------------------------------------------------------------------------- *)
+(* Modified version of the simple version with minimal amount of checking and postprocessing*)
+(* ------------------------------------------------------------------------- *)
+
+fun simple_prove_goal_cterm2 G tacs =
+  let
+    fun check None = error "prove_goal: tactic failed"
+      | check (Some (thm, _)) = (case nprems_of thm of
+            0 => thm
+          | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
+  in check (Seq.pull (EVERY tacs (trivial G))) end;
+
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+(*-------------------------------------------------------------*)
+
+fun cert_Trueprop sg t = cterm_of sg (HOLogic.mk_Trueprop t);
+
+(* ------------------------------------------------------------------------- *)
+(*This function proove elementar will be used to generate proofs at runtime*)
+(*It is is based on the isabelle function proove_goalw_cterm and is thought to *)
+(*prove properties such as a dvd b (essentially) that are only to make at
+runtime.*)
+(* ------------------------------------------------------------------------- *)
+fun prove_elementar sg s fm2 = case s of 
+  (*"ss" like simplification with simpset*)
+  "ss" =>
+    let
+      val ss = presburger_ss addsimps
+        [zdvd_iff_zmod_eq_0,unity_coeff_ex]
+      val ct =  cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]
+    end
+
+  (*"bl" like blast tactic*)
+  (* Is only used in the harrisons like proof procedure *)
+  | "bl" =>
+     let val ct = cert_Trueprop sg fm2
+     in
+       simple_prove_goal_cterm2 ct [blast_tac HOL_cs 1]
+     end
+
+  (*"ed" like Existence disjunctions ...*)
+  (* Is only used in the harrisons like proof procedure *)
+  | "ed" =>
+    let
+      val ex_disj_tacs =
+        let
+          val tac1 = EVERY[REPEAT(resolve_tac [disjI1,disjI2] 1), etac exI 1]
+          val tac2 = EVERY[etac exE 1, rtac exI 1,
+            REPEAT(resolve_tac [disjI1,disjI2] 1), assumption 1]
+	in [rtac iffI 1,
+          etac exE 1, REPEAT(EVERY[etac disjE 1, tac1]), tac1,
+          REPEAT(EVERY[etac disjE 1, tac2]), tac2]
+        end
+
+      val ct = cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct ex_disj_tacs
+    end
+
+  | "fa" =>
+    let val ct = cert_Trueprop sg fm2
+    in simple_prove_goal_cterm2 ct [simple_arith_tac 1]
+    end
+
+  | "sa" =>
+    let
+      val ss = presburger_ss addsimps zadd_ac
+      val ct = cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]
+    end
+
+  | "ac" =>
+    let
+      val ss = HOL_basic_ss addsimps zadd_ac
+      val ct = cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct [simp_tac ss 1]
+    end
+
+  | "lf" =>
+    let
+      val ss = presburger_ss addsimps zadd_ac
+      val ct = cert_Trueprop sg fm2
+    in 
+      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]
+    end;
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function return an Isabelle proof, of the adjustcoffeq result.*)
+(* The proofs are in Presburger.thy and are generally based on the arithmetic *)
+(* ------------------------------------------------------------------------- *)
+fun proof_of_adjustcoeffeq sg (prt,rs) = case prt of
+   ACfm fm => instantiate' [Some cboolT]
+    [Some (cterm_of sg fm)] refl
+ | ACAt (k,at as (Const(p,_) $a $( Const ("op +", _)$(Const ("op *",_) $ 
+      c $ x ) $t ))) => 
+   let
+     val ck = cterm_of sg (mk_numeral k)
+     val cc = cterm_of sg c
+     val ct = cterm_of sg t
+     val cx = cterm_of sg x
+     val ca = cterm_of sg a
+   in case p of
+     "op <" => let val pre = prove_elementar sg "ss" 
+	                  (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral k)))
+	           val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_lt_eq)))
+		      in [th1,(prove_elementar sg "ss" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans
+                   end
+    |"op =" =>let val pre = prove_elementar sg "ss" 
+	    (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
+	          in let val th1 = (pre RS(instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_eq_eq)))
+	             in [th1,(prove_elementar sg "ss" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans
+                      end
+                  end
+    |"Divides.op dvd" =>let val pre = prove_elementar sg "ss" 
+	   (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
+	                 val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct]) (ac_dvd_eq))
+                         in [th1,(prove_elementar sg "ss" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans
+                        
+                          end
+  end
+ |ACPI(k,at as (Const("Not",_)$(Const("op <",_) $a $( Const ("op +", _)$(Const ("op *",_) $ c $ x ) $t )))) => 
+   let
+     val ck = cterm_of sg (mk_numeral k)
+     val cc = cterm_of sg c
+     val ct = cterm_of sg t
+     val cx = cterm_of sg x
+     val pre = prove_elementar sg "ss" 
+       (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral k)))
+       val th1 = (pre RS (instantiate' [] [Some ck,Some cc, Some cx, Some ct] (ac_pi_eq)))
+
+         in [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans
+   end
+ |ACNeg(pr) => let val (Const("Not",_)$nrs) = rs
+               in (proof_of_adjustcoeffeq sg (pr,nrs)) RS (qe_Not) 
+               end
+ |ACConst(s,pr1,pr2) =>
+   let val (Const(_,_)$rs1$rs2) = rs
+       val th1 = proof_of_adjustcoeffeq sg (pr1,rs1)
+       val th2 = proof_of_adjustcoeffeq sg (pr2,rs2)
+       in case s of 
+	 "CJ" => [th1,th2] MRS (qe_conjI)
+         |"DJ" => [th1,th2] MRS (qe_disjI)
+         |"IM" => [th1,th2] MRS (qe_impI)
+         |"EQ" => [th1,th2] MRS (qe_eqI)
+   end;
+
+
+
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function return an Isabelle proof, of some properties on the atoms*)
+(* The proofs are in Presburger.thy and are generally based on the arithmetic *)
+(* This function doese only instantiate the the theorems in the theory *)
+(* ------------------------------------------------------------------------- *)
+fun atomar_minf_proof_of sg dlcm (Modd_minf (x,fm1)) =
+  let
+    (*Some certified Terms*)
+    
+   val ctrue = cterm_of sg HOLogic.true_const
+   val cfalse = cterm_of sg HOLogic.false_const
+   val fm = norm_zero_one fm1
+  in  case fm1 of 
+      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+         if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
+           else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+
+      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	   if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one) 
+	   then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf))
+	 	 else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf)) 
+
+      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+           if (y=x) andalso (c1 = zero) then 
+            if (pm1 = one) then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf)) else
+	     (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
+	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+  
+      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
+			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
+	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf)))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+      |(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
+      c $ y ) $ z))) => 
+         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
+			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
+	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf)))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
+		
+    
+   |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf)
+   end	
+
+ |atomar_minf_proof_of sg dlcm (Eq_minf (x,fm1)) =  let
+       (*Some certified types*)
+   val fm = norm_zero_one fm1
+    in  case fm1 of 
+      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+         if  (x=y) andalso (c1=zero) andalso (c2=one) 
+	   then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_minf))
+           else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+
+      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	   if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
+	     then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_minf))
+	     else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf)) 
+
+      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+           if (y=x) andalso (c1 =zero) then 
+            if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else
+	     (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_minf))
+	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
+	 		  val cz = cterm_of sg (norm_zero_one z)
+	 	      in(instantiate' [] [Some cd,  Some cz] (not_dvd_eq_minf)) 
+		      end
+
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+		
+      |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
+	 		  val cz = cterm_of sg (norm_zero_one z)
+	 	      in(instantiate' [] [Some cd, Some cz ] (dvd_eq_minf))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+
+      		
+    |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
+ end;
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function combines proofs of some special form already synthetised from the subtrees to make*)
+(* a new proof of the same form. The combination occures whith isabelle theorems which have been already prooved *)
+(*these Theorems are in Presburger.thy and mostly do not relay on the arithmetic.*)
+(* These are Theorems for the Property of P_{-infty}*)
+(* ------------------------------------------------------------------------- *)
+fun combine_minf_proof s pr1 pr2 = case s of
+    "ECJ" => [pr1 , pr2] MRS (eq_minf_conjI)
+
+   |"EDJ" => [pr1 , pr2] MRS (eq_minf_disjI)
+   
+   |"MCJ" => [pr1 , pr2] MRS (modd_minf_conjI)
+
+   |"MDJ" => [pr1 , pr2] MRS (modd_minf_disjI);
+
+(* ------------------------------------------------------------------------- *)
+(*This function return an isabelle Proof for the minusinfinity theorem*)
+(* It interpretates the protool and gives the protokoles property of P_{...} as a theorem*)
+(* ------------------------------------------------------------------------- *)
+fun minf_proof_ofh sg dlcm prl = case prl of 
+
+    Eq_minf (_) => atomar_minf_proof_of sg dlcm prl
+    
+   |Modd_minf (_) => atomar_minf_proof_of sg dlcm prl
+   
+   |Eq_minf_conjI (prl1,prl2) => let val pr1 = minf_proof_ofh sg dlcm prl1
+   				    val pr2 = minf_proof_ofh sg dlcm prl2
+				 in (combine_minf_proof "ECJ" pr1 pr2)
+				 end
+				 
+   |Eq_minf_disjI (prl1,prl2) => let val pr1 = minf_proof_ofh sg dlcm prl1
+   				    val pr2 = minf_proof_ofh sg dlcm prl2
+				 in (combine_minf_proof "EDJ" pr1 pr2)
+				 end
+				 
+   |Modd_minf_conjI (prl1,prl2) => let val pr1 = minf_proof_ofh sg dlcm prl1
+   				    val pr2 = minf_proof_ofh sg dlcm prl2
+				 in (combine_minf_proof "MCJ" pr1 pr2)
+				 end
+				 
+   |Modd_minf_disjI (prl1,prl2) => let val pr1 = minf_proof_ofh sg dlcm prl1
+   				    val pr2 = minf_proof_ofh sg dlcm prl2
+				 in (combine_minf_proof "MDJ" pr1 pr2)
+				 end;
+(* ------------------------------------------------------------------------- *)
+(* Main function For the rest both properies of P_{..} are needed and here both theorems are returned.*)				 
+(* ------------------------------------------------------------------------- *)
+fun  minf_proof_of sg dlcm (Minusinf (prl1,prl2))  = 
+  let val pr1 = minf_proof_ofh sg dlcm prl1
+      val pr2 = minf_proof_ofh sg dlcm prl2
+  in (pr1, pr2)
+end;
+				 
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function return an Isabelle proof, of some properties on the atoms*)
+(* The proofs are in Presburger.thy and are generally based on the arithmetic *)
+(* This function doese only instantiate the the theorems in the theory *)
+(* ------------------------------------------------------------------------- *)
+fun atomar_pinf_proof_of sg dlcm (Modd_minf (x,fm1)) =
+ let
+    (*Some certified Terms*)
+    
+  val ctrue = cterm_of sg HOLogic.true_const
+  val cfalse = cterm_of sg HOLogic.false_const
+  val fm = norm_zero_one fm1
+ in  case fm1 of 
+      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+         if ((x=y) andalso (c1= zero) andalso (c2= one))
+	 then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf))
+         else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+
+      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	if ((is_arith_rel fm) andalso (x = y) andalso (c1 = zero)  andalso (c2 = one)) 
+	then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
+	else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+
+      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+        if ((y=x) andalso (c1 = zero)) then 
+          if (pm1 = one) 
+	  then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf)) 
+	  else (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
+	else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+  
+      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
+			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
+	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf)))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+      |(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
+      c $ y ) $ z))) => 
+         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
+			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
+	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf)))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
+		
+    
+   |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf)
+   end	
+
+ |atomar_pinf_proof_of sg dlcm (Eq_minf (x,fm1)) =  let
+					val fm = norm_zero_one fm1
+    in  case fm1 of 
+      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+         if  (x=y) andalso (c1=zero) andalso (c2=one) 
+	   then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_pinf))
+           else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+
+      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
+  	   if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
+	     then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_pinf))
+	     else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf)) 
+
+      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+           if (y=x) andalso (c1 =zero) then 
+            if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else
+	     (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_pinf))
+	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
+	 		  val cz = cterm_of sg (norm_zero_one z)
+	 	      in(instantiate' [] [Some cd,  Some cz] (not_dvd_eq_pinf)) 
+		      end
+
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+		
+      |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
+	 		  val cz = cterm_of sg (norm_zero_one z)
+	 	      in(instantiate' [] [Some cd, Some cz ] (dvd_eq_pinf))
+		      end
+		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+
+      		
+    |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
+ end;
+
+
+(* ------------------------------------------------------------------------- *)
+(* This function combines proofs of some special form already synthetised from the subtrees to make*)
+(* a new proof of the same form. The combination occures whith isabelle theorems which have been already prooved *)
+(*these Theorems are in Presburger.thy and mostly do not relay on the arithmetic.*)
+(* These are Theorems for the Property of P_{+infty}*)
+(* ------------------------------------------------------------------------- *)
+fun combine_pinf_proof s pr1 pr2 = case s of
+    "ECJ" => [pr1 , pr2] MRS (eq_pinf_conjI)
+
+   |"EDJ" => [pr1 , pr2] MRS (eq_pinf_disjI)
+   
+   |"MCJ" => [pr1 , pr2] MRS (modd_pinf_conjI)
+
+   |"MDJ" => [pr1 , pr2] MRS (modd_pinf_disjI);
+
+(* ------------------------------------------------------------------------- *)
+(*This function return an isabelle Proof for the minusinfinity theorem*)
+(* It interpretates the protool and gives the protokoles property of P_{...} as a theorem*)
+(* ------------------------------------------------------------------------- *)
+fun pinf_proof_ofh sg dlcm prl = case prl of 
+
+    Eq_minf (_) => atomar_pinf_proof_of sg dlcm prl
+    
+   |Modd_minf (_) => atomar_pinf_proof_of sg dlcm prl
+   
+   |Eq_minf_conjI (prl1,prl2) => let val pr1 = pinf_proof_ofh sg dlcm prl1
+   				    val pr2 = pinf_proof_ofh sg dlcm prl2
+				 in (combine_pinf_proof "ECJ" pr1 pr2)
+				 end
+				 
+   |Eq_minf_disjI (prl1,prl2) => let val pr1 = pinf_proof_ofh sg dlcm prl1
+   				    val pr2 = pinf_proof_ofh sg dlcm prl2
+				 in (combine_pinf_proof "EDJ" pr1 pr2)
+				 end
+				 
+   |Modd_minf_conjI (prl1,prl2) => let val pr1 = pinf_proof_ofh sg dlcm prl1
+   				    val pr2 = pinf_proof_ofh sg dlcm prl2
+				 in (combine_pinf_proof "MCJ" pr1 pr2)
+				 end
+				 
+   |Modd_minf_disjI (prl1,prl2) => let val pr1 = pinf_proof_ofh sg dlcm prl1
+   				    val pr2 = pinf_proof_ofh sg dlcm prl2
+				 in (combine_pinf_proof "MDJ" pr1 pr2)
+				 end;
+(* ------------------------------------------------------------------------- *)
+(* Main function For the rest both properies of P_{..} are needed and here both theorems are returned.*)				 
+(* ------------------------------------------------------------------------- *)
+fun pinf_proof_of sg dlcm (Minusinf (prl1,prl2))  = 
+  let val pr1 = pinf_proof_ofh sg dlcm prl1
+      val pr2 = pinf_proof_ofh sg dlcm prl2
+  in (pr1, pr2)
+end;
+				 
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* Here we generate the theorem for the Bset Property in the simple direction*)
+(* It is just an instantiation*)
+(* ------------------------------------------------------------------------- *)
+fun bsetproof_of sg (Bset(x as Free(xn,xT),fm,bs,dlcm))   = 
+  let
+    val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
+    val cdlcm = cterm_of sg dlcm
+    val cB = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one bs))
+  in instantiate' [] [Some cdlcm,Some cB, Some cp] (bst_thm)
+    end;
+
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* Here we generate the theorem for the Bset Property in the simple direction*)
+(* It is just an instantiation*)
+(* ------------------------------------------------------------------------- *)
+fun asetproof_of sg (Aset(x as Free(xn,xT),fm,ast,dlcm))   = 
+  let
+    val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
+    val cdlcm = cterm_of sg dlcm
+    val cA = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one ast))
+  in instantiate' [] [Some cdlcm,Some cA, Some cp] (ast_thm)
+end;
+
+
+
+
+(* ------------------------------------------------------------------------- *)    
+(* Protokol interpretation function for the backwards direction for cooper's Theorem*)
+
+(* For the generation of atomic Theorems*)
+(* Prove the premisses on runtime and then make RS*)
+(* ------------------------------------------------------------------------- *)
+fun generate_atomic_not_bst_p sg (x as Free(xn,xT)) fm dlcm B at = 
+  let
+    val cdlcm = cterm_of sg dlcm
+    val cB = cterm_of sg B
+    val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
+    val cat = cterm_of sg (norm_zero_one at)
+  in
+  case at of 
+   (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+      if  (x=y) andalso (c1=zero) andalso (c2=one) 
+	 then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
+	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
+		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	 in  (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_bst_p_ne)))
+	 end
+         else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+
+   |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
+     if (is_arith_rel at) andalso (x=y)
+	then let val bst_z = norm_zero_one (linear_neg (linear_add [] z (mk_numeral 1)))
+	         in let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ bst_z $ B)
+	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq (bst_z,Const("op -",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
+		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	 in  (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_bst_p_eq)))
+	 end
+       end
+         else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+
+   |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+        if (y=x) andalso (c1 =zero) then 
+        if pm1 = one then 
+	  let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
+              val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
+	  in  (instantiate' [] [Some cfma,  Some cdlcm]([th1,th2] MRS (not_bst_p_gt)))
+	    end
+	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	      in (instantiate' [] [Some cfma, Some cB,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt)))
+	      end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+
+   |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+      if y=x then  
+           let val cz = cterm_of sg (norm_zero_one z)
+	       val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
+ 	     in (instantiate' []  [Some cfma, Some cB,Some cz] (th1 RS (not_bst_p_ndvd)))
+	     end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+
+   |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+       if y=x then  
+	 let val cz = cterm_of sg (norm_zero_one z)
+	     val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
+ 	    in (instantiate' []  [Some cfma,Some cB,Some cz] (th1 RS (not_bst_p_dvd)))
+	  end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+      		
+   |_ => (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
+      		
+    end;
+    
+(* ------------------------------------------------------------------------- *)    
+(* Main interpretation function for this backwards dirction*)
+(* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*)
+(*Help Function*)
+(* ------------------------------------------------------------------------- *)
+fun not_bst_p_proof_of_h sg x fm dlcm B prt = case prt of 
+	(Not_bst_p_atomic(fm2)) => (generate_atomic_not_bst_p sg x fm dlcm B fm2)
+	
+	|(Not_bst_p_conjI(pr1,pr2)) => 
+			let val th1 = (not_bst_p_proof_of_h sg x fm dlcm B pr1)
+			    val th2 = (not_bst_p_proof_of_h sg x fm dlcm B pr2)
+			    in ([th1,th2] MRS (not_bst_p_conjI))
+			    end
+
+	|(Not_bst_p_disjI(pr1,pr2)) => 
+			let val th1 = (not_bst_p_proof_of_h sg x fm dlcm B pr1)
+			    val th2 = (not_bst_p_proof_of_h sg x fm dlcm B pr2)
+			    in ([th1,th2] MRS not_bst_p_disjI)
+			    end;
+(* Main function*)
+fun not_bst_p_proof_of sg (Not_bst_p(x as Free(xn,xT),fm,dlcm,B,prl)) =
+  let val th =  not_bst_p_proof_of_h sg x fm dlcm B prl
+      val fma = absfree (xn,xT, norm_zero_one fm)
+  in let val th1 =  prove_elementar sg "ss"  (HOLogic.mk_eq (fma,fma))
+     in [th,th1] MRS (not_bst_p_Q_elim)
+     end
+  end;
+
+
+(* ------------------------------------------------------------------------- *)    
+(* Protokol interpretation function for the backwards direction for cooper's Theorem*)
+
+(* For the generation of atomic Theorems*)
+(* Prove the premisses on runtime and then make RS*)
+(* ------------------------------------------------------------------------- *)
+fun generate_atomic_not_ast_p sg (x as Free(xn,xT)) fm dlcm A at = 
+  let
+    val cdlcm = cterm_of sg dlcm
+    val cA = cterm_of sg A
+    val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
+    val cat = cterm_of sg (norm_zero_one at)
+  in
+  case at of 
+   (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
+      if  (x=y) andalso (c1=zero) andalso (c2=one) 
+	 then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ A)
+	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
+		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	 in  (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_ast_p_ne)))
+	 end
+         else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+
+   |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
+     if (is_arith_rel at) andalso (x=y)
+	then let val ast_z = norm_zero_one (linear_sub [] one z )
+	         val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ ast_z $ A)
+	         val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq (ast_z,Const("op +",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
+		 val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	 in  (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_ast_p_eq)))
+       end
+         else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+
+   |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
+        if (y=x) andalso (c1 =zero) then 
+        if pm1 = (mk_numeral ~1) then 
+	  let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one z) $ A)
+              val th2 =  prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm))
+	  in  (instantiate' [] [Some cfma]([th2,th1] MRS (not_ast_p_lt)))
+	    end
+	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
+	      in (instantiate' [] [Some cfma, Some cA,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt)))
+	      end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+
+   |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+      if y=x then  
+           let val cz = cterm_of sg (norm_zero_one z)
+	       val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
+ 	     in (instantiate' []  [Some cfma, Some cA,Some cz] (th1 RS (not_ast_p_ndvd)))
+	     end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+
+   |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
+       if y=x then  
+	 let val cz = cterm_of sg (norm_zero_one z)
+	     val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
+ 	    in (instantiate' []  [Some cfma,Some cA,Some cz] (th1 RS (not_ast_p_dvd)))
+	  end
+      else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+      		
+   |_ => (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
+      		
+    end;
+    
+(* ------------------------------------------------------------------------- *)    
+(* Main interpretation function for this backwards dirction*)
+(* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*)
+(*Help Function*)
+(* ------------------------------------------------------------------------- *)
+fun not_ast_p_proof_of_h sg x fm dlcm A prt = case prt of 
+	(Not_ast_p_atomic(fm2)) => (generate_atomic_not_ast_p sg x fm dlcm A fm2)
+	
+	|(Not_ast_p_conjI(pr1,pr2)) => 
+			let val th1 = (not_ast_p_proof_of_h sg x fm dlcm A pr1)
+			    val th2 = (not_ast_p_proof_of_h sg x fm dlcm A pr2)
+			    in ([th1,th2] MRS (not_ast_p_conjI))
+			    end
+
+	|(Not_ast_p_disjI(pr1,pr2)) => 
+			let val th1 = (not_ast_p_proof_of_h sg x fm dlcm A pr1)
+			    val th2 = (not_ast_p_proof_of_h sg x fm dlcm A pr2)
+			    in ([th1,th2] MRS (not_ast_p_disjI))
+			    end;
+(* Main function*)
+fun not_ast_p_proof_of sg (Not_ast_p(x as Free(xn,xT),fm,dlcm,A,prl)) =
+  let val th =  not_ast_p_proof_of_h sg x fm dlcm A prl
+      val fma = absfree (xn,xT, norm_zero_one fm)
+      val th1 =  prove_elementar sg "ss"  (HOLogic.mk_eq (fma,fma))
+  in [th,th1] MRS (not_ast_p_Q_elim)
+end;
+
+
+
+
+(* ------------------------------------------------------------------------- *)
+(* Interpretaion of Protocols of the cooper procedure : minusinfinity version*)
+(* ------------------------------------------------------------------------- *)
+
+
+fun coopermi_proof_of sg x (Cooper (dlcm,Simp(fm,miprt),bsprt,nbst_p_prt)) =
+  (* Get the Bset thm*)
+  let val bst = bsetproof_of sg bsprt
+      val (mit1,mit2) = minf_proof_of sg dlcm miprt
+      val fm1 = norm_zero_one (simpl fm) 
+      val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm));
+      val nbstpthm = not_bst_p_proof_of sg nbst_p_prt
+    (* Return the four theorems needed to proove the whole Cooper Theorem*)
+  in (dpos,mit2,bst,nbstpthm,mit1)
+end;
+
+
+(* ------------------------------------------------------------------------- *)
+(* Interpretaion of Protocols of the cooper procedure : plusinfinity version *)
+(* ------------------------------------------------------------------------- *)
+
+
+fun cooperpi_proof_of sg x (Cooper (dlcm,Simp(fm,miprt),bsprt,nast_p_prt)) =
+  let val ast = asetproof_of sg bsprt
+      val (mit1,mit2) = pinf_proof_of sg dlcm miprt
+      val fm1 = norm_zero_one (simpl fm) 
+      val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm));
+      val nastpthm = not_ast_p_proof_of sg nast_p_prt
+  in (dpos,mit2,ast,nastpthm,mit1)
+end;
+
+
+(* ------------------------------------------------------------------------- *)
+(* Interpretaion of Protocols of the cooper procedure : full version*)
+(* ------------------------------------------------------------------------- *)
+
+
+
+fun cooper_thm sg s (x as Free(xn,xT)) vars cfm = case s of
+  "pi" => let val (rs,prt) = cooperpi_wp (xn::vars) (HOLogic.mk_exists(xn,xT,cfm))
+	      val (dpsthm,th1,th2,nbpth,th3) = cooperpi_proof_of sg x prt
+		   in [dpsthm,th1,th2,nbpth,th3] MRS (cppi_eq)
+           end
+  |"mi" => let val (rs,prt) = coopermi_wp (xn::vars) (HOLogic.mk_exists(xn,xT,cfm))
+	       val (dpsthm,th1,th2,nbpth,th3) = coopermi_proof_of sg x prt
+		   in [dpsthm,th1,th2,nbpth,th3] MRS (cpmi_eq)
+                end
+ |_ => error "parameter error";
+
+(* ------------------------------------------------------------------------- *)
+(* This function should evoluate to the end prove Procedure for one quantifier elimination for Presburger arithmetic*)
+(* It shoud be plugged in the qfnp argument of the quantifier elimination proof function*)
+(* ------------------------------------------------------------------------- *)
+
+fun cooper_prv sg (x as Free(xn,xT)) efm vars = let 
+   val l = formlcm x efm
+   val ac_thm = proof_of_adjustcoeffeq sg (adjustcoeffeq_wp  x l efm)
+   val fm = snd (qe_get_terms ac_thm)
+   val  cfm = unitycoeff x fm
+   val afm = adjustcoeff x l fm
+   val P = absfree(xn,xT,afm)
+   val ss = presburger_ss addsimps
+     [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
+   val uth = instantiate' [] [Some (cterm_of sg P) , Some (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
+   val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
+   val cms = if ((length (aset x cfm)) < (length (bset x cfm))) then "pi" else "mi"
+   val cp_thm = cooper_thm sg cms x vars cfm
+   val exp_cp_thm = refl RS (simplify ss (cp_thm RSN (2,trans)))
+   val (lsuth,rsuth) = qe_get_terms (uth)
+   val (lseacth,rseacth) = qe_get_terms(e_ac_thm)
+   val (lscth,rscth) = qe_get_terms (exp_cp_thm)
+   val  u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans
+ in  ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans)
+   end
+|cooper_prv _ _ _ _ = error "Parameters format";
+
+
+(*====================================================*)
+(*Interpretation function for the evaluation protokol *)
+(*====================================================*)
+
+fun proof_of_evalc sg fm =
+let
+fun proof_of_evalch prt = case prt of
+  EvalAt(at) => prove_elementar sg "ss" at
+ |Evalfm(fm) => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl
+ |EvalConst(s,pr1,pr2) => 
+   let val th1 = proof_of_evalch pr1
+       val th2 = proof_of_evalch pr2
+   in case s of
+     "CJ" =>[th1,th2] MRS (qe_conjI)
+    |"DJ" =>[th1,th2] MRS (qe_disjI)
+    |"IM" =>[th1,th2] MRS (qe_impI)
+    |"EQ" =>[th1,th2] MRS (qe_eqI)
+    end
+in proof_of_evalch (evalc_wp fm)
+end;
+
+(*============================================================*)
+(*Interpretation function for the NNF-Transformation protokol *)
+(*============================================================*)
+
+fun proof_of_cnnf sg fm pf = 
+let fun proof_of_cnnfh prt pat = case prt of
+  NNFAt(at) => pat at
+ |NNFSimp (pr) => let val th1 = proof_of_cnnfh pr pat
+                  in let val fm2 = snd (qe_get_terms th1) 
+		     in [th1,prove_elementar sg "ss" (HOLogic.mk_eq(fm2 ,simpl fm2))] MRS trans
+                     end
+                  end
+ |NNFNN (pr) => (proof_of_cnnfh pr pat) RS (nnf_nn)
+ |NNFConst (s,pr1,pr2) =>
+   let val th1 = proof_of_cnnfh pr1 pat
+       val th2 = proof_of_cnnfh pr2 pat
+   in case s of
+     "CJ" => [th1,th2] MRS (qe_conjI)
+    |"DJ" => [th1,th2] MRS (qe_disjI)
+    |"IM" => [th1,th2] MRS (nnf_im)
+    |"EQ" => [th1,th2] MRS (nnf_eq)
+    |"SDJ" => let val (Const("op &",_)$A$_) = fst (qe_get_terms th1)
+	          val (Const("op &",_)$C$_) = fst (qe_get_terms th2)
+	      in [th1,th2,prove_elementar sg "ss" (HOLogic.mk_eq (A,HOLogic.Not $ C))] MRS (nnf_sdj)
+	      end
+    |"NCJ" => [th1,th2] MRS (nnf_ncj)
+    |"NIM" => [th1,th2] MRS (nnf_nim)
+    |"NEQ" => [th1,th2] MRS (nnf_neq)
+    |"NDJ" => [th1,th2] MRS (nnf_ndj)
+   end
+in proof_of_cnnfh (cnnf_wp fm) pf
+end;
+
+
+
+
+(*====================================================*)
+(* Interpretation function for the linform protokol   *)
+(*====================================================*)
+
+
+fun proof_of_linform sg vars f = 
+  let fun proof_of_linformh prt = 
+  case prt of
+    (LfAt (at)) =>  prove_elementar sg "lf" (HOLogic.mk_eq (at, linform vars at))
+   |(LfAtdvd (Const("Divides.op dvd",_)$d$t)) => (prove_elementar sg "lf" (HOLogic.mk_eq (t, lint vars t))) RS (instantiate' [] [None , None, Some (cterm_of sg d)](linearize_dvd))
+   |(Lffm (fm)) => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl)
+   |(LfConst (s,pr1,pr2)) =>
+     let val th1 = proof_of_linformh pr1
+	 val th2 = proof_of_linformh pr2
+     in case s of
+       "CJ" => [th1,th2] MRS (qe_conjI)
+      |"DJ" =>[th1,th2] MRS (qe_disjI)
+      |"IM" =>[th1,th2] MRS (qe_impI)
+      |"EQ" =>[th1,th2] MRS (qe_eqI)
+     end
+   |(LfNot(pr)) => 
+     let val th = proof_of_linformh pr
+     in (th RS (qe_Not))
+     end
+   |(LfQ(s,xn,xT,pr)) => 
+     let val th = forall_intr (cterm_of sg (Free(xn,xT)))(proof_of_linformh pr)
+     in if s = "Ex" 
+        then (th COMP(qe_exI) )
+        else (th COMP(qe_ALLI) )
+     end
+in
+ proof_of_linformh (linform_wp f)
+end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Presburger/presburger.ML	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,231 @@
+(*  Title:      HOL/Integ/presburger.ML
+    ID:         $Id$
+    Author:     Amine Chaieb and Stefan Berghofer, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+Tactic for solving arithmetical Goals in Presburger Arithmetic
+*)
+
+signature PRESBURGER = 
+sig
+ val presburger_tac : bool -> int -> tactic
+ val presburger_method : bool -> int -> Proof.method
+ val setup : (theory -> theory) list
+ val trace : bool ref
+end;
+
+structure Presburger: PRESBURGER =
+struct
+
+val trace = ref false;
+fun trace_msg s = if !trace then tracing s else ();
+
+(*-----------------------------------------------------------------*)
+(*cooper_pp: provefunction for the one-exstance quantifier elimination*)
+(* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*)
+(*-----------------------------------------------------------------*)
+
+val presburger_ss = simpset_of (theory "Presburger");
+
+fun cooper_pp sg vrl (fm as e$Abs(xn,xT,p)) = 
+  let val (xn1,p1) = variant_abs (xn,xT,p)
+  in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1 vrl) end;
+
+fun mnnf_pp sg fm = CooperProof.proof_of_cnnf sg fm
+  (CooperProof.proof_of_evalc sg);
+
+fun mproof_of_int_qelim sg fm =
+  Qelim.proof_of_mlift_qelim sg CooperDec.is_arith_rel
+    (CooperProof.proof_of_linform sg) (mnnf_pp sg) (cooper_pp sg) fm;
+
+(* Theorems to be used in this tactic*)
+
+val zdvd_int = thm "zdvd_int";
+val zdiff_int_split = thm "zdiff_int_split";
+val all_nat = thm "all_nat";
+val ex_nat = thm "ex_nat";
+val number_of1 = thm "number_of1";
+val number_of2 = thm "number_of2";
+val split_zdiv = thm "split_zdiv";
+val split_zmod = thm "split_zmod";
+val mod_div_equality' = thm "mod_div_equality'";
+val split_div' = thm "split_div'";
+val Suc_plus1 = thm "Suc_plus1";
+val imp_le_cong = thm "imp_le_cong";
+val conj_le_cong = thm "conj_le_cong";
+
+(* extract all the constants in a term*)
+fun add_term_typed_consts (Const (c, T), cs) = (c,T) ins cs
+  | add_term_typed_consts (t $ u, cs) =
+      add_term_typed_consts (t, add_term_typed_consts (u, cs))
+  | add_term_typed_consts (Abs (_, _, t), cs) = add_term_typed_consts (t, cs)
+  | add_term_typed_consts (_, cs) = cs;
+
+fun term_typed_consts t = add_term_typed_consts(t,[]);
+
+(* put a term into eta long beta normal form *)
+fun eta_long Ts (Abs (s, T, t)) = Abs (s, T, eta_long (T :: Ts) t)
+  | eta_long Ts t = (case strip_comb t of
+      (Abs _, _) => eta_long Ts (Envir.beta_norm t)
+    | (u, ts) => 
+      let val Us = binder_types (fastype_of1 (Ts, t))
+      in list_abs (map (pair "x") Us, Unify.combound
+        (list_comb (u, map (eta_long Ts) ts), 0, length Us))
+      end);
+
+(* Some Types*)
+val bT = HOLogic.boolT;
+val iT = HOLogic.intT;
+val binT = HOLogic.binT;
+val nT = HOLogic.natT;
+
+(* Allowed Consts in formulae for presburger tactic*)
+
+val allowed_consts =
+  [("All", (iT --> bT) --> bT),
+   ("Ex", (iT --> bT) --> bT),
+   ("All", (nT --> bT) --> bT),
+   ("Ex", (nT --> bT) --> bT),
+
+   ("op &", bT --> bT --> bT),
+   ("op |", bT --> bT --> bT),
+   ("op -->", bT --> bT --> bT),
+   ("op =", bT --> bT --> bT),
+   ("Not", bT --> bT),
+
+   ("op <=", iT --> iT --> bT),
+   ("op =", iT --> iT --> bT),
+   ("op <", iT --> iT --> bT),
+   ("Divides.op dvd", iT --> iT --> bT),
+   ("Divides.op div", iT --> iT --> iT),
+   ("Divides.op mod", iT --> iT --> iT),
+   ("op +", iT --> iT --> iT),
+   ("op -", iT --> iT --> iT),
+   ("op *", iT --> iT --> iT), 
+   ("HOL.abs", iT --> iT),
+   ("uminus", iT --> iT),
+
+   ("op <=", nT --> nT --> bT),
+   ("op =", nT --> nT --> bT),
+   ("op <", nT --> nT --> bT),
+   ("Divides.op dvd", nT --> nT --> bT),
+   ("Divides.op div", nT --> nT --> nT),
+   ("Divides.op mod", nT --> nT --> nT),
+   ("op +", nT --> nT --> nT),
+   ("op -", nT --> nT --> nT),
+   ("op *", nT --> nT --> nT), 
+   ("Suc", nT --> nT),
+
+   ("Numeral.bin.Bit", binT --> bT --> binT),
+   ("Numeral.bin.Pls", binT),
+   ("Numeral.bin.Min", binT),
+   ("Numeral.number_of", binT --> iT),
+   ("Numeral.number_of", binT --> nT),
+   ("0", nT),
+   ("0", iT),
+   ("1", nT),
+   ("1", iT),
+
+   ("False", bT),
+   ("True", bT)];
+
+(*returns true if the formula is relevant for presburger arithmetic tactic*)
+fun relevant t = (term_typed_consts t) subset allowed_consts;
+
+(* Preparation of the formula to be sent to the Integer quantifier *)
+(* elimination procedure                                           *)
+(* Transforms meta implications and meta quantifiers to object     *)
+(* implications and object quantifiers                             *)
+
+fun prepare_for_presburger q fm = 
+  let
+    val ps = Logic.strip_params fm
+    val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm)
+    val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl fm)
+    val _ = if relevant c then () else raise CooperDec.COOPER
+    fun mk_all ((s, T), (P,n)) =
+      if 0 mem loose_bnos P then
+        (HOLogic.all_const T $ Abs (s, T, P), n)
+      else (incr_boundvars ~1 P, n-1)
+    fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t;
+    val (rhs,irhs) = partition relevant hs
+    val np = length ps
+    val (fm',np) =  foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n)))
+      (ps,(foldr HOLogic.mk_imp (rhs, c), np))
+    val (vs, _) = partition (fn t => q orelse (type_of t) = nT)
+      (term_frees fm' @ term_vars fm');
+    val fm2 = foldr mk_all2 (vs, fm')
+  in (fm2, np + length vs, length rhs) end;
+
+(*Object quantifier to meta --*)
+fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ;
+
+(* object implication to meta---*)
+fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp;
+
+(* the presburger tactic*)
+fun presburger_tac q i st =
+  let
+    val g = BasisLibrary.List.nth (prems_of st, i - 1);
+    val sg = sign_of_thm st;
+    (* Transform the term*)
+    val (t,np,nh) = prepare_for_presburger q g
+    (* Some simpsets for dealing with mod div abs and nat*)
+
+    val simpset0 = HOL_basic_ss
+      addsimps [mod_div_equality', Suc_plus1]
+      addsplits [split_zdiv, split_zmod, split_div']
+    (* Simp rules for changing (n::int) to int n *)
+    val simpset1 = HOL_basic_ss
+      addsimps [nat_number_of_def, zdvd_int] @ map (fn r => r RS sym)
+        [int_int_eq, zle_int, zless_int, zadd_int, zmult_int]
+      addsplits [zdiff_int_split]
+    (*simp rules for elimination of int n*)
+
+    val simpset2 = HOL_basic_ss
+      addsimps [nat_0_le, all_nat, ex_nat, number_of1, number_of2, int_0, int_1]
+      addcongs [conj_le_cong, imp_le_cong]
+    (* simp rules for elimination of abs *)
+
+    val simpset3 = HOL_basic_ss addsplits [zabs_split]
+	      
+    val ct = cterm_of sg (HOLogic.mk_Trueprop t)
+
+    (* Theorem for the nat --> int transformation *)
+    val pre_thm = Seq.hd (EVERY
+      [simp_tac simpset0 i,
+       TRY (simp_tac simpset1 i), TRY (simp_tac simpset2 i),
+       TRY (simp_tac simpset3 i), TRY (simp_tac presburger_ss i)]
+      (trivial ct))
+
+    fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i);
+
+    (* The result of the quantifier elimination *)
+    val (th, tac) = case (prop_of pre_thm) of
+        Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
+          (trace_msg ("calling procedure with term:\n" ^
+             Sign.string_of_term sg t1);
+           ((mproof_of_int_qelim sg (eta_long [] t1) RS iffD2) RS pre_thm,
+            assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i)))
+      | _ => (pre_thm, assm_tac i)
+  in (rtac (((mp_step nh) o (spec_step np)) th) i THEN tac) st
+  end handle Subscript => no_tac st | CooperDec.COOPER => no_tac st;
+
+fun presburger_args meth =
+  Method.simple_args (Scan.optional (Args.$$$ "no_quantify" >> K false) true)
+    (fn q => fn _ => meth q 1);
+
+fun presburger_method q i = Method.METHOD (fn facts =>
+  Method.insert_tac facts 1 THEN presburger_tac q i)
+
+val setup =
+  [Method.add_method ("presburger",
+     presburger_args presburger_method,
+     "decision procedure for Presburger arithmetic"),
+   ArithTheoryData.map (fn {splits, inj_consts, discrete, presburger} =>
+     {splits = splits, inj_consts = inj_consts, discrete = discrete,
+      presburger = Some (presburger_tac true)})];
+
+end;
+
+val presburger_tac = Presburger.presburger_tac true;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Presburger/qelim.ML	Tue Mar 25 09:47:05 2003 +0100
@@ -0,0 +1,176 @@
+(*  Title:      HOL/Integ/qelim.ML
+    ID:         $Id$
+    Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen
+    License:    GPL (GNU GENERAL PUBLIC LICENSE)
+
+File containing the implementation of the proof protocoling
+and proof generation for multiple quantified formulae.
+*)
+
+signature QELIM =
+sig
+ val proof_of_mlift_qelim: Sign.sg -> (term -> bool) ->
+   (string list -> term -> thm) -> (term -> thm) ->
+   (string list -> term -> thm) -> term -> thm
+end;
+
+structure Qelim : QELIM =
+struct
+open CooperDec;
+open CooperProof;
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---         Protocoling part                                  ---*)
+(*---                                                           ---*)
+(*---           includes the protocolling datastructure         ---*)
+(*---                                                           ---*)
+(*---          and the protocolling fuctions                    ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+
+val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT;
+
+(* List of the theorems to be used in the following*)
+
+val qe_ex_conj = thm "qe_ex_conj";
+val qe_ex_nconj = thm "qe_ex_nconj";
+val qe_ALL = thm "qe_ALL";
+
+
+(*Datatype declaration for the protocoling procedure.*)
+
+
+datatype QeLog = AFN of term*(string list)
+                |QFN of term*(string list)
+                |ExConj of term*QeLog
+                |ExDisj of (string*typ*term)*term*QeLog*QeLog
+                |QeConst of string*QeLog*QeLog
+                |QeNot of QeLog
+                |QeAll of QeLog
+                |Lift_Qelim of term*QeLog
+		|QeUnk of term;
+
+datatype mQeLog = mQeProof of (string list)*mQeLog
+		|mAFN of term
+                |mNFN of mQeLog
+                |mQeConst of string*mQeLog*mQeLog
+                |mQeNot of mQeLog
+		|mQelim of term*(string list)*mQeLog
+		|mQeAll of mQeLog
+		|mQefm of term;
+
+(* This is the protokoling my function for the quantifier elimination*)
+fun mlift_qelim_wp isat vars = 
+  let fun mqelift_wp vars fm = if (isat fm) then mAFN(fm)
+    else  
+    (case fm of 
+     ( Const ("Not",_) $ p) => mQeNot(mqelift_wp vars p)
+    |( Const ("op &",_) $ p $q) => mQeConst("CJ", mqelift_wp vars p,mqelift_wp vars q) 
+
+    |( Const ("op |",_) $ p $q) => mQeConst("DJ", mqelift_wp vars p,mqelift_wp vars q) 
+
+    |( Const ("op -->",_) $ p $q) => mQeConst("IM", mqelift_wp vars p,mqelift_wp vars q) 
+
+    |( Const ("op =",Type ("fun",[Type ("bool", []),_])) $ p $q) =>mQeConst("EQ", mqelift_wp vars p,mqelift_wp vars q) 
+
+    				
+    |( Const ("All",QT) $ Abs(x,T,p)) =>mQeAll (mqelift_wp vars (Const("Ex",QT) $ Abs(x,T,(HOLogic.Not $ p))))
+
+    |(Const ("Ex",_) $ Abs (x,T,p))  => 
+      let val (x1,p1) = variant_abs (x,T,p)
+          val prt = mqelift_wp (x1::vars) p1
+      in mQelim(Free(x1,T),vars,mNFN(prt))
+      end
+    | _ => mQefm(fm)
+   ) 
+     
+  in (fn fm => mQeProof(vars,mNFN(mqelift_wp vars fm )))
+  end;  
+
+
+
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---      Interpretation and Proofgeneration Part              ---*)
+(*---                                                           ---*)
+(*---      Protocole interpretation functions                   ---*)
+(*---                                                           ---*)
+(*---      and proofgeneration functions                        ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*---                                                           ---*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+(*function that interpretates the protokol generated by the _wp function*)
+
+
+(* proof_of_lift_qelim interpretates a protokol for the quantifier elimination one some quantified formula. It uses the functions afnp nfnp and qfnp as proof functions to generate a prove for the hole quantifier elimination.*)
+(* afnp : must retun a proof for the transformations on the atomic formalae*)
+(* nfnp : must return a proof for the post one-quatifiers elimination process*)
+(* qfnp mus return a proof for the one quantifier elimination (existential) *)
+(* All these function are independent of the theory on whitch we are trying to prove quantifier elimination*)
+(* But the following invariants mus be respected : *)
+(* afnp : term -> string list -> thm*)
+(*   nfnp : term -> thm*)
+(*   qfnp : term -> string list -> thm*)
+(*For all theorms generated by these function must hold :*)
+(*    All of them are logical equivalences.*)
+(*    on left side of the equivalence must appear the term exactely as ist was given as a parameter (or eventually modulo Gamma, where Gamma are the rules whitch are allowed to be used during unification ex. beta reduction.....)*)
+(* qfnp must take as an argument for the term an existential quantified formula*)
+(*-----------------------------------------------------------------*)
+(*-----------------------------------------------------------------*)
+
+fun proof_of_mlift_qelim sg isat afnp nfnp qfnp =
+ let fun h_proof_of_mlift_qelim isat afnp nfnp qfnp prtkl vrl = 
+   (case prtkl of 
+   mAFN (fm) => afnp vrl fm 
+   |mNFN (prt) => let val th1 = h_proof_of_mlift_qelim isat  afnp nfnp qfnp prt vrl 
+                  val th2 = nfnp (snd (qe_get_terms th1)) 
+                    in [th1,th2] MRS trans 
+                 end 
+   |mQeConst (s,pr1,pr2) =>  
+     let val th1 =  h_proof_of_mlift_qelim isat afnp nfnp qfnp pr1 vrl  
+         val th2 =  h_proof_of_mlift_qelim isat afnp nfnp qfnp pr2 vrl  
+     in (case s of 
+        "CJ" => [th1,th2] MRS (qe_conjI)  
+       |"DJ" => [th1,th2] MRS (qe_disjI)  
+       |"IM" => [th1,th2] MRS (qe_impI)  
+       |"EQ" => [th1,th2] MRS (qe_eqI)  
+       )  
+    end  
+   |mQeNot (pr) =>(h_proof_of_mlift_qelim isat afnp nfnp qfnp pr vrl ) RS (qe_Not)  
+   |mQeAll(pr) => (h_proof_of_mlift_qelim isat afnp nfnp qfnp pr vrl ) RS (qe_ALL) 
+   |mQelim (x as (Free(xn,xT)),vl,pr) => 
+     let val th_1 = h_proof_of_mlift_qelim isat afnp nfnp qfnp pr vl
+         val mQeProof(l2,pr2) = mlift_qelim_wp isat (xn::vrl) (snd(qe_get_terms th_1))
+         val th_2 = [th_1,(h_proof_of_mlift_qelim isat afnp nfnp qfnp pr2 l2)] MRS trans
+         val th1 = (forall_intr (cterm_of sg x) th_2) COMP (qe_exI)
+	 val th2 = qfnp vl (snd (qe_get_terms th1)) 
+       in [th1,th2] MRS trans 
+       end
+   |mQefm (fm) => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl
+)
+in (fn fm => let val mQeProof(vars,prt) = (mlift_qelim_wp isat (fv fm) fm)
+                 in (h_proof_of_mlift_qelim isat afnp nfnp qfnp prt vars)
+                 end)
+end;
+
+end;