src/HOL/Integ/cooper_proof.ML
author chaieb
Fri, 06 Aug 2004 17:29:24 +0200
changeset 15123 4c49281dc9a8
parent 15122 4b52eeb62807
child 15164 5d7c96e0f9dc
permissions -rw-r--r--
*** empty log message ***
     1 (*  Title:      HOL/Integ/cooper_proof.ML
     2     ID:         $Id$
     3     Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen
     4 
     5 File containing the implementation of the proof
     6 generation for Cooper Algorithm
     7 *)
     8 
     9 
    10 signature COOPER_PROOF =
    11 sig
    12   val qe_Not : thm
    13   val qe_conjI : thm
    14   val qe_disjI : thm
    15   val qe_impI : thm
    16   val qe_eqI : thm
    17   val qe_exI : thm
    18   val list_to_set : typ -> term list -> term
    19   val qe_get_terms : thm -> term * term
    20   val cooper_prv  : Sign.sg -> term -> term -> thm
    21   val proof_of_evalc : Sign.sg -> term -> thm
    22   val proof_of_cnnf : Sign.sg -> term -> (term -> thm) -> thm
    23   val proof_of_linform : Sign.sg -> string list -> term -> thm
    24   val proof_of_adjustcoeffeq : Sign.sg -> term -> int -> term -> thm
    25   val prove_elementar : Sign.sg -> string -> term -> thm
    26   val thm_of : Sign.sg -> (term -> (term list * (thm list -> thm))) -> term -> thm
    27 end;
    28 
    29 structure CooperProof : COOPER_PROOF =
    30 struct
    31 open CooperDec;
    32 
    33 (*
    34 val presburger_ss = simpset_of (theory "Presburger")
    35   addsimps [zdiff_def] delsimps [symmetric zdiff_def];
    36 *)
    37 
    38 val presburger_ss = simpset_of (theory "Presburger")
    39   addsimps[diff_int_def] delsimps [thm"diff_int_def_symmetric"];
    40 val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT;
    41 
    42 (*Theorems that will be used later for the proofgeneration*)
    43 
    44 val zdvd_iff_zmod_eq_0 = thm "zdvd_iff_zmod_eq_0";
    45 val unity_coeff_ex = thm "unity_coeff_ex";
    46 
    47 (* Thorems for proving the adjustment of the coeffitients*)
    48 
    49 val ac_lt_eq =  thm "ac_lt_eq";
    50 val ac_eq_eq = thm "ac_eq_eq";
    51 val ac_dvd_eq = thm "ac_dvd_eq";
    52 val ac_pi_eq = thm "ac_pi_eq";
    53 
    54 (* The logical compination of the sythetised properties*)
    55 val qe_Not = thm "qe_Not";
    56 val qe_conjI = thm "qe_conjI";
    57 val qe_disjI = thm "qe_disjI";
    58 val qe_impI = thm "qe_impI";
    59 val qe_eqI = thm "qe_eqI";
    60 val qe_exI = thm "qe_exI";
    61 val qe_ALLI = thm "qe_ALLI";
    62 
    63 (*Modulo D property for Pminusinf an Plusinf *)
    64 val fm_modd_minf = thm "fm_modd_minf";
    65 val not_dvd_modd_minf = thm "not_dvd_modd_minf";
    66 val dvd_modd_minf = thm "dvd_modd_minf";
    67 
    68 val fm_modd_pinf = thm "fm_modd_pinf";
    69 val not_dvd_modd_pinf = thm "not_dvd_modd_pinf";
    70 val dvd_modd_pinf = thm "dvd_modd_pinf";
    71 
    72 (* the minusinfinity proprty*)
    73 
    74 val fm_eq_minf = thm "fm_eq_minf";
    75 val neq_eq_minf = thm "neq_eq_minf";
    76 val eq_eq_minf = thm "eq_eq_minf";
    77 val le_eq_minf = thm "le_eq_minf";
    78 val len_eq_minf = thm "len_eq_minf";
    79 val not_dvd_eq_minf = thm "not_dvd_eq_minf";
    80 val dvd_eq_minf = thm "dvd_eq_minf";
    81 
    82 (* the Plusinfinity proprty*)
    83 
    84 val fm_eq_pinf = thm "fm_eq_pinf";
    85 val neq_eq_pinf = thm "neq_eq_pinf";
    86 val eq_eq_pinf = thm "eq_eq_pinf";
    87 val le_eq_pinf = thm "le_eq_pinf";
    88 val len_eq_pinf = thm "len_eq_pinf";
    89 val not_dvd_eq_pinf = thm "not_dvd_eq_pinf";
    90 val dvd_eq_pinf = thm "dvd_eq_pinf";
    91 
    92 (*Logical construction of the Property*)
    93 val eq_minf_conjI = thm "eq_minf_conjI";
    94 val eq_minf_disjI = thm "eq_minf_disjI";
    95 val modd_minf_disjI = thm "modd_minf_disjI";
    96 val modd_minf_conjI = thm "modd_minf_conjI";
    97 
    98 val eq_pinf_conjI = thm "eq_pinf_conjI";
    99 val eq_pinf_disjI = thm "eq_pinf_disjI";
   100 val modd_pinf_disjI = thm "modd_pinf_disjI";
   101 val modd_pinf_conjI = thm "modd_pinf_conjI";
   102 
   103 (*Cooper Backwards...*)
   104 (*Bset*)
   105 val not_bst_p_fm = thm "not_bst_p_fm";
   106 val not_bst_p_ne = thm "not_bst_p_ne";
   107 val not_bst_p_eq = thm "not_bst_p_eq";
   108 val not_bst_p_gt = thm "not_bst_p_gt";
   109 val not_bst_p_lt = thm "not_bst_p_lt";
   110 val not_bst_p_ndvd = thm "not_bst_p_ndvd";
   111 val not_bst_p_dvd = thm "not_bst_p_dvd";
   112 
   113 (*Aset*)
   114 val not_ast_p_fm = thm "not_ast_p_fm";
   115 val not_ast_p_ne = thm "not_ast_p_ne";
   116 val not_ast_p_eq = thm "not_ast_p_eq";
   117 val not_ast_p_gt = thm "not_ast_p_gt";
   118 val not_ast_p_lt = thm "not_ast_p_lt";
   119 val not_ast_p_ndvd = thm "not_ast_p_ndvd";
   120 val not_ast_p_dvd = thm "not_ast_p_dvd";
   121 
   122 (*Logical construction of the prop*)
   123 (*Bset*)
   124 val not_bst_p_conjI = thm "not_bst_p_conjI";
   125 val not_bst_p_disjI = thm "not_bst_p_disjI";
   126 val not_bst_p_Q_elim = thm "not_bst_p_Q_elim";
   127 
   128 (*Aset*)
   129 val not_ast_p_conjI = thm "not_ast_p_conjI";
   130 val not_ast_p_disjI = thm "not_ast_p_disjI";
   131 val not_ast_p_Q_elim = thm "not_ast_p_Q_elim";
   132 
   133 (*Cooper*)
   134 val cppi_eq = thm "cppi_eq";
   135 val cpmi_eq = thm "cpmi_eq";
   136 
   137 (*Others*)
   138 val simp_from_to = thm "simp_from_to";
   139 val P_eqtrue = thm "P_eqtrue";
   140 val P_eqfalse = thm "P_eqfalse";
   141 
   142 (*For Proving NNF*)
   143 
   144 val nnf_nn = thm "nnf_nn";
   145 val nnf_im = thm "nnf_im";
   146 val nnf_eq = thm "nnf_eq";
   147 val nnf_sdj = thm "nnf_sdj";
   148 val nnf_ncj = thm "nnf_ncj";
   149 val nnf_nim = thm "nnf_nim";
   150 val nnf_neq = thm "nnf_neq";
   151 val nnf_ndj = thm "nnf_ndj";
   152 
   153 (*For Proving term linearizition*)
   154 val linearize_dvd = thm "linearize_dvd";
   155 val lf_lt = thm "lf_lt";
   156 val lf_eq = thm "lf_eq";
   157 val lf_dvd = thm "lf_dvd";
   158 
   159 
   160 (* ------------------------------------------------------------------------- *)
   161 (*This function norm_zero_one  replaces the occurences of Numeral1 and Numeral0*)
   162 (*Respectively by their abstract representation Const("1",..) and COnst("0",..)*)
   163 (*this is necessary because the theorems use this representation.*)
   164 (* This function should be elminated in next versions...*)
   165 (* ------------------------------------------------------------------------- *)
   166 
   167 fun norm_zero_one fm = case fm of
   168   (Const ("op *",_) $ c $ t) => 
   169     if c = one then (norm_zero_one t)
   170     else if (dest_numeral c = ~1) 
   171          then (Const("uminus",HOLogic.intT --> HOLogic.intT) $ (norm_zero_one t))
   172          else (HOLogic.mk_binop "op *" (norm_zero_one c,norm_zero_one t))
   173   |(node $ rest) => ((norm_zero_one node)$(norm_zero_one rest))
   174   |(Abs(x,T,p)) => (Abs(x,T,(norm_zero_one p)))
   175   |_ => fm;
   176 
   177 (* ------------------------------------------------------------------------- *)
   178 (*function list to Set, constructs a set containing all elements of a given list.*)
   179 (* ------------------------------------------------------------------------- *)
   180 fun list_to_set T1 l = let val T = (HOLogic.mk_setT T1) in 
   181 	case l of 
   182 		[] => Const ("{}",T)
   183 		|(h::t) => Const("insert", T1 --> (T --> T)) $ h $(list_to_set T1 t)
   184 		end;
   185 		
   186 (* ------------------------------------------------------------------------- *)
   187 (* Returns both sides of an equvalence in the theorem*)
   188 (* ------------------------------------------------------------------------- *)
   189 fun qe_get_terms th = let val (_$(Const("op =",Type ("fun",[Type ("bool", []),_])) $ A $ B )) = prop_of th in (A,B) end;
   190 
   191 (* ------------------------------------------------------------------------- *)
   192 (* Modified version of the simple version with minimal amount of checking and postprocessing*)
   193 (* ------------------------------------------------------------------------- *)
   194 
   195 fun simple_prove_goal_cterm2 G tacs =
   196   let
   197     fun check None = error "prove_goal: tactic failed"
   198       | check (Some (thm, _)) = (case nprems_of thm of
   199             0 => thm
   200           | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
   201   in check (Seq.pull (EVERY tacs (trivial G))) end;
   202 
   203 (*-------------------------------------------------------------*)
   204 (*-------------------------------------------------------------*)
   205 
   206 fun cert_Trueprop sg t = cterm_of sg (HOLogic.mk_Trueprop t);
   207 
   208 (* ------------------------------------------------------------------------- *)
   209 (*This function proove elementar will be used to generate proofs at runtime*)
   210 (*It is is based on the isabelle function proove_goalw_cterm and is thought to *)
   211 (*prove properties such as a dvd b (essentially) that are only to make at
   212 runtime.*)
   213 (* ------------------------------------------------------------------------- *)
   214 fun prove_elementar sg s fm2 = case s of 
   215   (*"ss" like simplification with simpset*)
   216   "ss" =>
   217     let
   218       val ss = presburger_ss addsimps
   219         [zdvd_iff_zmod_eq_0,unity_coeff_ex]
   220       val ct =  cert_Trueprop sg fm2
   221     in 
   222       simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)] 
   223 
   224     end
   225 
   226   (*"bl" like blast tactic*)
   227   (* Is only used in the harrisons like proof procedure *)
   228   | "bl" =>
   229      let val ct = cert_Trueprop sg fm2
   230      in
   231        simple_prove_goal_cterm2 ct [blast_tac HOL_cs 1]
   232      end
   233 
   234   (*"ed" like Existence disjunctions ...*)
   235   (* Is only used in the harrisons like proof procedure *)
   236   | "ed" =>
   237     let
   238       val ex_disj_tacs =
   239         let
   240           val tac1 = EVERY[REPEAT(resolve_tac [disjI1,disjI2] 1), etac exI 1]
   241           val tac2 = EVERY[etac exE 1, rtac exI 1,
   242             REPEAT(resolve_tac [disjI1,disjI2] 1), assumption 1]
   243 	in [rtac iffI 1,
   244           etac exE 1, REPEAT(EVERY[etac disjE 1, tac1]), tac1,
   245           REPEAT(EVERY[etac disjE 1, tac2]), tac2]
   246         end
   247 
   248       val ct = cert_Trueprop sg fm2
   249     in 
   250       simple_prove_goal_cterm2 ct ex_disj_tacs
   251     end
   252 
   253   | "fa" =>
   254     let val ct = cert_Trueprop sg fm2
   255     in simple_prove_goal_cterm2 ct [simple_arith_tac 1]  
   256 
   257     end
   258 
   259   | "sa" =>
   260     let
   261       val ss = presburger_ss addsimps zadd_ac
   262       val ct = cert_Trueprop sg fm2
   263     in 
   264       simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]  
   265 
   266     end
   267   (* like Existance Conjunction *)
   268   | "ec" =>
   269     let
   270       val ss = presburger_ss addsimps zadd_ac
   271       val ct = cert_Trueprop sg fm2
   272     in 
   273       simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (blast_tac HOL_cs 1)]
   274     end
   275 
   276   | "ac" =>
   277     let
   278       val ss = HOL_basic_ss addsimps zadd_ac
   279       val ct = cert_Trueprop sg fm2
   280     in 
   281       simple_prove_goal_cterm2 ct [simp_tac ss 1]
   282     end
   283 
   284   | "lf" =>
   285     let
   286       val ss = presburger_ss addsimps zadd_ac
   287       val ct = cert_Trueprop sg fm2
   288     in 
   289       simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]  
   290 
   291     end;
   292 
   293 (*=============================================================*)
   294 (*-------------------------------------------------------------*)
   295 (*              The new compact model                          *)
   296 (*-------------------------------------------------------------*)
   297 (*=============================================================*)
   298 
   299 fun thm_of sg decomp t = 
   300     let val (ts,recomb) = decomp t 
   301     in recomb (map (thm_of sg decomp) ts) 
   302     end;
   303 
   304 (*==================================================*)
   305 (*     Compact Version for adjustcoeffeq            *)
   306 (*==================================================*)
   307 
   308 fun decomp_adjustcoeffeq sg x l fm = case fm of
   309     (Const("Not",_)$(Const("op <",_) $(Const("0",_)) $(rt as (Const ("op +", _)$(Const ("op *",_) $    c $ y ) $z )))) => 
   310      let  
   311         val m = l div (dest_numeral c) 
   312         val n = if (x = y) then abs (m) else 1
   313         val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x)) 
   314         val rs = if (x = y) 
   315                  then (HOLogic.mk_binrel "op <" (zero,linear_sub [] (mk_numeral n) (HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
   316                  else HOLogic.mk_binrel "op <" (zero,linear_sub [] one rt )
   317         val ck = cterm_of sg (mk_numeral n)
   318         val cc = cterm_of sg c
   319         val ct = cterm_of sg z
   320         val cx = cterm_of sg y
   321         val pre = prove_elementar sg "lf" 
   322             (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral n)))
   323         val th1 = (pre RS (instantiate' [] [Some ck,Some cc, Some cx, Some ct] (ac_pi_eq)))
   324         in ([], fn [] => [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
   325         end
   326 
   327   |(Const(p,_) $a $( Const ("op +", _)$(Const ("op *",_) $ 
   328       c $ y ) $t )) => 
   329    if (is_arith_rel fm) andalso (x = y) 
   330    then  
   331         let val m = l div (dest_numeral c) 
   332            val k = (if p = "op <" then abs(m) else m)  
   333            val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div k)*l) ), x))
   334            val rs = (HOLogic.mk_binrel p ((linear_cmul k a),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul k t) )))) 
   335 
   336            val ck = cterm_of sg (mk_numeral k)
   337            val cc = cterm_of sg c
   338            val ct = cterm_of sg t
   339            val cx = cterm_of sg x
   340            val ca = cterm_of sg a
   341 
   342 	   in 
   343 	case p of
   344 	  "op <" => 
   345 	let val pre = prove_elementar sg "lf" 
   346 	    (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral k)))
   347             val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_lt_eq)))
   348 	in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
   349          end
   350 
   351            |"op =" =>
   352 	     let val pre = prove_elementar sg "lf" 
   353 	    (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
   354 	         val th1 = (pre RS(instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct] (ac_eq_eq)))
   355 	     in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
   356              end
   357 
   358              |"Divides.op dvd" =>
   359 	       let val pre = prove_elementar sg "lf" 
   360 	   (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
   361                    val th1 = (pre RS (instantiate' [] [Some ck,Some ca,Some cc, Some cx, Some ct]) (ac_dvd_eq))
   362                in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
   363                         
   364                end
   365               end
   366   else ([], fn [] => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl)
   367 
   368  |( Const ("Not", _) $ p) => ([p], fn [th] => th RS qe_Not)
   369   |( Const ("op &",_) $ p $ q) => ([p,q], fn [th1,th2] => [th1,th2] MRS qe_conjI)
   370   |( Const ("op |",_) $ p $ q) =>([p,q], fn [th1,th2] => [th1,th2] MRS qe_disjI)
   371 
   372   |_ => ([], fn [] => instantiate' [Some cboolT] [Some (cterm_of sg fm)] refl);
   373 
   374 fun proof_of_adjustcoeffeq sg x l = thm_of sg (decomp_adjustcoeffeq sg x l);
   375 
   376 
   377 
   378 (*==================================================*)
   379 (*   Finding rho for modd_minusinfinity             *)
   380 (*==================================================*)
   381 fun rho_for_modd_minf x dlcm sg fm1 =
   382 let
   383     (*Some certified Terms*)
   384     
   385    val ctrue = cterm_of sg HOLogic.true_const
   386    val cfalse = cterm_of sg HOLogic.false_const
   387    val fm = norm_zero_one fm1
   388   in  case fm1 of 
   389       (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
   390          if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
   391            else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
   392 
   393       |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
   394   	   if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one) 
   395 	   then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf))
   396 	 	 else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf)) 
   397 
   398       |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
   399            if (y=x) andalso (c1 = zero) then 
   400             if (pm1 = one) then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_minf)) else
   401 	     (instantiate' [Some cboolT] [Some ctrue] (fm_modd_minf))
   402 	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
   403   
   404       |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   405          if y=x then  let val cz = cterm_of sg (norm_zero_one z)
   406 			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
   407 	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf)))
   408 		      end
   409 		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
   410       |(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
   411       c $ y ) $ z))) => 
   412          if y=x then  let val cz = cterm_of sg (norm_zero_one z)
   413 			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
   414 	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf)))
   415 		      end
   416 		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf))
   417 		
   418     
   419    |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_minf)
   420    end;	 
   421 (*=========================================================================*)
   422 (*=========================================================================*)
   423 fun rho_for_eq_minf x dlcm  sg fm1 =  
   424    let
   425    val fm = norm_zero_one fm1
   426     in  case fm1 of 
   427       (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
   428          if  (x=y) andalso (c1=zero) andalso (c2=one) 
   429 	   then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_minf))
   430            else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
   431 
   432       |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
   433   	   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))
   434 	     then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_minf))
   435 	     else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf)) 
   436 
   437       |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
   438            if (y=x) andalso (c1 =zero) then 
   439             if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else
   440 	     (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_minf))
   441 	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
   442       |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   443          if y=x then  let val cd = cterm_of sg (norm_zero_one d)
   444 	 		  val cz = cterm_of sg (norm_zero_one z)
   445 	 	      in(instantiate' [] [Some cd,  Some cz] (not_dvd_eq_minf)) 
   446 		      end
   447 
   448 		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
   449 		
   450       |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   451          if y=x then  let val cd = cterm_of sg (norm_zero_one d)
   452 	 		  val cz = cterm_of sg (norm_zero_one z)
   453 	 	      in(instantiate' [] [Some cd, Some cz ] (dvd_eq_minf))
   454 		      end
   455 		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
   456 
   457       		
   458     |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_minf))
   459  end;
   460 
   461 (*=====================================================*)
   462 (*=====================================================*)
   463 (*=========== minf proofs with the compact version==========*)
   464 fun decomp_minf_eq x dlcm sg t =  case t of
   465    Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_conjI)
   466    |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_disjI)
   467    |_ => ([],fn [] => rho_for_eq_minf x dlcm sg t);
   468 
   469 fun decomp_minf_modd x dlcm sg t = case t of
   470    Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_conjI)
   471    |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_disjI)
   472    |_ => ([],fn [] => rho_for_modd_minf x dlcm sg t);
   473 
   474 (* -------------------------------------------------------------*)
   475 (*                    Finding rho for pinf_modd                 *)
   476 (* -------------------------------------------------------------*)
   477 fun rho_for_modd_pinf x dlcm sg fm1 = 
   478 let
   479     (*Some certified Terms*)
   480     
   481   val ctrue = cterm_of sg HOLogic.true_const
   482   val cfalse = cterm_of sg HOLogic.false_const
   483   val fm = norm_zero_one fm1
   484  in  case fm1 of 
   485       (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
   486          if ((x=y) andalso (c1= zero) andalso (c2= one))
   487 	 then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf))
   488          else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
   489 
   490       |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
   491   	if ((is_arith_rel fm) andalso (x = y) andalso (c1 = zero)  andalso (c2 = one)) 
   492 	then (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
   493 	else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
   494 
   495       |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
   496         if ((y=x) andalso (c1 = zero)) then 
   497           if (pm1 = one) 
   498 	  then (instantiate' [Some cboolT] [Some ctrue] (fm_modd_pinf)) 
   499 	  else (instantiate' [Some cboolT] [Some cfalse] (fm_modd_pinf))
   500 	else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
   501   
   502       |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   503          if y=x then  let val cz = cterm_of sg (norm_zero_one z)
   504 			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
   505 	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf)))
   506 		      end
   507 		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
   508       |(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
   509       c $ y ) $ z))) => 
   510          if y=x then  let val cz = cterm_of sg (norm_zero_one z)
   511 			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
   512 	 	      in(instantiate' [] [Some cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf)))
   513 		      end
   514 		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf))
   515 		
   516     
   517    |_ => instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_modd_pinf)
   518    end;	
   519 (* -------------------------------------------------------------*)
   520 (*                    Finding rho for pinf_eq                 *)
   521 (* -------------------------------------------------------------*)
   522 fun rho_for_eq_pinf x dlcm sg fm1 = 
   523   let
   524 					val fm = norm_zero_one fm1
   525     in  case fm1 of 
   526       (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
   527          if  (x=y) andalso (c1=zero) andalso (c2=one) 
   528 	   then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (neq_eq_pinf))
   529            else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
   530 
   531       |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
   532   	   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))
   533 	     then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (eq_eq_pinf))
   534 	     else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf)) 
   535 
   536       |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
   537            if (y=x) andalso (c1 =zero) then 
   538             if pm1 = one then (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else
   539 	     (instantiate' [] [Some (cterm_of sg (norm_zero_one z))] (len_eq_pinf))
   540 	    else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
   541       |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   542          if y=x then  let val cd = cterm_of sg (norm_zero_one d)
   543 	 		  val cz = cterm_of sg (norm_zero_one z)
   544 	 	      in(instantiate' [] [Some cd,  Some cz] (not_dvd_eq_pinf)) 
   545 		      end
   546 
   547 		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
   548 		
   549       |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   550          if y=x then  let val cd = cterm_of sg (norm_zero_one d)
   551 	 		  val cz = cterm_of sg (norm_zero_one z)
   552 	 	      in(instantiate' [] [Some cd, Some cz ] (dvd_eq_pinf))
   553 		      end
   554 		else (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
   555 
   556       		
   557     |_ => (instantiate' [Some cboolT] [Some (cterm_of sg fm)] (fm_eq_pinf))
   558  end;
   559 
   560 
   561 
   562 fun  minf_proof_of_c sg x dlcm t =
   563   let val minf_eqth   = thm_of sg (decomp_minf_eq x dlcm sg) t
   564       val minf_moddth = thm_of sg (decomp_minf_modd x dlcm sg) t
   565   in (minf_eqth, minf_moddth)
   566 end;
   567 
   568 (*=========== pinf proofs with the compact version==========*)
   569 fun decomp_pinf_eq x dlcm sg t = case t of
   570    Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_conjI)
   571    |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_disjI)
   572    |_ =>([],fn [] => rho_for_eq_pinf x dlcm sg t) ;
   573 
   574 fun decomp_pinf_modd x dlcm sg t =  case t of
   575    Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_conjI)
   576    |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_disjI)
   577    |_ => ([],fn [] => rho_for_modd_pinf x dlcm sg t);
   578 
   579 fun  pinf_proof_of_c sg x dlcm t =
   580   let val pinf_eqth   = thm_of sg (decomp_pinf_eq x dlcm sg) t
   581       val pinf_moddth = thm_of sg (decomp_pinf_modd x dlcm sg) t
   582   in (pinf_eqth,pinf_moddth)
   583 end;
   584 
   585 
   586 (* ------------------------------------------------------------------------- *)
   587 (* Here we generate the theorem for the Bset Property in the simple direction*)
   588 (* It is just an instantiation*)
   589 (* ------------------------------------------------------------------------- *)
   590 (*
   591 fun bsetproof_of sg (x as Free(xn,xT)) fm bs dlcm   = 
   592   let
   593     val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
   594     val cdlcm = cterm_of sg dlcm
   595     val cB = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one bs))
   596   in instantiate' [] [Some cdlcm,Some cB, Some cp] (bst_thm)
   597 end;
   598 
   599 fun asetproof_of sg (x as Free(xn,xT)) fm ast dlcm = 
   600   let
   601     val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
   602     val cdlcm = cterm_of sg dlcm
   603     val cA = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one ast))
   604   in instantiate' [] [Some cdlcm,Some cA, Some cp] (ast_thm)
   605 end;
   606 *)
   607 
   608 (* For the generation of atomic Theorems*)
   609 (* Prove the premisses on runtime and then make RS*)
   610 (* ------------------------------------------------------------------------- *)
   611 
   612 (*========= this is rho ============*)
   613 fun generate_atomic_not_bst_p sg (x as Free(xn,xT)) fm dlcm B at = 
   614   let
   615     val cdlcm = cterm_of sg dlcm
   616     val cB = cterm_of sg B
   617     val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
   618     val cat = cterm_of sg (norm_zero_one at)
   619   in
   620   case at of 
   621    (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
   622       if  (x=y) andalso (c1=zero) andalso (c2=one) 
   623 	 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)
   624 	          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)))
   625 		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
   626 	 in  (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_bst_p_ne)))
   627 	 end
   628          else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
   629 
   630    |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
   631      if (is_arith_rel at) andalso (x=y)
   632 	then let val bst_z = norm_zero_one (linear_neg (linear_add [] z (mk_numeral 1)))
   633 	         in let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ bst_z $ B)
   634 	          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))))
   635 		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
   636 	 in  (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_bst_p_eq)))
   637 	 end
   638        end
   639          else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
   640 
   641    |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
   642         if (y=x) andalso (c1 =zero) then 
   643         if pm1 = one then 
   644 	  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)
   645               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)))
   646 	  in  (instantiate' [] [Some cfma,  Some cdlcm]([th1,th2] MRS (not_bst_p_gt)))
   647 	    end
   648 	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
   649 	      in (instantiate' [] [Some cfma, Some cB,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt)))
   650 	      end
   651       else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
   652 
   653    |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   654       if y=x then  
   655            let val cz = cterm_of sg (norm_zero_one z)
   656 	       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)
   657  	     in (instantiate' []  [Some cfma, Some cB,Some cz] (th1 RS (not_bst_p_ndvd)))
   658 	     end
   659       else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
   660 
   661    |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   662        if y=x then  
   663 	 let val cz = cterm_of sg (norm_zero_one z)
   664 	     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)
   665  	    in (instantiate' []  [Some cfma,Some cB,Some cz] (th1 RS (not_bst_p_dvd)))
   666 	  end
   667       else (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
   668       		
   669    |_ => (instantiate' [] [Some cfma,  Some cdlcm, Some cB,Some cat] (not_bst_p_fm))
   670       		
   671     end;
   672     
   673 
   674 (* ------------------------------------------------------------------------- *)    
   675 (* Main interpretation function for this backwards dirction*)
   676 (* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*)
   677 (*Help Function*)
   678 (* ------------------------------------------------------------------------- *)
   679 
   680 (*==================== Proof with the compact version   *)
   681 
   682 fun decomp_nbstp sg x dlcm B fm t = case t of 
   683    Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_conjI )
   684   |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_disjI)
   685   |_ => ([], fn [] => generate_atomic_not_bst_p sg x fm dlcm B t);
   686 
   687 fun not_bst_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm B t =
   688   let 
   689        val th =  thm_of sg (decomp_nbstp sg x dlcm (list_to_set xT (map norm_zero_one B)) fm) t
   690       val fma = absfree (xn,xT, norm_zero_one fm)
   691   in let val th1 =  prove_elementar sg "ss"  (HOLogic.mk_eq (fma,fma))
   692      in [th,th1] MRS (not_bst_p_Q_elim)
   693      end
   694   end;
   695 
   696 
   697 (* ------------------------------------------------------------------------- *)    
   698 (* Protokol interpretation function for the backwards direction for cooper's Theorem*)
   699 
   700 (* For the generation of atomic Theorems*)
   701 (* Prove the premisses on runtime and then make RS*)
   702 (* ------------------------------------------------------------------------- *)
   703 (*========= this is rho ============*)
   704 fun generate_atomic_not_ast_p sg (x as Free(xn,xT)) fm dlcm A at = 
   705   let
   706     val cdlcm = cterm_of sg dlcm
   707     val cA = cterm_of sg A
   708     val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
   709     val cat = cterm_of sg (norm_zero_one at)
   710   in
   711   case at of 
   712    (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
   713       if  (x=y) andalso (c1=zero) andalso (c2=one) 
   714 	 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)
   715 	          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)))
   716 		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
   717 	 in  (instantiate' [] [Some cfma]([th3,th1,th2] MRS (not_ast_p_ne)))
   718 	 end
   719          else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
   720 
   721    |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
   722      if (is_arith_rel at) andalso (x=y)
   723 	then let val ast_z = norm_zero_one (linear_sub [] one z )
   724 	         val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ ast_z $ A)
   725 	         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))))
   726 		 val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
   727 	 in  (instantiate' [] [Some cfma] ([th3,th1,th2] MRS (not_ast_p_eq)))
   728        end
   729          else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
   730 
   731    |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
   732         if (y=x) andalso (c1 =zero) then 
   733         if pm1 = (mk_numeral ~1) then 
   734 	  let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one z) $ A)
   735               val th2 =  prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm))
   736 	  in  (instantiate' [] [Some cfma]([th2,th1] MRS (not_ast_p_lt)))
   737 	    end
   738 	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
   739 	      in (instantiate' [] [Some cfma, Some cA,Some (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt)))
   740 	      end
   741       else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
   742 
   743    |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   744       if y=x then  
   745            let val cz = cterm_of sg (norm_zero_one z)
   746 	       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)
   747  	     in (instantiate' []  [Some cfma, Some cA,Some cz] (th1 RS (not_ast_p_ndvd)))
   748 	     end
   749       else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
   750 
   751    |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
   752        if y=x then  
   753 	 let val cz = cterm_of sg (norm_zero_one z)
   754 	     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)
   755  	    in (instantiate' []  [Some cfma,Some cA,Some cz] (th1 RS (not_ast_p_dvd)))
   756 	  end
   757       else (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
   758       		
   759    |_ => (instantiate' [] [Some cfma,  Some cdlcm, Some cA,Some cat] (not_ast_p_fm))
   760       		
   761     end;
   762 
   763 (* ------------------------------------------------------------------------ *)
   764 (* Main interpretation function for this backwards dirction*)
   765 (* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*)
   766 (*Help Function*)
   767 (* ------------------------------------------------------------------------- *)
   768 
   769 fun decomp_nastp sg x dlcm A fm t = case t of 
   770    Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_conjI )
   771   |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_disjI)
   772   |_ => ([], fn [] => generate_atomic_not_ast_p sg x fm dlcm A t);
   773 
   774 fun not_ast_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm A t =
   775   let 
   776        val th =  thm_of sg (decomp_nastp sg x dlcm (list_to_set xT (map norm_zero_one A)) fm) t
   777       val fma = absfree (xn,xT, norm_zero_one fm)
   778   in let val th1 =  prove_elementar sg "ss"  (HOLogic.mk_eq (fma,fma))
   779      in [th,th1] MRS (not_ast_p_Q_elim)
   780      end
   781   end;
   782 
   783 
   784 (* -------------------------------*)
   785 (* Finding rho and beta for evalc *)
   786 (* -------------------------------*)
   787 
   788 fun rho_for_evalc sg at = case at of  
   789     (Const (p,_) $ s $ t) =>(  
   790     case assoc (operations,p) of 
   791         Some f => 
   792            ((if (f ((dest_numeral s),(dest_numeral t))) 
   793              then prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)) 
   794              else prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const)))  
   795 		   handle _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl)
   796         | _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl )
   797      |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
   798        case assoc (operations,p) of 
   799          Some f => 
   800            ((if (f ((dest_numeral s),(dest_numeral t))) 
   801              then prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const))  
   802              else prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)))  
   803 		      handle _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl) 
   804          | _ => instantiate' [Some cboolT] [Some (cterm_of sg at)] refl ) 
   805      | _ =>   instantiate' [Some cboolT] [Some (cterm_of sg at)] refl;
   806 
   807 
   808 (*=========================================================*)
   809 fun decomp_evalc sg t = case t of
   810    (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI)
   811    |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI)
   812    |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI)
   813    |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
   814    |_ => ([], fn [] => rho_for_evalc sg t);
   815 
   816 
   817 fun proof_of_evalc sg fm = thm_of sg (decomp_evalc sg) fm;
   818 
   819 (*==================================================*)
   820 (*     Proof of linform with the compact model      *)
   821 (*==================================================*)
   822 
   823 
   824 fun decomp_linform sg vars t = case t of
   825    (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI)
   826    |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI)
   827    |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI)
   828    |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
   829    |(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not)
   830    |(Const("Divides.op dvd",_)$d$r) => ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [None , None, Some (cterm_of sg d)](linearize_dvd)))
   831    |_ => ([], fn [] => prove_elementar sg "lf" (HOLogic.mk_eq (t, linform vars t)));
   832 
   833 fun proof_of_linform sg vars f = thm_of sg (decomp_linform sg vars) f;
   834 
   835 (* ------------------------------------------------------------------------- *)
   836 (* Interpretaion of Protocols of the cooper procedure : minusinfinity version*)
   837 (* ------------------------------------------------------------------------- *)
   838 fun coopermi_proof_of sg (x as Free(xn,xT)) fm B dlcm =
   839   (* Get the Bset thm*)
   840   let val (minf_eqth, minf_moddth) = minf_proof_of_c sg x dlcm fm 
   841       val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm));
   842       val nbstpthm = not_bst_p_proof_of_c sg x fm dlcm B fm
   843   in (dpos,minf_eqth,nbstpthm,minf_moddth)
   844 end;
   845 
   846 (* ------------------------------------------------------------------------- *)
   847 (* Interpretaion of Protocols of the cooper procedure : plusinfinity version *)
   848 (* ------------------------------------------------------------------------- *)
   849 fun cooperpi_proof_of sg (x as Free(xn,xT)) fm A dlcm =
   850   let val (pinf_eqth,pinf_moddth) = pinf_proof_of_c sg x dlcm fm
   851       val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm));
   852       val nastpthm = not_ast_p_proof_of_c sg x fm dlcm A fm
   853   in (dpos,pinf_eqth,nastpthm,pinf_moddth)
   854 end;
   855 
   856 (* ------------------------------------------------------------------------- *)
   857 (* Interpretaion of Protocols of the cooper procedure : full version*)
   858 (* ------------------------------------------------------------------------- *)
   859 fun cooper_thm sg s (x as Free(xn,xT)) cfm dlcm ast bst= case s of
   860   "pi" => let val (dpsthm,pinf_eqth,nbpth,pinf_moddth) = cooperpi_proof_of sg x cfm ast dlcm 
   861 	      in [dpsthm,pinf_eqth,nbpth,pinf_moddth] MRS (cppi_eq)
   862            end
   863   |"mi" => let val (dpsthm,minf_eqth,nbpth,minf_moddth) = coopermi_proof_of sg x cfm bst dlcm
   864 	       in [dpsthm,minf_eqth,nbpth,minf_moddth] MRS (cpmi_eq)
   865                 end
   866  |_ => error "parameter error";
   867 
   868 (* ------------------------------------------------------------------------- *)
   869 (* This function should evoluate to the end prove Procedure for one quantifier elimination for Presburger arithmetic*)
   870 (* It shoud be plugged in the qfnp argument of the quantifier elimination proof function*)
   871 (* ------------------------------------------------------------------------- *)
   872 
   873 fun cooper_prv sg (x as Free(xn,xT)) efm = let 
   874    (* lfm_thm : efm = linearized form of efm*)
   875    val lfm_thm = proof_of_linform sg [xn] efm
   876    (*efm2 is the linearized form of efm *) 
   877    val efm2 = snd(qe_get_terms lfm_thm)
   878    (* l is the lcm of all coefficients of x *)
   879    val l = formlcm x efm2
   880    (*ac_thm: efm = efm2 with adjusted coefficients of x *)
   881    val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans
   882    (* fm is efm2 with adjusted coefficients of x *)
   883    val fm = snd (qe_get_terms ac_thm)
   884   (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*)
   885    val  cfm = unitycoeff x fm
   886    (*afm is fm where c*x is replaced by 1*x or -1*x *)
   887    val afm = adjustcoeff x l fm
   888    (* P = %x.afm*)
   889    val P = absfree(xn,xT,afm)
   890    (* This simpset allows the elimination of the sets in bex {1..d} *)
   891    val ss = presburger_ss addsimps
   892      [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
   893    (* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
   894    val uth = instantiate' [] [Some (cterm_of sg P) , Some (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
   895    (* e_ac_thm : Ex x. efm = EX x. fm*)
   896    val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
   897    (* A and B set of the formula*)
   898    val A = aset x cfm
   899    val B = bset x cfm
   900    (* the divlcm (delta) of the formula*)
   901    val dlcm = mk_numeral (divlcm x cfm)
   902    (* Which set is smaller to generate the (hoepfully) shorter proof*)
   903    val cms = if ((length A) < (length B )) then "pi" else "mi"
   904    (* synthesize the proof of cooper's theorem*)
   905     (* cp_thm: EX x. cfm = Q*)
   906    val cp_thm = cooper_thm sg cms x cfm dlcm A B
   907    (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*)
   908    (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*)
   909    val exp_cp_thm = refl RS (simplify ss (cp_thm RSN (2,trans)))
   910    (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*)
   911    val (lsuth,rsuth) = qe_get_terms (uth)
   912    (* lseacth = EX x. efm; rseacth = EX x. fm*)
   913    val (lseacth,rseacth) = qe_get_terms(e_ac_thm)
   914    (* lscth = EX x. cfm; rscth = Q' *)
   915    val (lscth,rscth) = qe_get_terms (exp_cp_thm)
   916    (* u_c_thm: EX x. P(l*x) = Q'*)
   917    val  u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans
   918    (* result: EX x. efm = Q'*)
   919  in  ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans)
   920    end
   921 |cooper_prv _ _ _ =  error "Parameters format";
   922 
   923 (* **************************************** *)
   924 (*    An Other Version of cooper proving    *)
   925 (*     by giving a withness for EX          *)
   926 (* **************************************** *)
   927 
   928 
   929 
   930 fun cooper_prv_w sg (x as Free(xn,xT)) efm = let 
   931    (* lfm_thm : efm = linearized form of efm*)
   932    val lfm_thm = proof_of_linform sg [xn] efm
   933    (*efm2 is the linearized form of efm *) 
   934    val efm2 = snd(qe_get_terms lfm_thm)
   935    (* l is the lcm of all coefficients of x *)
   936    val l = formlcm x efm2
   937    (*ac_thm: efm = efm2 with adjusted coefficients of x *)
   938    val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans
   939    (* fm is efm2 with adjusted coefficients of x *)
   940    val fm = snd (qe_get_terms ac_thm)
   941   (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*)
   942    val  cfm = unitycoeff x fm
   943    (*afm is fm where c*x is replaced by 1*x or -1*x *)
   944    val afm = adjustcoeff x l fm
   945    (* P = %x.afm*)
   946    val P = absfree(xn,xT,afm)
   947    (* This simpset allows the elimination of the sets in bex {1..d} *)
   948    val ss = presburger_ss addsimps
   949      [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
   950    (* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
   951    val uth = instantiate' [] [Some (cterm_of sg P) , Some (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
   952    (* e_ac_thm : Ex x. efm = EX x. fm*)
   953    val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
   954    (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*)
   955    val (lsuth,rsuth) = qe_get_terms (uth)
   956    (* lseacth = EX x. efm; rseacth = EX x. fm*)
   957    val (lseacth,rseacth) = qe_get_terms(e_ac_thm)
   958 
   959    val (w,rs) = cooper_w [] cfm
   960    val exp_cp_thm =  case w of 
   961      (* FIXME - e_ac_thm just tipped to test syntactical correctness of the program!!!!*)
   962     Some n =>  e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*)
   963    |_ => let 
   964     (* A and B set of the formula*)
   965     val A = aset x cfm
   966     val B = bset x cfm
   967     (* the divlcm (delta) of the formula*)
   968     val dlcm = mk_numeral (divlcm x cfm)
   969     (* Which set is smaller to generate the (hoepfully) shorter proof*)
   970     val cms = if ((length A) < (length B )) then "pi" else "mi"
   971     (* synthesize the proof of cooper's theorem*)
   972      (* cp_thm: EX x. cfm = Q*)
   973     val cp_thm = cooper_thm sg cms x cfm dlcm A B
   974      (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*)
   975     (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*)
   976     in refl RS (simplify ss (cp_thm RSN (2,trans)))
   977     end
   978    (* lscth = EX x. cfm; rscth = Q' *)
   979    val (lscth,rscth) = qe_get_terms (exp_cp_thm)
   980    (* u_c_thm: EX x. P(l*x) = Q'*)
   981    val  u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans
   982    (* result: EX x. efm = Q'*)
   983  in  ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans)
   984    end
   985 |cooper_prv_w _ _ _ =  error "Parameters format";
   986 
   987 
   988 
   989 fun decomp_cnnf sg lfnp P = case P of 
   990      Const ("op &",_) $ p $q => ([p,q] , fn [th1,th2] => [th1,th2] MRS qe_conjI )
   991    |Const ("op |",_) $ p $q => ([p,q] , fn [th1,th2] => [th1,th2] MRS  qe_disjI)
   992    |Const ("Not",_) $ (Const("Not",_) $ p) => ([p], fn [th] => th RS nnf_nn)
   993    |Const("Not",_) $ (Const(opn,T) $ p $ q) => 
   994      if opn = "op |" 
   995       then case (p,q) of 
   996          (A as (Const ("op &",_) $ r $ s),B as (Const ("op &",_) $ r1 $ t)) =>
   997           if r1 = negate r 
   998           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)
   999 
  1000           else ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj)
  1001         |(_,_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj)
  1002       else (
  1003          case (opn,T) of 
  1004            ("op &",_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_ncj )
  1005            |("op -->",_) => ([p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_nim )
  1006            |("op =",Type ("fun",[Type ("bool", []),_])) => 
  1007            ([HOLogic.conj $ p $ (HOLogic.Not $ q),HOLogic.conj $ (HOLogic.Not $ p) $ q], fn [th1,th2] => [th1,th2] MRS nnf_neq)
  1008             |(_,_) => ([], fn [] => lfnp P)
  1009 )
  1010 
  1011    |(Const ("op -->",_) $ p $ q) => ([HOLogic.Not$p,q], fn [th1,th2] => [th1,th2] MRS nnf_im)
  1012 
  1013    |(Const ("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q) =>
  1014      ([HOLogic.conj $ p $ q,HOLogic.conj $ (HOLogic.Not $ p) $ (HOLogic.Not $ q) ], fn [th1,th2] =>[th1,th2] MRS nnf_eq )
  1015    |_ => ([], fn [] => lfnp P);
  1016 
  1017 
  1018 
  1019 
  1020 fun proof_of_cnnf sg p lfnp = 
  1021  let val th1 = thm_of sg (decomp_cnnf sg lfnp) p
  1022      val rs = snd(qe_get_terms th1)
  1023      val th2 = prove_elementar sg "ss" (HOLogic.mk_eq(rs,simpl rs))
  1024   in [th1,th2] MRS trans
  1025   end;
  1026 
  1027 end;
  1028