| author | mengj | 
| Fri, 28 Apr 2006 05:59:32 +0200 | |
| changeset 19491 | cd6c71c57f53 | 
| parent 19277 | f7602e74d948 | 
| child 20052 | 3d4ff822d0b3 | 
| permissions | -rw-r--r-- | 
| 13876 | 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 | ||
| 14920 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 chaieb parents: 
14877diff
changeset | 9 | |
| 13876 | 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 | |
| 14877 | 18 | val list_to_set : typ -> term list -> term | 
| 13876 | 19 | val qe_get_terms : thm -> term * term | 
| 19250 | 20 | val cooper_prv : theory -> term -> term -> thm | 
| 21 | val proof_of_evalc : theory -> term -> thm | |
| 22 | val proof_of_cnnf : theory -> term -> (term -> thm) -> thm | |
| 23 | val proof_of_linform : theory -> string list -> term -> thm | |
| 24 | val proof_of_adjustcoeffeq : theory -> term -> IntInf.int -> term -> thm | |
| 25 | val prove_elementar : theory -> string -> term -> thm | |
| 26 | val thm_of : theory -> (term -> (term list * (thm list -> thm))) -> term -> thm | |
| 13876 | 27 | end; | 
| 28 | ||
| 29 | structure CooperProof : COOPER_PROOF = | |
| 30 | struct | |
| 31 | open CooperDec; | |
| 32 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 33 | (* | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 34 | val presburger_ss = simpset_of (theory "Presburger") | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 35 | addsimps [zdiff_def] delsimps [symmetric zdiff_def]; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 36 | *) | 
| 13876 | 37 | |
| 38 | val presburger_ss = simpset_of (theory "Presburger") | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 39 | addsimps[diff_int_def] delsimps [thm"diff_int_def_symmetric"]; | 
| 13876 | 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 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 47 | (* Thorems for proving the adjustment of the coeffitients*) | 
| 13876 | 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 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 63 | (*Modulo D property for Pminusinf an Plusinf *) | 
| 13876 | 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 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 72 | (* the minusinfinity proprty*) | 
| 13876 | 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 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 82 | (* the Plusinfinity proprty*) | 
| 13876 | 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 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 168 |   (Const ("HOL.times",_) $ c $ t) => 
 | 
| 13876 | 169 | if c = one then (norm_zero_one t) | 
| 170 | else if (dest_numeral c = ~1) | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 171 |          then (Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $ (norm_zero_one t))
 | 
| 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 172 | else (HOLogic.mk_binop "HOL.times" (norm_zero_one c,norm_zero_one t)) | 
| 13876 | 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 | (* ------------------------------------------------------------------------- *) | |
| 17985 | 192 | (*This function proove elementar will be used to generate proofs at | 
| 193 | runtime*) (*It is thought to prove properties such as a dvd b | |
| 194 | (essentially) that are only to make at runtime.*) | |
| 13876 | 195 | (* ------------------------------------------------------------------------- *) | 
| 17985 | 196 | fun prove_elementar thy s fm2 = | 
| 197 | Goal.prove thy [] [] (HOLogic.mk_Trueprop fm2) (fn _ => EVERY | |
| 198 | (case s of | |
| 13876 | 199 | (*"ss" like simplification with simpset*) | 
| 200 | "ss" => | |
| 17985 | 201 | let val ss = presburger_ss addsimps [zdvd_iff_zmod_eq_0,unity_coeff_ex] | 
| 202 | in [simp_tac ss 1, TRY (simple_arith_tac 1)] end | |
| 13876 | 203 | |
| 204 | (*"bl" like blast tactic*) | |
| 205 | (* Is only used in the harrisons like proof procedure *) | |
| 17985 | 206 | | "bl" => [blast_tac HOL_cs 1] | 
| 13876 | 207 | |
| 208 | (*"ed" like Existence disjunctions ...*) | |
| 209 | (* Is only used in the harrisons like proof procedure *) | |
| 210 | | "ed" => | |
| 211 | let | |
| 212 | val ex_disj_tacs = | |
| 213 | let | |
| 214 | val tac1 = EVERY[REPEAT(resolve_tac [disjI1,disjI2] 1), etac exI 1] | |
| 215 | val tac2 = EVERY[etac exE 1, rtac exI 1, | |
| 216 | REPEAT(resolve_tac [disjI1,disjI2] 1), assumption 1] | |
| 217 | in [rtac iffI 1, | |
| 218 | etac exE 1, REPEAT(EVERY[etac disjE 1, tac1]), tac1, | |
| 219 | REPEAT(EVERY[etac disjE 1, tac2]), tac2] | |
| 220 | end | |
| 17985 | 221 | in ex_disj_tacs end | 
| 13876 | 222 | |
| 17985 | 223 | | "fa" => [simple_arith_tac 1] | 
| 13876 | 224 | |
| 225 | | "sa" => | |
| 17985 | 226 | let val ss = presburger_ss addsimps zadd_ac | 
| 227 | in [simp_tac ss 1, TRY (simple_arith_tac 1)] end | |
| 15107 | 228 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 229 | (* like Existance Conjunction *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 230 | | "ec" => | 
| 17985 | 231 | let val ss = presburger_ss addsimps zadd_ac | 
| 232 | in [simp_tac ss 1, TRY (blast_tac HOL_cs 1)] end | |
| 13876 | 233 | |
| 234 | | "ac" => | |
| 17985 | 235 | let val ss = HOL_basic_ss addsimps zadd_ac | 
| 236 | in [simp_tac ss 1] end | |
| 13876 | 237 | |
| 238 | | "lf" => | |
| 17985 | 239 | let val ss = presburger_ss addsimps zadd_ac | 
| 240 | in [simp_tac ss 1, TRY (simple_arith_tac 1)] end)); | |
| 13876 | 241 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 242 | (*=============================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 243 | (*-------------------------------------------------------------*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 244 | (* The new compact model *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 245 | (*-------------------------------------------------------------*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 246 | (*=============================================================*) | 
| 13876 | 247 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 248 | fun thm_of sg decomp t = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 249 | let val (ts,recomb) = decomp t | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 250 | in recomb (map (thm_of sg decomp) ts) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 251 | end; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 252 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 253 | (*==================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 254 | (* Compact Version for adjustcoeffeq *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 255 | (*==================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 256 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 257 | fun decomp_adjustcoeffeq sg x l fm = case fm of | 
| 19277 | 258 |     (Const("Not",_)$(Const("Orderings.less",_) $(Const("0",_)) $(rt as (Const ("HOL.plus", _)$(Const ("HOL.times",_) $    c $ y ) $z )))) => 
 | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 259 | let | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 260 | val m = l div (dest_numeral c) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 261 | val n = if (x = y) then abs (m) else 1 | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 262 | val xtm = (HOLogic.mk_binop "HOL.times" ((mk_numeral ((m div n)*l) ), x)) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 263 | val rs = if (x = y) | 
| 19277 | 264 | then (HOLogic.mk_binrel "Orderings.less" (zero,linear_sub [] (mk_numeral n) (HOLogic.mk_binop "HOL.plus" ( xtm ,( linear_cmul n z) )))) | 
| 265 | else HOLogic.mk_binrel "Orderings.less" (zero,linear_sub [] one rt ) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 266 | val ck = cterm_of sg (mk_numeral n) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 267 | val cc = cterm_of sg c | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 268 | val ct = cterm_of sg z | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 269 | val cx = cterm_of sg y | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 270 | val pre = prove_elementar sg "lf" | 
| 19277 | 271 |             (HOLogic.mk_binrel "Orderings.less" (Const("0",HOLogic.intT),(mk_numeral n)))
 | 
| 15531 | 272 | val th1 = (pre RS (instantiate' [] [SOME ck,SOME cc, SOME cx, SOME ct] (ac_pi_eq))) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 273 | in ([], fn [] => [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 274 | end | 
| 13876 | 275 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 276 |   |(Const(p,_) $a $( Const ("HOL.plus", _)$(Const ("HOL.times",_) $ 
 | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 277 | c $ y ) $t )) => | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 278 | if (is_arith_rel fm) andalso (x = y) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 279 | then | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 280 | let val m = l div (dest_numeral c) | 
| 19277 | 281 | val k = (if p = "Orderings.less" then abs(m) else m) | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 282 | val xtm = (HOLogic.mk_binop "HOL.times" ((mk_numeral ((m div k)*l) ), x)) | 
| 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 283 | val rs = (HOLogic.mk_binrel p ((linear_cmul k a),(HOLogic.mk_binop "HOL.plus" ( xtm ,( linear_cmul k t) )))) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 284 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 285 | val ck = cterm_of sg (mk_numeral k) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 286 | val cc = cterm_of sg c | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 287 | val ct = cterm_of sg t | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 288 | val cx = cterm_of sg x | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 289 | val ca = cterm_of sg a | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 290 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 291 | in | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 292 | case p of | 
| 19277 | 293 | "Orderings.less" => | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 294 | let val pre = prove_elementar sg "lf" | 
| 19277 | 295 | 	    (HOLogic.mk_binrel "Orderings.less" (Const("0",HOLogic.intT),(mk_numeral k)))
 | 
| 15531 | 296 | val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_lt_eq))) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 297 | in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 298 | end | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 299 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 300 | |"op =" => | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 301 | let val pre = prove_elementar sg "lf" | 
| 13876 | 302 | 	    (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
 | 
| 15531 | 303 | val th1 = (pre RS(instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_eq_eq))) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 304 | in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 305 | end | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 306 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 307 | |"Divides.op dvd" => | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 308 | let val pre = prove_elementar sg "lf" | 
| 13876 | 309 | 	   (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
 | 
| 15531 | 310 | val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct]) (ac_dvd_eq)) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 311 | in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans) | 
| 13876 | 312 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 313 | end | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 314 | end | 
| 15531 | 315 | else ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 316 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 317 |  |( Const ("Not", _) $ p) => ([p], fn [th] => th RS qe_Not)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 318 |   |( Const ("op &",_) $ p $ q) => ([p,q], fn [th1,th2] => [th1,th2] MRS qe_conjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 319 |   |( Const ("op |",_) $ p $ q) =>([p,q], fn [th1,th2] => [th1,th2] MRS qe_disjI)
 | 
| 13876 | 320 | |
| 15531 | 321 | |_ => ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl); | 
| 13876 | 322 | |
| 14877 | 323 | fun proof_of_adjustcoeffeq sg x l = thm_of sg (decomp_adjustcoeffeq sg x l); | 
| 324 | ||
| 325 | ||
| 326 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 327 | (*==================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 328 | (* Finding rho for modd_minusinfinity *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 329 | (*==================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 330 | fun rho_for_modd_minf x dlcm sg fm1 = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 331 | let | 
| 15661 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 wenzelm parents: 
15531diff
changeset | 332 | (*Some certified Terms*) | 
| 13876 | 333 | |
| 334 | val ctrue = cterm_of sg HOLogic.true_const | |
| 335 | val cfalse = cterm_of sg HOLogic.false_const | |
| 336 | val fm = norm_zero_one fm1 | |
| 337 | in case fm1 of | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 338 |       (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z))) => 
 | 
| 15531 | 339 | if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf)) | 
| 340 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) | |
| 13876 | 341 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 342 |       |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z)) =>
 | 
| 13876 | 343 | if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one) | 
| 15531 | 344 | then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf)) | 
| 345 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) | |
| 13876 | 346 | |
| 19277 | 347 |       |(Const("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ pm1 $ y ) $ z )) =>
 | 
| 13876 | 348 | if (y=x) andalso (c1 = zero) then | 
| 15531 | 349 | if (pm1 = one) then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf)) else | 
| 350 | (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf)) | |
| 351 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) | |
| 13876 | 352 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 353 |       |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 354 | if y=x then let val cz = cterm_of sg (norm_zero_one z) | 
| 355 | val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero) | |
| 15531 | 356 | in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf))) | 
| 13876 | 357 | end | 
| 15531 | 358 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 359 |       |(Const("Divides.op dvd",_)$ d $ (db as (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $
 | 
| 13876 | 360 | c $ y ) $ z))) => | 
| 361 | if y=x then let val cz = cterm_of sg (norm_zero_one z) | |
| 362 | val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero) | |
| 15531 | 363 | in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf))) | 
| 13876 | 364 | end | 
| 15531 | 365 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) | 
| 13876 | 366 | |
| 367 | ||
| 15531 | 368 | |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 369 | end; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 370 | (*=========================================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 371 | (*=========================================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 372 | fun rho_for_eq_minf x dlcm sg fm1 = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 373 | let | 
| 13876 | 374 | val fm = norm_zero_one fm1 | 
| 375 | in case fm1 of | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 376 |       (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z))) => 
 | 
| 13876 | 377 | if (x=y) andalso (c1=zero) andalso (c2=one) | 
| 15531 | 378 | then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_minf)) | 
| 379 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) | |
| 13876 | 380 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 381 |       |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z)) =>
 | 
| 13876 | 382 | 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)) | 
| 15531 | 383 | then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_minf)) | 
| 384 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) | |
| 13876 | 385 | |
| 19277 | 386 |       |(Const("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ pm1 $ y ) $ z )) =>
 | 
| 13876 | 387 | if (y=x) andalso (c1 =zero) then | 
| 15531 | 388 | if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else | 
| 389 | (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_minf)) | |
| 390 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 391 |       |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 392 | if y=x then let val cd = cterm_of sg (norm_zero_one d) | 
| 393 | val cz = cterm_of sg (norm_zero_one z) | |
| 15531 | 394 | in(instantiate' [] [SOME cd, SOME cz] (not_dvd_eq_minf)) | 
| 13876 | 395 | end | 
| 396 | ||
| 15531 | 397 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) | 
| 13876 | 398 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 399 |       |(Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 400 | if y=x then let val cd = cterm_of sg (norm_zero_one d) | 
| 401 | val cz = cterm_of sg (norm_zero_one z) | |
| 15531 | 402 | in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_minf)) | 
| 13876 | 403 | end | 
| 15531 | 404 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) | 
| 13876 | 405 | |
| 406 | ||
| 15531 | 407 | |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) | 
| 13876 | 408 | end; | 
| 409 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 410 | (*=====================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 411 | (*=====================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 412 | (*=========== minf proofs with the compact version==========*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 413 | fun decomp_minf_eq x dlcm sg t = case t of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 414 |    Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_conjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 415 |    |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 416 | |_ => ([],fn [] => rho_for_eq_minf x dlcm sg t); | 
| 13876 | 417 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 418 | fun decomp_minf_modd x dlcm sg t = case t of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 419 |    Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_conjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 420 |    |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 421 | |_ => ([],fn [] => rho_for_modd_minf x dlcm sg t); | 
| 13876 | 422 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 423 | (* -------------------------------------------------------------*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 424 | (* Finding rho for pinf_modd *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 425 | (* -------------------------------------------------------------*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 426 | fun rho_for_modd_pinf x dlcm sg fm1 = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 427 | let | 
| 15661 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 wenzelm parents: 
15531diff
changeset | 428 | (*Some certified Terms*) | 
| 13876 | 429 | |
| 430 | val ctrue = cterm_of sg HOLogic.true_const | |
| 431 | val cfalse = cterm_of sg HOLogic.false_const | |
| 432 | val fm = norm_zero_one fm1 | |
| 433 | in case fm1 of | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 434 |       (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z))) => 
 | 
| 13876 | 435 | if ((x=y) andalso (c1= zero) andalso (c2= one)) | 
| 15531 | 436 | then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf)) | 
| 437 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) | |
| 13876 | 438 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 439 |       |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z)) =>
 | 
| 13876 | 440 | if ((is_arith_rel fm) andalso (x = y) andalso (c1 = zero) andalso (c2 = one)) | 
| 15531 | 441 | then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf)) | 
| 442 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) | |
| 13876 | 443 | |
| 19277 | 444 |       |(Const("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ pm1 $ y ) $ z )) =>
 | 
| 13876 | 445 | if ((y=x) andalso (c1 = zero)) then | 
| 446 | if (pm1 = one) | |
| 15531 | 447 | then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf)) | 
| 448 | else (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf)) | |
| 449 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) | |
| 13876 | 450 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 451 |       |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 452 | if y=x then let val cz = cterm_of sg (norm_zero_one z) | 
| 453 | val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero) | |
| 15531 | 454 | in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf))) | 
| 13876 | 455 | end | 
| 15531 | 456 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 457 |       |(Const("Divides.op dvd",_)$ d $ (db as (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $
 | 
| 13876 | 458 | c $ y ) $ z))) => | 
| 459 | if y=x then let val cz = cterm_of sg (norm_zero_one z) | |
| 460 | val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero) | |
| 15531 | 461 | in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf))) | 
| 13876 | 462 | end | 
| 15531 | 463 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)) | 
| 13876 | 464 | |
| 465 | ||
| 15531 | 466 | |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 467 | end; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 468 | (* -------------------------------------------------------------*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 469 | (* Finding rho for pinf_eq *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 470 | (* -------------------------------------------------------------*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 471 | fun rho_for_eq_pinf x dlcm sg fm1 = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 472 | let | 
| 13876 | 473 | val fm = norm_zero_one fm1 | 
| 474 | in case fm1 of | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 475 |       (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z))) => 
 | 
| 13876 | 476 | if (x=y) andalso (c1=zero) andalso (c2=one) | 
| 15531 | 477 | then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_pinf)) | 
| 478 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) | |
| 13876 | 479 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 480 |       |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z)) =>
 | 
| 13876 | 481 | 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)) | 
| 15531 | 482 | then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_pinf)) | 
| 483 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) | |
| 13876 | 484 | |
| 19277 | 485 |       |(Const("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ pm1 $ y ) $ z )) =>
 | 
| 13876 | 486 | if (y=x) andalso (c1 =zero) then | 
| 15531 | 487 | if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else | 
| 488 | (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_pinf)) | |
| 489 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 490 |       |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 491 | if y=x then let val cd = cterm_of sg (norm_zero_one d) | 
| 492 | val cz = cterm_of sg (norm_zero_one z) | |
| 15531 | 493 | in(instantiate' [] [SOME cd, SOME cz] (not_dvd_eq_pinf)) | 
| 13876 | 494 | end | 
| 495 | ||
| 15531 | 496 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) | 
| 13876 | 497 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 498 |       |(Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 499 | if y=x then let val cd = cterm_of sg (norm_zero_one d) | 
| 500 | val cz = cterm_of sg (norm_zero_one z) | |
| 15531 | 501 | in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_pinf)) | 
| 13876 | 502 | end | 
| 15531 | 503 | else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) | 
| 13876 | 504 | |
| 505 | ||
| 15531 | 506 | |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) | 
| 13876 | 507 | end; | 
| 508 | ||
| 509 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 510 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 511 | fun minf_proof_of_c sg x dlcm t = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 512 | let val minf_eqth = thm_of sg (decomp_minf_eq x dlcm sg) t | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 513 | val minf_moddth = thm_of sg (decomp_minf_modd x dlcm sg) t | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 514 | in (minf_eqth, minf_moddth) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 515 | end; | 
| 13876 | 516 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 517 | (*=========== pinf proofs with the compact version==========*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 518 | fun decomp_pinf_eq x dlcm sg t = case t of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 519 |    Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_conjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 520 |    |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 521 | |_ =>([],fn [] => rho_for_eq_pinf x dlcm sg t) ; | 
| 13876 | 522 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 523 | fun decomp_pinf_modd x dlcm sg t = case t of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 524 |    Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_conjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 525 |    |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 526 | |_ => ([],fn [] => rho_for_modd_pinf x dlcm sg t); | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 527 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 528 | fun pinf_proof_of_c sg x dlcm t = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 529 | let val pinf_eqth = thm_of sg (decomp_pinf_eq x dlcm sg) t | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 530 | val pinf_moddth = thm_of sg (decomp_pinf_modd x dlcm sg) t | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 531 | in (pinf_eqth,pinf_moddth) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 532 | end; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 533 | |
| 13876 | 534 | |
| 535 | (* ------------------------------------------------------------------------- *) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 536 | (* Here we generate the theorem for the Bset Property in the simple direction*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 537 | (* It is just an instantiation*) | 
| 13876 | 538 | (* ------------------------------------------------------------------------- *) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 539 | (* | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 540 | fun bsetproof_of sg (x as Free(xn,xT)) fm bs dlcm = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 541 | let | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 542 | val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm))) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 543 | val cdlcm = cterm_of sg dlcm | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 544 | val cB = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one bs)) | 
| 15531 | 545 | in instantiate' [] [SOME cdlcm,SOME cB, SOME cp] (bst_thm) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 546 | end; | 
| 13876 | 547 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 548 | fun asetproof_of sg (x as Free(xn,xT)) fm ast dlcm = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 549 | let | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 550 | val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm))) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 551 | val cdlcm = cterm_of sg dlcm | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 552 | val cA = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one ast)) | 
| 15531 | 553 | in instantiate' [] [SOME cdlcm,SOME cA, SOME cp] (ast_thm) | 
| 13876 | 554 | end; | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 555 | *) | 
| 13876 | 556 | |
| 557 | (* For the generation of atomic Theorems*) | |
| 558 | (* Prove the premisses on runtime and then make RS*) | |
| 559 | (* ------------------------------------------------------------------------- *) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 560 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 561 | (*========= this is rho ============*) | 
| 13876 | 562 | fun generate_atomic_not_bst_p sg (x as Free(xn,xT)) fm dlcm B at = | 
| 563 | let | |
| 564 | val cdlcm = cterm_of sg dlcm | |
| 565 | val cB = cterm_of sg B | |
| 566 | val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm))) | |
| 567 | val cat = cterm_of sg (norm_zero_one at) | |
| 568 | in | |
| 569 | case at of | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 570 |    (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z))) => 
 | 
| 13876 | 571 | if (x=y) andalso (c1=zero) andalso (c2=one) | 
| 572 | 	 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)
 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 573 | 	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
 | 
| 19277 | 574 | 		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("0",HOLogic.intT),dlcm))
 | 
| 15531 | 575 | in (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_bst_p_ne))) | 
| 13876 | 576 | end | 
| 15531 | 577 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) | 
| 13876 | 578 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 579 |    |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("HOL.plus", T) $(Const ("HOL.times",_) $ c2 $ y) $z)) =>
 | 
| 13876 | 580 | if (is_arith_rel at) andalso (x=y) | 
| 581 | then let val bst_z = norm_zero_one (linear_neg (linear_add [] z (mk_numeral 1))) | |
| 582 | 	         in let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ bst_z $ B)
 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 583 | 	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq (bst_z,Const("HOL.minus",T) $ (Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
 | 
| 19277 | 584 | 		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("0",HOLogic.intT),dlcm))
 | 
| 15531 | 585 | in (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_bst_p_eq))) | 
| 13876 | 586 | end | 
| 587 | end | |
| 15531 | 588 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) | 
| 13876 | 589 | |
| 19277 | 590 |    |(Const("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ pm1 $ y ) $ z )) =>
 | 
| 13876 | 591 | if (y=x) andalso (c1 =zero) then | 
| 592 | if pm1 = one then | |
| 593 | 	  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)
 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 594 |               val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
 | 
| 15531 | 595 | in (instantiate' [] [SOME cfma, SOME cdlcm]([th1,th2] MRS (not_bst_p_gt))) | 
| 13876 | 596 | end | 
| 19277 | 597 | 	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("0",HOLogic.intT),dlcm))
 | 
| 15531 | 598 | in (instantiate' [] [SOME cfma, SOME cB,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt))) | 
| 13876 | 599 | end | 
| 15531 | 600 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) | 
| 13876 | 601 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 602 |    |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 603 | if y=x then | 
| 604 | let val cz = cterm_of sg (norm_zero_one z) | |
| 605 | 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) | |
| 15531 | 606 | in (instantiate' [] [SOME cfma, SOME cB,SOME cz] (th1 RS (not_bst_p_ndvd))) | 
| 13876 | 607 | end | 
| 15531 | 608 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) | 
| 13876 | 609 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 610 |    |(Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 611 | if y=x then | 
| 612 | let val cz = cterm_of sg (norm_zero_one z) | |
| 613 | 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) | |
| 15531 | 614 | in (instantiate' [] [SOME cfma,SOME cB,SOME cz] (th1 RS (not_bst_p_dvd))) | 
| 13876 | 615 | end | 
| 15531 | 616 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) | 
| 13876 | 617 | |
| 15531 | 618 | |_ => (instantiate' [] [SOME cfma, SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm)) | 
| 13876 | 619 | |
| 620 | end; | |
| 621 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 622 | |
| 13876 | 623 | (* ------------------------------------------------------------------------- *) | 
| 624 | (* Main interpretation function for this backwards dirction*) | |
| 625 | (* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*) | |
| 626 | (*Help Function*) | |
| 627 | (* ------------------------------------------------------------------------- *) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 628 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 629 | (*==================== Proof with the compact version *) | 
| 13876 | 630 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 631 | fun decomp_nbstp sg x dlcm B fm t = case t of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 632 |    Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_conjI )
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 633 |   |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 634 | |_ => ([], fn [] => generate_atomic_not_bst_p sg x fm dlcm B t); | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 635 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 636 | fun not_bst_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm B t = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 637 | let | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 638 | val th = thm_of sg (decomp_nbstp sg x dlcm (list_to_set xT (map norm_zero_one B)) fm) t | 
| 13876 | 639 | val fma = absfree (xn,xT, norm_zero_one fm) | 
| 640 | in let val th1 = prove_elementar sg "ss" (HOLogic.mk_eq (fma,fma)) | |
| 641 | in [th,th1] MRS (not_bst_p_Q_elim) | |
| 642 | end | |
| 643 | end; | |
| 644 | ||
| 645 | ||
| 646 | (* ------------------------------------------------------------------------- *) | |
| 647 | (* Protokol interpretation function for the backwards direction for cooper's Theorem*) | |
| 648 | ||
| 649 | (* For the generation of atomic Theorems*) | |
| 650 | (* Prove the premisses on runtime and then make RS*) | |
| 651 | (* ------------------------------------------------------------------------- *) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 652 | (*========= this is rho ============*) | 
| 13876 | 653 | fun generate_atomic_not_ast_p sg (x as Free(xn,xT)) fm dlcm A at = | 
| 654 | let | |
| 655 | val cdlcm = cterm_of sg dlcm | |
| 656 | val cA = cterm_of sg A | |
| 657 | val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm))) | |
| 658 | val cat = cterm_of sg (norm_zero_one at) | |
| 659 | in | |
| 660 | case at of | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 661 |    (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("HOL.plus", _) $(Const ("HOL.times",_) $ c2 $ y) $z))) => 
 | 
| 13876 | 662 | if (x=y) andalso (c1=zero) andalso (c2=one) | 
| 663 | 	 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)
 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 664 | 	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
 | 
| 19277 | 665 | 		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("0",HOLogic.intT),dlcm))
 | 
| 15531 | 666 | in (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_ast_p_ne))) | 
| 13876 | 667 | end | 
| 15531 | 668 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) | 
| 13876 | 669 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 670 |    |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("HOL.plus", T) $(Const ("HOL.times",_) $ c2 $ y) $z)) =>
 | 
| 13876 | 671 | if (is_arith_rel at) andalso (x=y) | 
| 672 | then let val ast_z = norm_zero_one (linear_sub [] one z ) | |
| 673 | 	         val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ ast_z $ A)
 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 674 | 	         val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq (ast_z,Const("HOL.plus",T) $ (Const("HOL.uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
 | 
| 19277 | 675 | 		 val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("0",HOLogic.intT),dlcm))
 | 
| 15531 | 676 | in (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_ast_p_eq))) | 
| 13876 | 677 | end | 
| 15531 | 678 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) | 
| 13876 | 679 | |
| 19277 | 680 |    |(Const("Orderings.less",_) $ c1 $(Const ("HOL.plus", _) $(Const ("HOL.times",_) $ pm1 $ y ) $ z )) =>
 | 
| 13876 | 681 | if (y=x) andalso (c1 =zero) then | 
| 682 | if pm1 = (mk_numeral ~1) then | |
| 683 | 	  let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one z) $ A)
 | |
| 19277 | 684 | val th2 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (zero,dlcm)) | 
| 15531 | 685 | in (instantiate' [] [SOME cfma]([th2,th1] MRS (not_ast_p_lt))) | 
| 13876 | 686 | end | 
| 19277 | 687 | 	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (Const("0",HOLogic.intT),dlcm))
 | 
| 15531 | 688 | in (instantiate' [] [SOME cfma, SOME cA,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt))) | 
| 13876 | 689 | end | 
| 15531 | 690 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) | 
| 13876 | 691 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 692 |    |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 693 | if y=x then | 
| 694 | let val cz = cterm_of sg (norm_zero_one z) | |
| 695 | 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) | |
| 15531 | 696 | in (instantiate' [] [SOME cfma, SOME cA,SOME cz] (th1 RS (not_ast_p_ndvd))) | 
| 13876 | 697 | end | 
| 15531 | 698 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) | 
| 13876 | 699 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
17985diff
changeset | 700 |    |(Const("Divides.op dvd",_)$ d $ (Const ("HOL.plus",_) $ (Const ("HOL.times",_) $ c $ y ) $ z)) => 
 | 
| 13876 | 701 | if y=x then | 
| 702 | let val cz = cterm_of sg (norm_zero_one z) | |
| 703 | 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) | |
| 15531 | 704 | in (instantiate' [] [SOME cfma,SOME cA,SOME cz] (th1 RS (not_ast_p_dvd))) | 
| 13876 | 705 | end | 
| 15531 | 706 | else (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) | 
| 13876 | 707 | |
| 15531 | 708 | |_ => (instantiate' [] [SOME cfma, SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm)) | 
| 13876 | 709 | |
| 710 | end; | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 711 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 712 | (* ------------------------------------------------------------------------ *) | 
| 13876 | 713 | (* Main interpretation function for this backwards dirction*) | 
| 714 | (* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*) | |
| 715 | (*Help Function*) | |
| 716 | (* ------------------------------------------------------------------------- *) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 717 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 718 | fun decomp_nastp sg x dlcm A fm t = case t of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 719 |    Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_conjI )
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 720 |   |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 721 | |_ => ([], fn [] => generate_atomic_not_ast_p sg x fm dlcm A t); | 
| 13876 | 722 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 723 | fun not_ast_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm A t = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 724 | let | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 725 | val th = thm_of sg (decomp_nastp sg x dlcm (list_to_set xT (map norm_zero_one A)) fm) t | 
| 13876 | 726 | val fma = absfree (xn,xT, norm_zero_one fm) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 727 | in let val th1 = prove_elementar sg "ss" (HOLogic.mk_eq (fma,fma)) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 728 | in [th,th1] MRS (not_ast_p_Q_elim) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 729 | end | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 730 | end; | 
| 13876 | 731 | |
| 732 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 733 | (* -------------------------------*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 734 | (* Finding rho and beta for evalc *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 735 | (* -------------------------------*) | 
| 13876 | 736 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 737 | fun rho_for_evalc sg at = case at of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 738 | (Const (p,_) $ s $ t) =>( | 
| 17485 | 739 | case AList.lookup (op =) operations p of | 
| 15531 | 740 | SOME f => | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 741 | ((if (f ((dest_numeral s),(dest_numeral t))) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 742 | then prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 743 | else prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const))) | 
| 15531 | 744 | handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl) | 
| 745 | | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl ) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 746 |      |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
 | 
| 17485 | 747 | case AList.lookup (op =) operations p of | 
| 15531 | 748 | SOME f => | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 749 | ((if (f ((dest_numeral s),(dest_numeral t))) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 750 | then prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const)) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 751 | else prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const))) | 
| 15531 | 752 | handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl) | 
| 753 | | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl ) | |
| 754 | | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl; | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 755 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 756 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 757 | (*=========================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 758 | fun decomp_evalc sg t = case t of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 759 |    (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 760 |    |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 761 |    |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 762 |    |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 763 | |_ => ([], fn [] => rho_for_evalc sg t); | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 764 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 765 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 766 | fun proof_of_evalc sg fm = thm_of sg (decomp_evalc sg) fm; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 767 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 768 | (*==================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 769 | (* Proof of linform with the compact model *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 770 | (*==================================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 771 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 772 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 773 | fun decomp_linform sg vars t = case t of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 774 |    (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 775 |    |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 776 |    |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 777 |    |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 778 |    |(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not)
 | 
| 15164 | 779 |    |(Const("Divides.op dvd",_)$d$r) => 
 | 
| 15531 | 780 | if is_numeral d then ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [NONE , NONE, SOME (cterm_of sg d)](linearize_dvd))) | 
| 15164 | 781 | else (warning "Nonlinear Term --- Non numeral leftside at dvd"; | 
| 782 | raise COOPER) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 783 | |_ => ([], fn [] => prove_elementar sg "lf" (HOLogic.mk_eq (t, linform vars t))); | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 784 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 785 | fun proof_of_linform sg vars f = thm_of sg (decomp_linform sg vars) f; | 
| 13876 | 786 | |
| 787 | (* ------------------------------------------------------------------------- *) | |
| 788 | (* Interpretaion of Protocols of the cooper procedure : minusinfinity version*) | |
| 789 | (* ------------------------------------------------------------------------- *) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 790 | fun coopermi_proof_of sg (x as Free(xn,xT)) fm B dlcm = | 
| 13876 | 791 | (* Get the Bset thm*) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 792 | let val (minf_eqth, minf_moddth) = minf_proof_of_c sg x dlcm fm | 
| 19277 | 793 | val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (zero,dlcm)); | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 794 | val nbstpthm = not_bst_p_proof_of_c sg x fm dlcm B fm | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 795 | in (dpos,minf_eqth,nbstpthm,minf_moddth) | 
| 13876 | 796 | end; | 
| 797 | ||
| 798 | (* ------------------------------------------------------------------------- *) | |
| 799 | (* Interpretaion of Protocols of the cooper procedure : plusinfinity version *) | |
| 800 | (* ------------------------------------------------------------------------- *) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 801 | fun cooperpi_proof_of sg (x as Free(xn,xT)) fm A dlcm = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 802 | let val (pinf_eqth,pinf_moddth) = pinf_proof_of_c sg x dlcm fm | 
| 19277 | 803 | val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "Orderings.less" (zero,dlcm)); | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 804 | val nastpthm = not_ast_p_proof_of_c sg x fm dlcm A fm | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 805 | in (dpos,pinf_eqth,nastpthm,pinf_moddth) | 
| 13876 | 806 | end; | 
| 807 | ||
| 808 | (* ------------------------------------------------------------------------- *) | |
| 809 | (* Interpretaion of Protocols of the cooper procedure : full version*) | |
| 810 | (* ------------------------------------------------------------------------- *) | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 811 | fun cooper_thm sg s (x as Free(xn,xT)) cfm dlcm ast bst= case s of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 812 | "pi" => let val (dpsthm,pinf_eqth,nbpth,pinf_moddth) = cooperpi_proof_of sg x cfm ast dlcm | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 813 | in [dpsthm,pinf_eqth,nbpth,pinf_moddth] MRS (cppi_eq) | 
| 13876 | 814 | end | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 815 | |"mi" => let val (dpsthm,minf_eqth,nbpth,minf_moddth) = coopermi_proof_of sg x cfm bst dlcm | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 816 | in [dpsthm,minf_eqth,nbpth,minf_moddth] MRS (cpmi_eq) | 
| 13876 | 817 | end | 
| 818 | |_ => error "parameter error"; | |
| 819 | ||
| 820 | (* ------------------------------------------------------------------------- *) | |
| 821 | (* This function should evoluate to the end prove Procedure for one quantifier elimination for Presburger arithmetic*) | |
| 822 | (* It shoud be plugged in the qfnp argument of the quantifier elimination proof function*) | |
| 823 | (* ------------------------------------------------------------------------- *) | |
| 824 | ||
| 15165 | 825 | (* val (timef:(unit->thm) -> thm,prtime,time_reset) = gen_timer();*) | 
| 826 | (* val (timef2:(unit->thm) -> thm,prtime2,time_reset2) = gen_timer(); *) | |
| 15164 | 827 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 828 | fun cooper_prv sg (x as Free(xn,xT)) efm = let | 
| 14877 | 829 | (* lfm_thm : efm = linearized form of efm*) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 830 | val lfm_thm = proof_of_linform sg [xn] efm | 
| 14877 | 831 | (*efm2 is the linearized form of efm *) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 832 | val efm2 = snd(qe_get_terms lfm_thm) | 
| 14877 | 833 | (* l is the lcm of all coefficients of x *) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 834 | val l = formlcm x efm2 | 
| 14877 | 835 | (*ac_thm: efm = efm2 with adjusted coefficients of x *) | 
| 836 | val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans | |
| 837 | (* fm is efm2 with adjusted coefficients of x *) | |
| 13876 | 838 | val fm = snd (qe_get_terms ac_thm) | 
| 14877 | 839 | (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*) | 
| 13876 | 840 | val cfm = unitycoeff x fm | 
| 14877 | 841 | (*afm is fm where c*x is replaced by 1*x or -1*x *) | 
| 13876 | 842 | val afm = adjustcoeff x l fm | 
| 14877 | 843 | (* P = %x.afm*) | 
| 13876 | 844 | val P = absfree(xn,xT,afm) | 
| 14877 | 845 |    (* This simpset allows the elimination of the sets in bex {1..d} *)
 | 
| 13876 | 846 | val ss = presburger_ss addsimps | 
| 847 | [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff] | |
| 14877 | 848 | (* uth : EX x.P(l*x) = EX x. l dvd x & P x*) | 
| 15531 | 849 | val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_numeral l))] (unity_coeff_ex) | 
| 14877 | 850 | (* e_ac_thm : Ex x. efm = EX x. fm*) | 
| 13876 | 851 | val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI) | 
| 14877 | 852 | (* A and B set of the formula*) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 853 | val A = aset x cfm | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 854 | val B = bset x cfm | 
| 14877 | 855 | (* the divlcm (delta) of the formula*) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 856 | val dlcm = mk_numeral (divlcm x cfm) | 
| 14877 | 857 | (* Which set is smaller to generate the (hoepfully) shorter proof*) | 
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 858 | val cms = if ((length A) < (length B )) then "pi" else "mi" | 
| 15165 | 859 | (* val _ = if cms = "pi" then writeln "Plusinfinity" else writeln "Minusinfinity"*) | 
| 14877 | 860 | (* synthesize the proof of cooper's theorem*) | 
| 861 | (* cp_thm: EX x. cfm = Q*) | |
| 15165 | 862 | val cp_thm = cooper_thm sg cms x cfm dlcm A B | 
| 14877 | 863 |    (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*)
 | 
| 864 | (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*) | |
| 15165 | 865 | (* | 
| 15164 | 866 | val _ = prth cp_thm | 
| 867 | val _ = writeln "Expanding the bounded EX..." | |
| 15165 | 868 | *) | 
| 869 | val exp_cp_thm = refl RS (simplify ss (cp_thm RSN (2,trans))) | |
| 870 | (* | |
| 871 | val _ = writeln "Expanded" *) | |
| 14877 | 872 | (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*) | 
| 13876 | 873 | val (lsuth,rsuth) = qe_get_terms (uth) | 
| 14877 | 874 | (* lseacth = EX x. efm; rseacth = EX x. fm*) | 
| 13876 | 875 | val (lseacth,rseacth) = qe_get_terms(e_ac_thm) | 
| 14877 | 876 | (* lscth = EX x. cfm; rscth = Q' *) | 
| 13876 | 877 | val (lscth,rscth) = qe_get_terms (exp_cp_thm) | 
| 14877 | 878 | (* u_c_thm: EX x. P(l*x) = Q'*) | 
| 13876 | 879 | val u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans | 
| 14877 | 880 | (* result: EX x. efm = Q'*) | 
| 13876 | 881 | in ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans) | 
| 882 | end | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 883 | |cooper_prv _ _ _ = error "Parameters format"; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 884 | |
| 15107 | 885 | (* **************************************** *) | 
| 886 | (* An Other Version of cooper proving *) | |
| 887 | (* by giving a withness for EX *) | |
| 888 | (* **************************************** *) | |
| 889 | ||
| 890 | ||
| 891 | ||
| 892 | fun cooper_prv_w sg (x as Free(xn,xT)) efm = let | |
| 893 | (* lfm_thm : efm = linearized form of efm*) | |
| 894 | val lfm_thm = proof_of_linform sg [xn] efm | |
| 895 | (*efm2 is the linearized form of efm *) | |
| 896 | val efm2 = snd(qe_get_terms lfm_thm) | |
| 897 | (* l is the lcm of all coefficients of x *) | |
| 898 | val l = formlcm x efm2 | |
| 899 | (*ac_thm: efm = efm2 with adjusted coefficients of x *) | |
| 900 | val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans | |
| 901 | (* fm is efm2 with adjusted coefficients of x *) | |
| 902 | val fm = snd (qe_get_terms ac_thm) | |
| 903 | (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*) | |
| 904 | val cfm = unitycoeff x fm | |
| 905 | (*afm is fm where c*x is replaced by 1*x or -1*x *) | |
| 906 | val afm = adjustcoeff x l fm | |
| 907 | (* P = %x.afm*) | |
| 908 | val P = absfree(xn,xT,afm) | |
| 909 |    (* This simpset allows the elimination of the sets in bex {1..d} *)
 | |
| 910 | val ss = presburger_ss addsimps | |
| 911 | [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff] | |
| 912 | (* uth : EX x.P(l*x) = EX x. l dvd x & P x*) | |
| 15531 | 913 | val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_numeral l))] (unity_coeff_ex) | 
| 15107 | 914 | (* e_ac_thm : Ex x. efm = EX x. fm*) | 
| 915 | val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI) | |
| 916 | (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*) | |
| 917 | val (lsuth,rsuth) = qe_get_terms (uth) | |
| 918 | (* lseacth = EX x. efm; rseacth = EX x. fm*) | |
| 919 | val (lseacth,rseacth) = qe_get_terms(e_ac_thm) | |
| 920 | ||
| 921 | val (w,rs) = cooper_w [] cfm | |
| 922 | val exp_cp_thm = case w of | |
| 923 | (* FIXME - e_ac_thm just tipped to test syntactical correctness of the program!!!!*) | |
| 15531 | 924 | SOME n => e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*) | 
| 15107 | 925 | |_ => let | 
| 926 | (* A and B set of the formula*) | |
| 927 | val A = aset x cfm | |
| 928 | val B = bset x cfm | |
| 929 | (* the divlcm (delta) of the formula*) | |
| 930 | val dlcm = mk_numeral (divlcm x cfm) | |
| 931 | (* Which set is smaller to generate the (hoepfully) shorter proof*) | |
| 932 | val cms = if ((length A) < (length B )) then "pi" else "mi" | |
| 933 | (* synthesize the proof of cooper's theorem*) | |
| 934 | (* cp_thm: EX x. cfm = Q*) | |
| 935 | val cp_thm = cooper_thm sg cms x cfm dlcm A B | |
| 936 |      (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*)
 | |
| 937 | (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*) | |
| 938 | in refl RS (simplify ss (cp_thm RSN (2,trans))) | |
| 939 | end | |
| 940 | (* lscth = EX x. cfm; rscth = Q' *) | |
| 941 | val (lscth,rscth) = qe_get_terms (exp_cp_thm) | |
| 942 | (* u_c_thm: EX x. P(l*x) = Q'*) | |
| 943 | val u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans | |
| 944 | (* result: EX x. efm = Q'*) | |
| 945 | in ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans) | |
| 946 | end | |
| 947 | |cooper_prv_w _ _ _ = error "Parameters format"; | |
| 948 | ||
| 949 | ||
| 13876 | 950 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 951 | fun decomp_cnnf sg lfnp P = case P of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 952 |      Const ("op &",_) $ p $q => ([p,q] , fn [th1,th2] => [th1,th2] MRS qe_conjI )
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 953 |    |Const ("op |",_) $ p $q => ([p,q] , fn [th1,th2] => [th1,th2] MRS  qe_disjI)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 954 |    |Const ("Not",_) $ (Const("Not",_) $ p) => ([p], fn [th] => th RS nnf_nn)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 955 |    |Const("Not",_) $ (Const(opn,T) $ p $ q) => 
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 956 | if opn = "op |" | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 957 | then case (p,q) of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 958 |          (A as (Const ("op &",_) $ r $ s),B as (Const ("op &",_) $ r1 $ t)) =>
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 959 | if r1 = negate r | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 960 | 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) | 
| 13876 | 961 | |
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 962 | else ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 963 | |(_,_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 964 | else ( | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 965 | case (opn,T) of | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 966 |            ("op &",_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_ncj )
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 967 |            |("op -->",_) => ([p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_nim )
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 968 |            |("op =",Type ("fun",[Type ("bool", []),_])) => 
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 969 | ([HOLogic.conj $ p $ (HOLogic.Not $ q),HOLogic.conj $ (HOLogic.Not $ p) $ q], fn [th1,th2] => [th1,th2] MRS nnf_neq) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 970 | |(_,_) => ([], fn [] => lfnp P) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 971 | ) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 972 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 973 |    |(Const ("op -->",_) $ p $ q) => ([HOLogic.Not$p,q], fn [th1,th2] => [th1,th2] MRS nnf_im)
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 974 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 975 |    |(Const ("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q) =>
 | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 976 | ([HOLogic.conj $ p $ q,HOLogic.conj $ (HOLogic.Not $ p) $ (HOLogic.Not $ q) ], fn [th1,th2] =>[th1,th2] MRS nnf_eq ) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 977 | |_ => ([], fn [] => lfnp P); | 
| 13876 | 978 | |
| 979 | ||
| 980 | ||
| 981 | ||
| 14758 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 982 | fun proof_of_cnnf sg p lfnp = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 983 | let val th1 = thm_of sg (decomp_cnnf sg lfnp) p | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 984 | val rs = snd(qe_get_terms th1) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 985 | val th2 = prove_elementar sg "ss" (HOLogic.mk_eq(rs,simpl rs)) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 986 | in [th1,th2] MRS trans | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14479diff
changeset | 987 | end; | 
| 13876 | 988 | |
| 989 | end; | |
| 14920 
a7525235e20f
An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
 chaieb parents: 
14877diff
changeset | 990 |