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