| author | huffman | 
| Sat, 30 Sep 2006 17:10:55 +0200 | |
| changeset 20792 | add17d26151b | 
| parent 20713 | 823967ef47f1 | 
| child 20854 | f9cf9e62d11c | 
| permissions | -rw-r--r-- | 
| 13876 | 1 | (* Title: HOL/Integ/presburger.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Amine Chaieb and Stefan Berghofer, TU Muenchen | |
| 4 | ||
| 17465 | 5 | Tactic for solving arithmetical Goals in Presburger Arithmetic. | 
| 13876 | 6 | |
| 17465 | 7 | This version of presburger deals with occurences of functional symbols | 
| 8 | in the subgoal and abstract over them to try to prove the more general | |
| 9 | formula. It then resolves with the subgoal. To enable this feature | |
| 10 | call the procedure with the parameter abs. | |
| 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: 
14353diff
changeset | 11 | *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 12 | |
| 13876 | 13 | signature PRESBURGER = | 
| 14 | sig | |
| 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: 
14353diff
changeset | 15 | val presburger_tac : bool -> bool -> int -> tactic | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 16 | val presburger_method : bool -> bool -> int -> Proof.method | 
| 18708 | 17 | val setup : theory -> theory | 
| 13876 | 18 | val trace : bool ref | 
| 19 | end; | |
| 20 | ||
| 21 | structure Presburger: PRESBURGER = | |
| 22 | struct | |
| 23 | ||
| 24 | val trace = ref false; | |
| 25 | fun trace_msg s = if !trace then tracing s else (); | |
| 26 | ||
| 27 | (*-----------------------------------------------------------------*) | |
| 28 | (*cooper_pp: provefunction for the one-exstance quantifier elimination*) | |
| 29 | (* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*) | |
| 30 | (*-----------------------------------------------------------------*) | |
| 31 | ||
| 14941 | 32 | |
| 20053 | 33 | val presburger_ss = simpset (); | 
| 18202 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 34 | val binarith = map thm | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 35 | ["Pls_0_eq", "Min_1_eq", | 
| 20485 | 36 | "pred_Pls","pred_Min","pred_1","pred_0", | 
| 37 | "succ_Pls", "succ_Min", "succ_1", "succ_0", | |
| 38 | "add_Pls", "add_Min", "add_BIT_0", "add_BIT_10", | |
| 39 | "add_BIT_11", "minus_Pls", "minus_Min", "minus_1", | |
| 40 | "minus_0", "mult_Pls", "mult_Min", "mult_num1", "mult_num0", | |
| 41 | "add_Pls_right", "add_Min_right"]; | |
| 18202 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 42 | val intarithrel = | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 43 | (map thm ["int_eq_number_of_eq","int_neg_number_of_BIT", | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 44 | "int_le_number_of_eq","int_iszero_number_of_0", | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 45 | "int_less_number_of_eq_neg"]) @ | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 46 | (map (fn s => thm s RS thm "lift_bool") | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 47 | ["int_iszero_number_of_Pls","int_iszero_number_of_1", | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 48 | "int_neg_number_of_Min"])@ | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 49 | (map (fn s => thm s RS thm "nlift_bool") | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 50 | ["int_nonzero_number_of_Min","int_not_neg_number_of_Pls"]); | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 51 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 52 | val intarith = map thm ["int_number_of_add_sym", "int_number_of_minus_sym", | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 53 | "int_number_of_diff_sym", "int_number_of_mult_sym"]; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 54 | val natarith = map thm ["add_nat_number_of", "diff_nat_number_of", | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 55 | "mult_nat_number_of", "eq_nat_number_of", | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 56 | "less_nat_number_of"] | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 57 | val powerarith = | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 58 | (map thm ["nat_number_of", "zpower_number_of_even", | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 59 | "zpower_Pls", "zpower_Min"]) @ | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 60 | [(Tactic.simplify true [thm "zero_eq_Numeral0_nring", | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 61 | thm "one_eq_Numeral1_nring"] | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 62 | (thm "zpower_number_of_odd"))] | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 63 | |
| 18393 | 64 | val comp_arith = binarith @ intarith @ intarithrel @ natarith | 
| 18202 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 65 | @ powerarith @[thm"not_false_eq_true", thm "not_true_eq_false"]; | 
| 14801 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 berghofe parents: 
14758diff
changeset | 66 | |
| 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: 
14353diff
changeset | 67 | fun cooper_pp sg (fm as e$Abs(xn,xT,p)) = | 
| 20194 | 68 | let val (xn1,p1) = Syntax.variant_abs (xn,xT,p) | 
| 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: 
14353diff
changeset | 69 | in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1) end; | 
| 13876 | 70 | |
| 71 | fun mnnf_pp sg fm = CooperProof.proof_of_cnnf sg fm | |
| 72 | (CooperProof.proof_of_evalc sg); | |
| 73 | ||
| 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: 
14353diff
changeset | 74 | fun tmproof_of_int_qelim 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: 
14353diff
changeset | 75 | Qelim.tproof_of_mlift_qelim sg CooperDec.is_arith_rel | 
| 13876 | 76 | (CooperProof.proof_of_linform sg) (mnnf_pp sg) (cooper_pp sg) fm; | 
| 77 | ||
| 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: 
14353diff
changeset | 78 | |
| 13876 | 79 | (* Theorems to be used in this tactic*) | 
| 80 | ||
| 81 | val zdvd_int = thm "zdvd_int"; | |
| 82 | val zdiff_int_split = thm "zdiff_int_split"; | |
| 83 | val all_nat = thm "all_nat"; | |
| 84 | val ex_nat = thm "ex_nat"; | |
| 85 | val number_of1 = thm "number_of1"; | |
| 86 | val number_of2 = thm "number_of2"; | |
| 87 | val split_zdiv = thm "split_zdiv"; | |
| 88 | val split_zmod = thm "split_zmod"; | |
| 89 | val mod_div_equality' = thm "mod_div_equality'"; | |
| 90 | val split_div' = thm "split_div'"; | |
| 91 | val Suc_plus1 = thm "Suc_plus1"; | |
| 92 | val imp_le_cong = thm "imp_le_cong"; | |
| 93 | val conj_le_cong = thm "conj_le_cong"; | |
| 18202 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 94 | val nat_mod_add_eq = mod_add1_eq RS sym; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 95 | val nat_mod_add_left_eq = mod_add_left_eq RS sym; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 96 | val nat_mod_add_right_eq = mod_add_right_eq RS sym; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 97 | val int_mod_add_eq = zmod_zadd1_eq RS sym; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 98 | val int_mod_add_left_eq = zmod_zadd_left_eq RS sym; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 99 | val int_mod_add_right_eq = zmod_zadd_right_eq RS sym; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 100 | val nat_div_add_eq = div_add1_eq RS sym; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 101 | val int_div_add_eq = zdiv_zadd1_eq RS sym; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 102 | val ZDIVISION_BY_ZERO_MOD = DIVISION_BY_ZERO RS conjunct2; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 103 | val ZDIVISION_BY_ZERO_DIV = DIVISION_BY_ZERO RS conjunct1; | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 104 | |
| 13876 | 105 | |
| 106 | (* extract all the constants in a term*) | |
| 107 | fun add_term_typed_consts (Const (c, T), cs) = (c,T) ins cs | |
| 108 | | add_term_typed_consts (t $ u, cs) = | |
| 109 | add_term_typed_consts (t, add_term_typed_consts (u, cs)) | |
| 110 | | add_term_typed_consts (Abs (_, _, t), cs) = add_term_typed_consts (t, cs) | |
| 111 | | add_term_typed_consts (_, cs) = cs; | |
| 112 | ||
| 113 | fun term_typed_consts t = add_term_typed_consts(t,[]); | |
| 114 | ||
| 15661 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 wenzelm parents: 
15620diff
changeset | 115 | (* Some Types*) | 
| 13876 | 116 | val bT = HOLogic.boolT; | 
| 15620 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 paulson parents: 
15574diff
changeset | 117 | val bitT = HOLogic.bitT; | 
| 13876 | 118 | val iT = HOLogic.intT; | 
| 119 | val nT = HOLogic.natT; | |
| 120 | ||
| 121 | (* Allowed Consts in formulae for presburger tactic*) | |
| 122 | ||
| 123 | val allowed_consts = | |
| 124 |   [("All", (iT --> bT) --> bT),
 | |
| 125 |    ("Ex", (iT --> bT) --> bT),
 | |
| 126 |    ("All", (nT --> bT) --> bT),
 | |
| 127 |    ("Ex", (nT --> bT) --> bT),
 | |
| 128 | ||
| 129 |    ("op &", bT --> bT --> bT),
 | |
| 130 |    ("op |", bT --> bT --> bT),
 | |
| 131 |    ("op -->", bT --> bT --> bT),
 | |
| 132 |    ("op =", bT --> bT --> bT),
 | |
| 133 |    ("Not", bT --> bT),
 | |
| 134 | ||
| 19277 | 135 |    ("Orderings.less_eq", iT --> iT --> bT),
 | 
| 13876 | 136 |    ("op =", iT --> iT --> bT),
 | 
| 19277 | 137 |    ("Orderings.less", iT --> iT --> bT),
 | 
| 13876 | 138 |    ("Divides.op dvd", iT --> iT --> bT),
 | 
| 139 |    ("Divides.op div", iT --> iT --> iT),
 | |
| 140 |    ("Divides.op mod", iT --> iT --> iT),
 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 141 |    ("HOL.plus", iT --> iT --> iT),
 | 
| 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 142 |    ("HOL.minus", iT --> iT --> iT),
 | 
| 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 143 |    ("HOL.times", iT --> iT --> iT), 
 | 
| 13876 | 144 |    ("HOL.abs", iT --> iT),
 | 
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 145 |    ("HOL.uminus", iT --> iT),
 | 
| 14801 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 berghofe parents: 
14758diff
changeset | 146 |    ("HOL.max", iT --> iT --> iT),
 | 
| 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 berghofe parents: 
14758diff
changeset | 147 |    ("HOL.min", iT --> iT --> iT),
 | 
| 13876 | 148 | |
| 19277 | 149 |    ("Orderings.less_eq", nT --> nT --> bT),
 | 
| 13876 | 150 |    ("op =", nT --> nT --> bT),
 | 
| 19277 | 151 |    ("Orderings.less", nT --> nT --> bT),
 | 
| 13876 | 152 |    ("Divides.op dvd", nT --> nT --> bT),
 | 
| 153 |    ("Divides.op div", nT --> nT --> nT),
 | |
| 154 |    ("Divides.op mod", nT --> nT --> nT),
 | |
| 19233 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 155 |    ("HOL.plus", nT --> nT --> nT),
 | 
| 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 156 |    ("HOL.minus", nT --> nT --> nT),
 | 
| 
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
 haftmann parents: 
18708diff
changeset | 157 |    ("HOL.times", nT --> nT --> nT), 
 | 
| 13876 | 158 |    ("Suc", nT --> nT),
 | 
| 14801 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 berghofe parents: 
14758diff
changeset | 159 |    ("HOL.max", nT --> nT --> nT),
 | 
| 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 berghofe parents: 
14758diff
changeset | 160 |    ("HOL.min", nT --> nT --> nT),
 | 
| 13876 | 161 | |
| 15620 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 paulson parents: 
15574diff
changeset | 162 |    ("Numeral.bit.B0", bitT),
 | 
| 
8ccdc8bc66a2
replaced bool by a new datatype "bit" for binary numerals
 paulson parents: 
15574diff
changeset | 163 |    ("Numeral.bit.B1", bitT),
 | 
| 20485 | 164 |    ("Numeral.Bit", iT --> bitT --> iT),
 | 
| 165 |    ("Numeral.Pls", iT),
 | |
| 166 |    ("Numeral.Min", iT),
 | |
| 167 |    ("Numeral.number_of", iT --> iT),
 | |
| 168 |    ("Numeral.number_of", iT --> nT),
 | |
| 20713 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 haftmann parents: 
20485diff
changeset | 169 |    ("HOL.zero", nT),
 | 
| 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 haftmann parents: 
20485diff
changeset | 170 |    ("HOL.zero", iT),
 | 
| 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 haftmann parents: 
20485diff
changeset | 171 |    ("HOL.one", nT),
 | 
| 
823967ef47f1
renamed 0 and 1 to HOL.zero and HOL.one respectivly; introduced corresponding syntactic classes
 haftmann parents: 
20485diff
changeset | 172 |    ("HOL.one", iT),
 | 
| 13876 | 173 |    ("False", bT),
 | 
| 174 |    ("True", bT)];
 | |
| 175 | ||
| 176 | (* Preparation of the formula to be sent to the Integer quantifier *) | |
| 177 | (* elimination procedure *) | |
| 178 | (* Transforms meta implications and meta quantifiers to object *) | |
| 179 | (* implications and object quantifiers *) | |
| 180 | ||
| 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: 
14353diff
changeset | 181 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 182 | (*==================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 183 | (* Abstracting on subterms ========*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 184 | (*==================================*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 185 | (* Returns occurences of terms that are function application of type int or nat*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 186 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 187 | fun getfuncs fm = case strip_comb fm 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: 
14353diff
changeset | 188 | (Free (_, T), ts as _ :: _) => | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 189 | if body_type T mem [iT, nT] | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 190 | andalso not (ts = []) andalso forall (null o loose_bnos) 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: 
14353diff
changeset | 191 | then [fm] | 
| 15570 | 192 | else Library.foldl op union ([], map getfuncs ts) | 
| 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: 
14353diff
changeset | 193 | | (Var (_, T), ts as _ :: _) => | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 194 | if body_type T mem [iT, nT] | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 195 | andalso not (ts = []) andalso forall (null o loose_bnos) ts then [fm] | 
| 15570 | 196 | else Library.foldl op union ([], map getfuncs ts) | 
| 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: 
14353diff
changeset | 197 | | (Const (s, T), 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: 
14353diff
changeset | 198 | if (s, T) mem allowed_consts orelse not (body_type T mem [iT, nT]) | 
| 15570 | 199 | then Library.foldl op union ([], map getfuncs ts) | 
| 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: 
14353diff
changeset | 200 | else [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: 
14353diff
changeset | 201 | | (Abs (s, T, t), _) => getfuncs 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: 
14353diff
changeset | 202 | | _ => []; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 203 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 204 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 205 | fun abstract_pres sg fm = | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 206 | foldr (fn (t, u) => | 
| 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: 
14353diff
changeset | 207 | let val T = fastype_of 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: 
14353diff
changeset | 208 |       in all T $ Abs ("x", T, abstract_over (t, u)) end)
 | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 209 | fm (getfuncs 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: 
14353diff
changeset | 210 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 211 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 212 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 213 | (* hasfuncs_on_bounds dont care of the type of the functions applied! | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 214 | It returns true if there is a subterm coresponding to the application 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: 
14353diff
changeset | 215 | a function on a bounded variable. | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 216 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 217 | Function applications are allowed only for well predefined functions 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: 
14353diff
changeset | 218 | consts*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 219 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 220 | fun has_free_funcs fm = case strip_comb fm 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: 
14353diff
changeset | 221 | (Free (_, T), ts as _ :: _) => | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 222 | if (body_type T mem [iT,nT]) andalso (not (T mem [iT,nT])) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 223 | then true | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 224 | else exists (fn x => x) (map has_free_funcs 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: 
14353diff
changeset | 225 | | (Var (_, T), ts as _ :: _) => | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 226 | if (body_type T mem [iT,nT]) andalso not (T mem [iT,nT]) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 227 | then true | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 228 | else exists (fn x => x) (map has_free_funcs 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: 
14353diff
changeset | 229 | | (Const (s, T), ts) => exists (fn x => x) (map has_free_funcs 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: 
14353diff
changeset | 230 | | (Abs (s, T, t), _) => has_free_funcs 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: 
14353diff
changeset | 231 | |_ => false; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 232 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 233 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 234 | (*returns true if the formula is relevant for presburger arithmetic tactic | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 235 | The constants occuring in term t should be a subset of the allowed_consts | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 236 | There also should be no occurences of application of functions on bounded | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 237 | variables. Whenever this function will be used, it will be ensured that 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: 
14353diff
changeset | 238 | will not contain subterms with function symbols that could have been | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 239 | abstracted over.*) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 240 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 241 | fun relevant ps t = (term_typed_consts t) subset allowed_consts andalso | 
| 15570 | 242 | map (fn i => snd (List.nth (ps, i))) (loose_bnos 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: 
14353diff
changeset | 243 | map (snd o dest_Free) (term_frees t) @ map (snd o dest_Var) (term_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: 
14353diff
changeset | 244 | subset [iT, nT] | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 245 | andalso not (has_free_funcs 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: 
14353diff
changeset | 246 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 247 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 248 | fun prepare_for_presburger sg q fm = | 
| 13876 | 249 | let | 
| 250 | val ps = Logic.strip_params fm | |
| 251 | val hs = map HOLogic.dest_Trueprop (Logic.strip_assums_hyp fm) | |
| 252 | val c = HOLogic.dest_Trueprop (Logic.strip_assums_concl 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: 
14353diff
changeset | 253 | val _ = if relevant (rev ps) c 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: 
14353diff
changeset | 254 |                else  (trace_msg ("Conclusion is not a presburger term:\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: 
14353diff
changeset | 255 | Sign.string_of_term sg c); raise CooperDec.COOPER) | 
| 13876 | 256 | fun mk_all ((s, T), (P,n)) = | 
| 257 | if 0 mem loose_bnos P then | |
| 258 | (HOLogic.all_const T $ Abs (s, T, P), n) | |
| 259 | else (incr_boundvars ~1 P, n-1) | |
| 260 | fun mk_all2 (v, t) = HOLogic.all_const (fastype_of v) $ lambda v t; | |
| 15570 | 261 | val (rhs,irhs) = List.partition (relevant (rev ps)) hs | 
| 13876 | 262 | val np = length ps | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 263 | val (fm',np) = foldr (fn ((x, T), (fm,n)) => mk_all ((x, T), (fm,n))) | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 264 | (foldr HOLogic.mk_imp c rhs, np) ps | 
| 15570 | 265 | val (vs, _) = List.partition (fn t => q orelse (type_of t) = nT) | 
| 13876 | 266 | (term_frees fm' @ term_vars fm'); | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 267 | val fm2 = foldr mk_all2 fm' vs | 
| 13876 | 268 | in (fm2, np + length vs, length rhs) end; | 
| 269 | ||
| 270 | (*Object quantifier to meta --*) | |
| 271 | fun spec_step n th = if (n=0) then th else (spec_step (n-1) th) RS spec ; | |
| 272 | ||
| 273 | (* object implication to meta---*) | |
| 274 | fun mp_step n th = if (n=0) then th else (mp_step (n-1) th) RS mp; | |
| 275 | ||
| 276 | (* the presburger tactic*) | |
| 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: 
14353diff
changeset | 277 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 278 | (* Parameters : q = flag for quantify ofer free variables ; | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 279 | a = flag for abstracting over function occurences | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 280 | i = subgoal *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 281 | |
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 282 | fun presburger_tac q a i = ObjectLogic.atomize_tac i THEN (fn st => | 
| 13876 | 283 | let | 
| 17465 | 284 | val g = List.nth (prems_of st, i - 1) | 
| 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: 
14353diff
changeset | 285 | val sg = sign_of_thm st | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 286 | (* The Abstraction step *) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 287 | val g' = if a then abstract_pres sg g else g | 
| 13876 | 288 | (* Transform the term*) | 
| 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: 
14353diff
changeset | 289 | val (t,np,nh) = prepare_for_presburger sg q g' | 
| 15661 
9ef583b08647
reverted renaming of Some/None in comments and strings;
 wenzelm parents: 
15620diff
changeset | 290 | (* Some simpsets for dealing with mod div abs and nat*) | 
| 20255 | 291 | val mod_div_simpset = Simplifier.theory_context sg HOL_basic_ss | 
| 18202 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 292 | addsimps [refl,nat_mod_add_eq, nat_mod_add_left_eq, | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 293 | nat_mod_add_right_eq, int_mod_add_eq, | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 294 | int_mod_add_right_eq, int_mod_add_left_eq, | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 295 | nat_div_add_eq, int_div_add_eq, | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 296 | mod_self, zmod_self, | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 297 | DIVISION_BY_ZERO_MOD,DIVISION_BY_ZERO_DIV, | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 298 | ZDIVISION_BY_ZERO_MOD,ZDIVISION_BY_ZERO_DIV, | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 299 | zdiv_zero,zmod_zero,div_0,mod_0, | 
| 18393 | 300 | zdiv_1,zmod_1,div_1,mod_1, | 
| 301 | Suc_plus1] | |
| 18202 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 302 | addsimps add_ac | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 303 | addsimprocs [cancel_div_mod_proc] | 
| 13876 | 304 | val simpset0 = HOL_basic_ss | 
| 305 | addsimps [mod_div_equality', Suc_plus1] | |
| 18393 | 306 | addsimps comp_arith | 
| 13997 | 307 | addsplits [split_zdiv, split_zmod, split_div', split_min, split_max] | 
| 13876 | 308 | (* Simp rules for changing (n::int) to int n *) | 
| 309 | val simpset1 = HOL_basic_ss | |
| 310 | addsimps [nat_number_of_def, zdvd_int] @ map (fn r => r RS sym) | |
| 311 | [int_int_eq, zle_int, zless_int, zadd_int, zmult_int] | |
| 312 | addsplits [zdiff_int_split] | |
| 313 | (*simp rules for elimination of int n*) | |
| 314 | ||
| 315 | val simpset2 = HOL_basic_ss | |
| 316 | addsimps [nat_0_le, all_nat, ex_nat, number_of1, number_of2, int_0, int_1] | |
| 317 | addcongs [conj_le_cong, imp_le_cong] | |
| 318 | (* simp rules for elimination of abs *) | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14130diff
changeset | 319 | val simpset3 = HOL_basic_ss addsplits [abs_split] | 
| 13876 | 320 | val ct = cterm_of sg (HOLogic.mk_Trueprop t) | 
| 321 | (* Theorem for the nat --> int transformation *) | |
| 322 | val pre_thm = Seq.hd (EVERY | |
| 18202 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
18155diff
changeset | 323 | [simp_tac mod_div_simpset 1, simp_tac simpset0 1, | 
| 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: 
14353diff
changeset | 324 | TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), | 
| 14801 
2d27b5ebc447
- deleted unneeded function eta_long (now in Pure/pattern.ML
 berghofe parents: 
14758diff
changeset | 325 | TRY (simp_tac simpset3 1), TRY (simp_tac presburger_ss 1)] | 
| 13876 | 326 | (trivial ct)) | 
| 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: 
14353diff
changeset | 327 | fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) | 
| 13876 | 328 | (* The result of the quantifier elimination *) | 
| 329 | val (th, tac) = case (prop_of pre_thm) of | |
| 330 |         Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ =>
 | |
| 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: 
14882diff
changeset | 331 | let val pth = | 
| 
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: 
14882diff
changeset | 332 | (* If quick_and_dirty then run without proof generation as oracle*) | 
| 
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: 
14882diff
changeset | 333 | if !quick_and_dirty | 
| 16836 | 334 | then presburger_oracle sg (Pattern.eta_long [] t1) | 
| 14941 | 335 | (* | 
| 336 | assume (cterm_of sg | |
| 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: 
14882diff
changeset | 337 | (HOLogic.mk_Trueprop(HOLogic.mk_eq(t1,CooperDec.integer_qelim (Pattern.eta_long [] t1))))) | 
| 14941 | 338 | *) | 
| 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: 
14882diff
changeset | 339 | else tmproof_of_int_qelim sg (Pattern.eta_long [] t1) | 
| 
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: 
14882diff
changeset | 340 | in | 
| 13876 | 341 |           (trace_msg ("calling procedure with term:\n" ^
 | 
| 342 | Sign.string_of_term sg t1); | |
| 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: 
14882diff
changeset | 343 | ((pth RS iffD2) RS pre_thm, | 
| 13876 | 344 | assm_tac (i + 1) THEN (if q then I else TRY) (rtac TrueI i))) | 
| 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: 
14882diff
changeset | 345 | end | 
| 13876 | 346 | | _ => (pre_thm, assm_tac i) | 
| 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: 
14353diff
changeset | 347 | in (rtac (((mp_step nh) o (spec_step np)) th) i | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 348 | THEN tac) st | 
| 14130 | 349 | end handle Subscript => no_tac st | CooperDec.COOPER => no_tac st); | 
| 13876 | 350 | |
| 351 | fun presburger_args meth = | |
| 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: 
14353diff
changeset | 352 | let val parse_flag = | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 353 | Args.$$$ "no_quantify" >> K (apfst (K false)) | 
| 18155 
e5ab63ca6b0f
old argument "abs" is replaced by "no_abs". Abstraction is turned on by default.
 chaieb parents: 
17465diff
changeset | 354 | || Args.$$$ "no_abs" >> K (apsnd (K false)); | 
| 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: 
14353diff
changeset | 355 | 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: 
14353diff
changeset | 356 | Method.simple_args | 
| 14882 | 357 |   (Scan.optional (Args.$$$ "(" |-- Scan.repeat1 parse_flag --| Args.$$$ ")") [] >>
 | 
| 18155 
e5ab63ca6b0f
old argument "abs" is replaced by "no_abs". Abstraction is turned on by default.
 chaieb parents: 
17465diff
changeset | 358 | curry (Library.foldl op |>) (true, true)) | 
| 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: 
14353diff
changeset | 359 | (fn (q,a) => fn _ => meth q a 1) | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 360 | end; | 
| 13876 | 361 | |
| 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: 
14353diff
changeset | 362 | fun presburger_method q a i = Method.METHOD (fn facts => | 
| 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 chaieb parents: 
14353diff
changeset | 363 | Method.insert_tac facts 1 THEN presburger_tac q a i) | 
| 13876 | 364 | |
| 20413 | 365 | val presburger_arith_tac = mk_arith_tactic "presburger" (fn i => fn st => | 
| 366 | (warning "Trying full Presburger arithmetic ..."; | |
| 367 | presburger_tac true true i st)); | |
| 368 | ||
| 13876 | 369 | val setup = | 
| 18708 | 370 |   Method.add_method ("presburger",
 | 
| 371 | presburger_args presburger_method, | |
| 372 | "decision procedure for Presburger arithmetic") #> | |
| 20413 | 373 | arith_tactic_add presburger_arith_tac; | 
| 13876 | 374 | |
| 375 | end; | |
| 376 | ||
| 18155 
e5ab63ca6b0f
old argument "abs" is replaced by "no_abs". Abstraction is turned on by default.
 chaieb parents: 
17465diff
changeset | 377 | val presburger_tac = Presburger.presburger_tac true true; |