src/HOL/Integ/cooper_proof.ML
author haftmann
Wed, 13 Dec 2006 15:45:31 +0100
changeset 21820 2f2b6a965ccc
parent 21416 f23e4e75dfd3
child 22997 d4f3b015b50b
permissions -rw-r--r--
introduced mk/dest_numeral/number for mk/dest_binum etc.

(*  Title:      HOL/Integ/cooper_proof.ML
    ID:         $Id$
    Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen

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 list_to_set : typ -> term list -> term
  val qe_get_terms : thm -> term * term
  val cooper_prv  : theory -> term -> term -> thm
  val proof_of_evalc : theory -> term -> thm
  val proof_of_cnnf : theory -> term -> (term -> thm) -> thm
  val proof_of_linform : theory -> string list -> term -> thm
  val proof_of_adjustcoeffeq : theory -> term -> IntInf.int -> term -> thm
  val prove_elementar : theory -> string -> term -> thm
  val thm_of : theory -> (term -> (term list * (thm list -> thm))) -> term -> thm
end;

structure CooperProof : COOPER_PROOF =
struct
open CooperDec;

val presburger_ss = simpset ()
  addsimps [diff_int_def] delsimps [thm "diff_int_def_symmetric"];

val cboolT = ctyp_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";

(* Theorems for proving the adjustment of the coefficients*)

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";

(*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";


(* ------------------------------------------------------------------------- *)
(*This function norm_zero_one  replaces the occurences of Numeral1 and Numeral0*)
(*Respectively by their abstract representation Const("HOL.one",..) and Const("HOL.zero",..)*)
(*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 ("HOL.times",_) $ c $ t) => 
    if c = one then (norm_zero_one t)
    else if (dest_number c = ~1) 
         then (Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $ (norm_zero_one t))
         else (HOLogic.mk_binop "HOL.times" (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;

(* ------------------------------------------------------------------------- *)
(*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;
		
(* ------------------------------------------------------------------------- *)
(* 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;

(* ------------------------------------------------------------------------- *)
(*This function proove elementar will be used to generate proofs at
  runtime*) (*It is thought to prove properties such as a dvd b
  (essentially) that are only to make at runtime.*)
(* ------------------------------------------------------------------------- *)
fun prove_elementar thy s fm2 =
  Goal.prove (ProofContext.init thy) [] [] (HOLogic.mk_Trueprop fm2) (fn _ => EVERY
  (case s of
  (*"ss" like simplification with simpset*)
  "ss" =>
    let val ss = presburger_ss addsimps [zdvd_iff_zmod_eq_0,unity_coeff_ex]
    in [simp_tac ss 1, TRY (simple_arith_tac 1)] end

  (*"bl" like blast tactic*)
  (* Is only used in the harrisons like proof procedure *)
  | "bl" => [blast_tac HOL_cs 1]

  (*"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
    in ex_disj_tacs end

  | "fa" => [simple_arith_tac 1]

  | "sa" =>
    let val ss = presburger_ss addsimps zadd_ac
    in [simp_tac ss 1, TRY (simple_arith_tac 1)] end

  (* like Existance Conjunction *)
  | "ec" =>
    let val ss = presburger_ss addsimps zadd_ac
    in [simp_tac ss 1, TRY (blast_tac HOL_cs 1)] end

  | "ac" =>
    let val ss = HOL_basic_ss addsimps zadd_ac
    in [simp_tac ss 1] end

  | "lf" =>
    let val ss = presburger_ss addsimps zadd_ac
    in [simp_tac ss 1, TRY (simple_arith_tac 1)] end));

(*=============================================================*)
(*-------------------------------------------------------------*)
(*              The new compact model                          *)
(*-------------------------------------------------------------*)
(*=============================================================*)

fun thm_of sg decomp t = 
    let val (ts,recomb) = decomp t 
    in recomb (map (thm_of sg decomp) ts) 
    end;

(*==================================================*)
(*     Compact Version for adjustcoeffeq            *)
(*==================================================*)

fun decomp_adjustcoeffeq sg x l fm = case fm of
    (Const("Not",_)$(Const("Orderings.less",_) $(Const("HOL.zero",_)) $(rt as (Const ("HOL.plus", _)$(Const ("HOL.times",_) $    c $ y ) $z )))) => 
     let  
        val m = l div (dest_number c) 
        val n = if (x = y) then abs (m) else 1
        val xtm = (HOLogic.mk_binop "HOL.times" ((mk_number ((m div n)*l) ), x)) 
        val rs = if (x = y) 
                 then (HOLogic.mk_binrel "Orderings.less" (zero,linear_sub [] (mk_number n) (HOLogic.mk_binop "HOL.plus" ( xtm ,( linear_cmul n z) )))) 
                 else HOLogic.mk_binrel "Orderings.less" (zero,linear_sub [] one rt )
        val ck = cterm_of sg (mk_number n)
        val cc = cterm_of sg c
        val ct = cterm_of sg z
        val cx = cterm_of sg y
        val pre = prove_elementar sg "lf" 
            (HOLogic.mk_binrel "Orderings.less" (Const("HOL.zero",HOLogic.intT),(mk_number n)))
        val th1 = (pre RS (instantiate' [] [SOME ck,SOME cc, SOME cx, SOME ct] (ac_pi_eq)))
        in ([], fn [] => [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
        end

  |(Const(p,_) $a $( Const ("HOL.plus", _)$(Const ("HOL.times",_) $ 
      c $ y ) $t )) => 
   if (is_arith_rel fm) andalso (x = y) 
   then  
        let val m = l div (dest_number c) 
           val k = (if p = "Orderings.less" then abs(m) else m)  
           val xtm = (HOLogic.mk_binop "HOL.times" ((mk_number ((m div k)*l) ), x))
           val rs = (HOLogic.mk_binrel p ((linear_cmul k a),(HOLogic.mk_binop "HOL.plus" ( xtm ,( linear_cmul k t) )))) 

           val ck = cterm_of sg (mk_number 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
	  "Orderings.less" => 
	let val pre = prove_elementar sg "lf" 
	    (HOLogic.mk_binrel "Orderings.less" (Const("HOL.zero",HOLogic.intT),(mk_number k)))
            val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_lt_eq)))
	in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
         end

           |"op =" =>
	     let val pre = prove_elementar sg "lf" 
	    (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("HOL.zero",HOLogic.intT),(mk_number k))))
	         val th1 = (pre RS(instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_eq_eq)))
	     in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
             end

             |"Divides.dvd" =>
	       let val pre = prove_elementar sg "lf" 
	   (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("HOL.zero",HOLogic.intT),(mk_number k))))
                   val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct]) (ac_dvd_eq))
               in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
                        
               end
              end
  else ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl)

 |( Const ("Not", _) $ p) => ([p], fn [th] => th RS qe_Not)
  |( Const ("op &",_) $ p $ q) => ([p,q], fn [th1,th2] => [th1,th2] MRS qe_conjI)
  |( Const ("op |",_) $ p $ q) =>([p,q], fn [th1,th2] => [th1,th2] MRS qe_disjI)

  |_ => ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl);

fun proof_of_adjustcoeffeq sg x l = thm_of sg (decomp_adjustcoeffeq sg x l);



(*==================================================*)
(*   Finding rho for modd_minusinfinity             *)
(*==================================================*)
fun rho_for_modd_minf x dlcm sg 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 ("HOL.plus", _) $(Const ("HOL.times",_) $ 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 ("HOL.plus", _) $(Const ("HOL.times",_) $ 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("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ 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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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.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.dvd",_)$ d $ (db as (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $
      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.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;	 
(*=========================================================================*)
(*=========================================================================*)
fun rho_for_eq_minf x dlcm  sg fm1 =  
   let
   val fm = norm_zero_one fm1
    in  case fm1 of 
      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("HOL.plus", _) $(Const ("HOL.times",_) $ 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 ("HOL.plus", _) $(Const ("HOL.times",_) $ 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("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ 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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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;

(*=====================================================*)
(*=====================================================*)
(*=========== minf proofs with the compact version==========*)
fun decomp_minf_eq x dlcm sg t =  case t of
   Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_conjI)
   |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_disjI)
   |_ => ([],fn [] => rho_for_eq_minf x dlcm sg t);

fun decomp_minf_modd x dlcm sg t = case t of
   Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_conjI)
   |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_disjI)
   |_ => ([],fn [] => rho_for_modd_minf x dlcm sg t);

(* -------------------------------------------------------------*)
(*                    Finding rho for pinf_modd                 *)
(* -------------------------------------------------------------*)
fun rho_for_modd_pinf x dlcm sg 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 ("HOL.plus", _) $(Const ("HOL.times",_) $ 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 ("HOL.plus", _) $(Const ("HOL.times",_) $ 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("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ 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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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.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.dvd",_)$ d $ (db as (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $
      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.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;	
(* -------------------------------------------------------------*)
(*                    Finding rho for pinf_eq                 *)
(* -------------------------------------------------------------*)
fun rho_for_eq_pinf x dlcm sg fm1 = 
  let
					val fm = norm_zero_one fm1
    in  case fm1 of 
      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("HOL.plus", _) $(Const ("HOL.times",_) $ 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 ("HOL.plus", _) $(Const ("HOL.times",_) $ 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("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ 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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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;



fun  minf_proof_of_c sg x dlcm t =
  let val minf_eqth   = thm_of sg (decomp_minf_eq x dlcm sg) t
      val minf_moddth = thm_of sg (decomp_minf_modd x dlcm sg) t
  in (minf_eqth, minf_moddth)
end;

(*=========== pinf proofs with the compact version==========*)
fun decomp_pinf_eq x dlcm sg t = case t of
   Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_conjI)
   |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_disjI)
   |_ =>([],fn [] => rho_for_eq_pinf x dlcm sg t) ;

fun decomp_pinf_modd x dlcm sg t =  case t of
   Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_conjI)
   |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_disjI)
   |_ => ([],fn [] => rho_for_modd_pinf x dlcm sg t);

fun  pinf_proof_of_c sg x dlcm t =
  let val pinf_eqth   = thm_of sg (decomp_pinf_eq x dlcm sg) t
      val pinf_moddth = thm_of sg (decomp_pinf_modd x dlcm sg) t
  in (pinf_eqth,pinf_moddth)
end;


(* ------------------------------------------------------------------------- *)
(* Here we generate the theorem for the Bset Property in the simple direction*)
(* It is just an instantiation*)
(* ------------------------------------------------------------------------- *)
(*
fun bsetproof_of sg (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;

fun asetproof_of sg (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;
*)

(* For the generation of atomic Theorems*)
(* Prove the premisses on runtime and then make RS*)
(* ------------------------------------------------------------------------- *)

(*========= this is rho ============*)
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 ("HOL.plus", _) $(Const ("HOL.times",_) $ 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("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("HOL.zero",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 ("HOL.plus", T) $(Const ("HOL.times",_) $ 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_number 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("HOL.minus",T) $ (Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("HOL.one",HOLogic.intT))))
		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("HOL.zero",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("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ 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("HOL.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 "Orderings.less" (Const("HOL.zero",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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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.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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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.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*)
(* ------------------------------------------------------------------------- *)

(*==================== Proof with the compact version   *)

fun decomp_nbstp sg x dlcm B fm t = case t of 
   Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_conjI )
  |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_disjI)
  |_ => ([], fn [] => generate_atomic_not_bst_p sg x fm dlcm B t);

fun not_bst_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm B t =
  let 
       val th =  thm_of sg (decomp_nbstp sg x dlcm (list_to_set xT (map norm_zero_one B)) fm) t
      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*)
(* ------------------------------------------------------------------------- *)
(*========= this is rho ============*)
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 ("HOL.plus", _) $(Const ("HOL.times",_) $ 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("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("HOL.zero",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 ("HOL.plus", T) $(Const ("HOL.times",_) $ 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("HOL.plus",T) $ (Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("HOL.one",HOLogic.intT))))
		 val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("HOL.zero",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("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ pm1 $ y ) $ z )) =>
        if (y=x) andalso (c1 =zero) then 
        if pm1 = (mk_number ~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 "Orderings.less" (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 "Orderings.less" (Const("HOL.zero",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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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.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.dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ 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.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 decomp_nastp sg x dlcm A fm t = case t of 
   Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_conjI )
  |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_disjI)
  |_ => ([], fn [] => generate_atomic_not_ast_p sg x fm dlcm A t);

fun not_ast_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm A t =
  let 
       val th =  thm_of sg (decomp_nastp sg x dlcm (list_to_set xT (map norm_zero_one A)) fm) t
      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_ast_p_Q_elim)
     end
  end;


(* -------------------------------*)
(* Finding rho and beta for evalc *)
(* -------------------------------*)

fun rho_for_evalc sg at = case at of  
    (Const (p,_) $ s $ t) =>(  
    case AList.lookup (op =) operations p of 
        SOME f => 
           ((if (f ((dest_number s),(dest_number t))) 
             then prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)) 
             else prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const)))  
		   handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl)
        | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl )
     |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
       case AList.lookup (op =) operations p of 
         SOME f => 
           ((if (f ((dest_number s),(dest_number t))) 
             then prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const))  
             else prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)))  
		      handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl) 
         | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl ) 
     | _ =>   instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl;


(*=========================================================*)
fun decomp_evalc sg t = case t of
   (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI)
   |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI)
   |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI)
   |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
   |_ => ([], fn [] => rho_for_evalc sg t);


fun proof_of_evalc sg fm = thm_of sg (decomp_evalc sg) fm;

(*==================================================*)
(*     Proof of linform with the compact model      *)
(*==================================================*)


fun decomp_linform sg vars t = case t of
   (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI)
   |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI)
   |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI)
   |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
   |(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not)
   |(Const("Divides.dvd",_)$d$r) => 
     if is_numeral d then ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [NONE , NONE, SOME (cterm_of sg d)](linearize_dvd)))
     else (warning "Nonlinear Term --- Non numeral leftside at dvd";
       raise COOPER)
   |_ => ([], fn [] => prove_elementar sg "lf" (HOLogic.mk_eq (t, linform vars t)));

fun proof_of_linform sg vars f = thm_of sg (decomp_linform sg vars) f;

(* ------------------------------------------------------------------------- *)
(* Interpretaion of Protocols of the cooper procedure : minusinfinity version*)
(* ------------------------------------------------------------------------- *)
fun coopermi_proof_of sg (x as Free(xn,xT)) fm B dlcm =
  (* Get the Bset thm*)
  let val (minf_eqth, minf_moddth) = minf_proof_of_c sg x dlcm fm 
      val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (zero,dlcm));
      val nbstpthm = not_bst_p_proof_of_c sg x fm dlcm B fm
  in (dpos,minf_eqth,nbstpthm,minf_moddth)
end;

(* ------------------------------------------------------------------------- *)
(* Interpretaion of Protocols of the cooper procedure : plusinfinity version *)
(* ------------------------------------------------------------------------- *)
fun cooperpi_proof_of sg (x as Free(xn,xT)) fm A dlcm =
  let val (pinf_eqth,pinf_moddth) = pinf_proof_of_c sg x dlcm fm
      val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (zero,dlcm));
      val nastpthm = not_ast_p_proof_of_c sg x fm dlcm A fm
  in (dpos,pinf_eqth,nastpthm,pinf_moddth)
end;

(* ------------------------------------------------------------------------- *)
(* Interpretaion of Protocols of the cooper procedure : full version*)
(* ------------------------------------------------------------------------- *)
fun cooper_thm sg s (x as Free(xn,xT)) cfm dlcm ast bst= case s of
  "pi" => let val (dpsthm,pinf_eqth,nbpth,pinf_moddth) = cooperpi_proof_of sg x cfm ast dlcm 
	      in [dpsthm,pinf_eqth,nbpth,pinf_moddth] MRS (cppi_eq)
           end
  |"mi" => let val (dpsthm,minf_eqth,nbpth,minf_moddth) = coopermi_proof_of sg x cfm bst dlcm
	       in [dpsthm,minf_eqth,nbpth,minf_moddth] 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*)
(* ------------------------------------------------------------------------- *)

(* val (timef:(unit->thm) -> thm,prtime,time_reset) = gen_timer();*)
(* val (timef2:(unit->thm) -> thm,prtime2,time_reset2) = gen_timer(); *)

fun cooper_prv sg (x as Free(xn,xT)) efm = let 
   (* lfm_thm : efm = linearized form of efm*)
   val lfm_thm = proof_of_linform sg [xn] efm
   (*efm2 is the linearized form of efm *) 
   val efm2 = snd(qe_get_terms lfm_thm)
   (* l is the lcm of all coefficients of x *)
   val l = formlcm x efm2
   (*ac_thm: efm = efm2 with adjusted coefficients of x *)
   val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans
   (* fm is efm2 with adjusted coefficients of x *)
   val fm = snd (qe_get_terms ac_thm)
  (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*)
   val  cfm = unitycoeff x fm
   (*afm is fm where c*x is replaced by 1*x or -1*x *)
   val afm = adjustcoeff x l fm
   (* P = %x.afm*)
   val P = absfree(xn,xT,afm)
   (* This simpset allows the elimination of the sets in bex {1..d} *)
   val ss = presburger_ss addsimps
     [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
   (* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
   val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_number l))] (unity_coeff_ex)
   (* e_ac_thm : Ex x. efm = EX x. fm*)
   val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
   (* A and B set of the formula*)
   val A = aset x cfm
   val B = bset x cfm
   (* the divlcm (delta) of the formula*)
   val dlcm = mk_number (divlcm x cfm)
   (* Which set is smaller to generate the (hoepfully) shorter proof*)
   val cms = if ((length A) < (length B )) then "pi" else "mi"
(*   val _ = if cms = "pi" then writeln "Plusinfinity" else writeln "Minusinfinity"*)
   (* synthesize the proof of cooper's theorem*)
    (* cp_thm: EX x. cfm = Q*)
   val cp_thm =  cooper_thm sg cms x cfm dlcm A B
   (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*)
   (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*)
(*
   val _ = prth cp_thm
   val _ = writeln "Expanding the bounded EX..."
*)
   val exp_cp_thm = refl RS (simplify ss (cp_thm RSN (2,trans)))
(*
   val _ = writeln "Expanded" *)
   (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*)
   val (lsuth,rsuth) = qe_get_terms (uth)
   (* lseacth = EX x. efm; rseacth = EX x. fm*)
   val (lseacth,rseacth) = qe_get_terms(e_ac_thm)
   (* lscth = EX x. cfm; rscth = Q' *)
   val (lscth,rscth) = qe_get_terms (exp_cp_thm)
   (* u_c_thm: EX x. P(l*x) = Q'*)
   val  u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans
   (* result: EX x. efm = Q'*)
 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";

(* **************************************** *)
(*    An Other Version of cooper proving    *)
(*     by giving a withness for EX          *)
(* **************************************** *)



fun cooper_prv_w sg (x as Free(xn,xT)) efm = let 
   (* lfm_thm : efm = linearized form of efm*)
   val lfm_thm = proof_of_linform sg [xn] efm
   (*efm2 is the linearized form of efm *) 
   val efm2 = snd(qe_get_terms lfm_thm)
   (* l is the lcm of all coefficients of x *)
   val l = formlcm x efm2
   (*ac_thm: efm = efm2 with adjusted coefficients of x *)
   val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans
   (* fm is efm2 with adjusted coefficients of x *)
   val fm = snd (qe_get_terms ac_thm)
  (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*)
   val  cfm = unitycoeff x fm
   (*afm is fm where c*x is replaced by 1*x or -1*x *)
   val afm = adjustcoeff x l fm
   (* P = %x.afm*)
   val P = absfree(xn,xT,afm)
   (* This simpset allows the elimination of the sets in bex {1..d} *)
   val ss = presburger_ss addsimps
     [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
   (* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
   val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_number l))] (unity_coeff_ex)
   (* e_ac_thm : Ex x. efm = EX x. fm*)
   val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
   (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*)
   val (lsuth,rsuth) = qe_get_terms (uth)
   (* lseacth = EX x. efm; rseacth = EX x. fm*)
   val (lseacth,rseacth) = qe_get_terms(e_ac_thm)

   val (w,rs) = cooper_w [] cfm
   val exp_cp_thm =  case w of 
     (* FIXME - e_ac_thm just tipped to test syntactical correctness of the program!!!!*)
    SOME n =>  e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*)
   |_ => let 
    (* A and B set of the formula*)
    val A = aset x cfm
    val B = bset x cfm
    (* the divlcm (delta) of the formula*)
    val dlcm = mk_number (divlcm x cfm)
    (* Which set is smaller to generate the (hoepfully) shorter proof*)
    val cms = if ((length A) < (length B )) then "pi" else "mi"
    (* synthesize the proof of cooper's theorem*)
     (* cp_thm: EX x. cfm = Q*)
    val cp_thm = cooper_thm sg cms x cfm dlcm A B
     (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*)
    (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*)
    in refl RS (simplify ss (cp_thm RSN (2,trans)))
    end
   (* lscth = EX x. cfm; rscth = Q' *)
   val (lscth,rscth) = qe_get_terms (exp_cp_thm)
   (* u_c_thm: EX x. P(l*x) = Q'*)
   val  u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans
   (* result: EX x. efm = Q'*)
 in  ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans)
   end
|cooper_prv_w _ _ _ =  error "Parameters format";



fun decomp_cnnf sg lfnp P = case P of 
     Const ("op &",_) $ p $q => ([p,q] , fn [th1,th2] => [th1,th2] MRS qe_conjI )
   |Const ("op |",_) $ p $q => ([p,q] , fn [th1,th2] => [th1,th2] MRS  qe_disjI)
   |Const ("Not",_) $ (Const("Not",_) $ p) => ([p], fn [th] => th RS nnf_nn)
   |Const("Not",_) $ (Const(opn,T) $ p $ q) => 
     if opn = "op |" 
      then case (p,q) of 
         (A as (Const ("op &",_) $ r $ s),B as (Const ("op &",_) $ r1 $ t)) =>
          if r1 = negate r 
          then  ([r,HOLogic.Not$s,r1,HOLogic.Not$t],fn [th1_1,th1_2,th2_1,th2_2] => [[th1_1,th1_1] MRS qe_conjI,[th2_1,th2_2] MRS qe_conjI] MRS nnf_sdj)

          else ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj)
        |(_,_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj)
      else (
         case (opn,T) of 
           ("op &",_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_ncj )
           |("op -->",_) => ([p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_nim )
           |("op =",Type ("fun",[Type ("bool", []),_])) => 
           ([HOLogic.conj $ p $ (HOLogic.Not $ q),HOLogic.conj $ (HOLogic.Not $ p) $ q], fn [th1,th2] => [th1,th2] MRS nnf_neq)
            |(_,_) => ([], fn [] => lfnp P)
)

   |(Const ("op -->",_) $ p $ q) => ([HOLogic.Not$p,q], fn [th1,th2] => [th1,th2] MRS nnf_im)

   |(Const ("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q) =>
     ([HOLogic.conj $ p $ q,HOLogic.conj $ (HOLogic.Not $ p) $ (HOLogic.Not $ q) ], fn [th1,th2] =>[th1,th2] MRS nnf_eq )
   |_ => ([], fn [] => lfnp P);




fun proof_of_cnnf sg p lfnp = 
 let val th1 = thm_of sg (decomp_cnnf sg lfnp) p
     val rs = snd(qe_get_terms th1)
     val th2 = prove_elementar sg "ss" (HOLogic.mk_eq(rs,simpl rs))
  in [th1,th2] MRS trans
  end;

end;