| author | wenzelm | 
| Wed, 25 Jan 2006 00:21:35 +0100 | |
| changeset 18777 | 9d98d5705433 | 
| parent 18202 | 46af82efd311 | 
| child 20051 | 859e7129961b | 
| permissions | -rw-r--r-- | 
| 13876 | 1 | (* Title: HOL/Integ/Presburger.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Amine Chaieb, Tobias Nipkow and Stefan Berghofer, TU Muenchen | |
| 4 | ||
| 5 | File containing necessary theorems for the proof | |
| 6 | generation for Cooper Algorithm | |
| 7 | *) | |
| 8 | ||
| 15131 | 9 | header {* Presburger Arithmetic: Cooper's Algorithm *}
 | 
| 14577 | 10 | |
| 15131 | 11 | theory Presburger | 
| 15140 | 12 | imports NatSimprocs SetInterval | 
| 17378 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 chaieb parents: 
16836diff
changeset | 13 | uses ("cooper_dec.ML") ("cooper_proof.ML") ("qelim.ML") 
 | 
| 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 chaieb parents: 
16836diff
changeset | 14 | 	("reflected_presburger.ML") ("reflected_cooper.ML") ("presburger.ML")
 | 
| 15131 | 15 | begin | 
| 13876 | 16 | |
| 14577 | 17 | text {* Theorem for unitifying the coeffitients of @{text x} in an existential formula*}
 | 
| 13876 | 18 | |
| 19 | theorem unity_coeff_ex: "(\<exists>x::int. P (l * x)) = (\<exists>x. l dvd (1*x+0) \<and> P x)" | |
| 20 | apply (rule iffI) | |
| 21 | apply (erule exE) | |
| 22 | apply (rule_tac x = "l * x" in exI) | |
| 23 | apply simp | |
| 24 | apply (erule exE) | |
| 25 | apply (erule conjE) | |
| 26 | apply (erule dvdE) | |
| 27 | apply (rule_tac x = k in exI) | |
| 28 | apply simp | |
| 29 | done | |
| 30 | ||
| 31 | lemma uminus_dvd_conv: "(d dvd (t::int)) = (-d dvd t)" | |
| 32 | apply(unfold dvd_def) | |
| 33 | apply(rule iffI) | |
| 34 | apply(clarsimp) | |
| 35 | apply(rename_tac k) | |
| 36 | apply(rule_tac x = "-k" in exI) | |
| 37 | apply simp | |
| 38 | apply(clarsimp) | |
| 39 | apply(rename_tac k) | |
| 40 | apply(rule_tac x = "-k" in exI) | |
| 41 | apply simp | |
| 42 | done | |
| 43 | ||
| 44 | lemma uminus_dvd_conv': "(d dvd (t::int)) = (d dvd -t)" | |
| 45 | apply(unfold dvd_def) | |
| 46 | apply(rule iffI) | |
| 47 | apply(clarsimp) | |
| 48 | apply(rule_tac x = "-k" in exI) | |
| 49 | apply simp | |
| 50 | apply(clarsimp) | |
| 51 | apply(rule_tac x = "-k" in exI) | |
| 52 | apply simp | |
| 53 | done | |
| 54 | ||
| 55 | ||
| 56 | ||
| 14577 | 57 | text {*Theorems for the combination of proofs of the equality of @{text P} and @{text P_m} for integers @{text x} less than some integer @{text z}.*}
 | 
| 13876 | 58 | |
| 59 | theorem eq_minf_conjI: "\<exists>z1::int. \<forall>x. x < z1 \<longrightarrow> (A1 x = A2 x) \<Longrightarrow> | |
| 60 | \<exists>z2::int. \<forall>x. x < z2 \<longrightarrow> (B1 x = B2 x) \<Longrightarrow> | |
| 61 | \<exists>z::int. \<forall>x. x < z \<longrightarrow> ((A1 x \<and> B1 x) = (A2 x \<and> B2 x))" | |
| 62 | apply (erule exE)+ | |
| 63 | apply (rule_tac x = "min z1 z2" in exI) | |
| 64 | apply simp | |
| 65 | done | |
| 66 | ||
| 67 | ||
| 68 | theorem eq_minf_disjI: "\<exists>z1::int. \<forall>x. x < z1 \<longrightarrow> (A1 x = A2 x) \<Longrightarrow> | |
| 69 | \<exists>z2::int. \<forall>x. x < z2 \<longrightarrow> (B1 x = B2 x) \<Longrightarrow> | |
| 70 | \<exists>z::int. \<forall>x. x < z \<longrightarrow> ((A1 x \<or> B1 x) = (A2 x \<or> B2 x))" | |
| 71 | ||
| 72 | apply (erule exE)+ | |
| 73 | apply (rule_tac x = "min z1 z2" in exI) | |
| 74 | apply simp | |
| 75 | done | |
| 76 | ||
| 77 | ||
| 14577 | 78 | text {*Theorems for the combination of proofs of the equality of @{text P} and @{text P_m} for integers @{text x} greather than some integer @{text z}.*}
 | 
| 13876 | 79 | |
| 80 | theorem eq_pinf_conjI: "\<exists>z1::int. \<forall>x. z1 < x \<longrightarrow> (A1 x = A2 x) \<Longrightarrow> | |
| 81 | \<exists>z2::int. \<forall>x. z2 < x \<longrightarrow> (B1 x = B2 x) \<Longrightarrow> | |
| 82 | \<exists>z::int. \<forall>x. z < x \<longrightarrow> ((A1 x \<and> B1 x) = (A2 x \<and> B2 x))" | |
| 83 | apply (erule exE)+ | |
| 84 | apply (rule_tac x = "max z1 z2" in exI) | |
| 85 | apply simp | |
| 86 | done | |
| 87 | ||
| 88 | ||
| 89 | theorem eq_pinf_disjI: "\<exists>z1::int. \<forall>x. z1 < x \<longrightarrow> (A1 x = A2 x) \<Longrightarrow> | |
| 90 | \<exists>z2::int. \<forall>x. z2 < x \<longrightarrow> (B1 x = B2 x) \<Longrightarrow> | |
| 91 | \<exists>z::int. \<forall>x. z < x \<longrightarrow> ((A1 x \<or> B1 x) = (A2 x \<or> B2 x))" | |
| 92 | apply (erule exE)+ | |
| 93 | apply (rule_tac x = "max z1 z2" in exI) | |
| 94 | apply simp | |
| 95 | done | |
| 14577 | 96 | |
| 97 | text {*
 | |
| 98 |   \medskip Theorems for the combination of proofs of the modulo @{text
 | |
| 99 |   D} property for @{text "P plusinfinity"}
 | |
| 100 | ||
| 101 |   FIXME: This is THE SAME theorem as for the @{text minusinf} version,
 | |
| 102 |   but with @{text "+k.."} instead of @{text "-k.."} In the future
 | |
| 103 | replace these both with only one. *} | |
| 13876 | 104 | |
| 105 | theorem modd_pinf_conjI: "\<forall>(x::int) k. A x = A (x+k*d) \<Longrightarrow> | |
| 106 | \<forall>(x::int) k. B x = B (x+k*d) \<Longrightarrow> | |
| 107 | \<forall>(x::int) (k::int). (A x \<and> B x) = (A (x+k*d) \<and> B (x+k*d))" | |
| 108 | by simp | |
| 109 | ||
| 110 | theorem modd_pinf_disjI: "\<forall>(x::int) k. A x = A (x+k*d) \<Longrightarrow> | |
| 111 | \<forall>(x::int) k. B x = B (x+k*d) \<Longrightarrow> | |
| 112 | \<forall>(x::int) (k::int). (A x \<or> B x) = (A (x+k*d) \<or> B (x+k*d))" | |
| 113 | by simp | |
| 114 | ||
| 14577 | 115 | text {*
 | 
| 116 | This is one of the cases where the simplifed formula is prooved to | |
| 117 |   habe some property (in relation to @{text P_m}) but we need to prove
 | |
| 118 |   the property for the original formula (@{text P_m})
 | |
| 119 | ||
| 120 |   FIXME: This is exaclty the same thm as for @{text minusinf}. *}
 | |
| 121 | ||
| 13876 | 122 | lemma pinf_simp_eq: "ALL x. P(x) = Q(x) ==> (EX (x::int). P(x)) --> (EX (x::int). F(x)) ==> (EX (x::int). Q(x)) --> (EX (x::int). F(x)) " | 
| 14577 | 123 | by blast | 
| 13876 | 124 | |
| 125 | ||
| 14577 | 126 | text {*
 | 
| 127 |   \medskip Theorems for the combination of proofs of the modulo @{text D}
 | |
| 128 |   property for @{text "P minusinfinity"} *}
 | |
| 13876 | 129 | |
| 130 | theorem modd_minf_conjI: "\<forall>(x::int) k. A x = A (x-k*d) \<Longrightarrow> | |
| 131 | \<forall>(x::int) k. B x = B (x-k*d) \<Longrightarrow> | |
| 132 | \<forall>(x::int) (k::int). (A x \<and> B x) = (A (x-k*d) \<and> B (x-k*d))" | |
| 133 | by simp | |
| 134 | ||
| 135 | theorem modd_minf_disjI: "\<forall>(x::int) k. A x = A (x-k*d) \<Longrightarrow> | |
| 136 | \<forall>(x::int) k. B x = B (x-k*d) \<Longrightarrow> | |
| 137 | \<forall>(x::int) (k::int). (A x \<or> B x) = (A (x-k*d) \<or> B (x-k*d))" | |
| 138 | by simp | |
| 139 | ||
| 14577 | 140 | text {*
 | 
| 141 | This is one of the cases where the simplifed formula is prooved to | |
| 142 |   have some property (in relation to @{text P_m}) but we need to
 | |
| 143 |   prove the property for the original formula (@{text P_m}). *}
 | |
| 13876 | 144 | |
| 145 | lemma minf_simp_eq: "ALL x. P(x) = Q(x) ==> (EX (x::int). P(x)) --> (EX (x::int). F(x)) ==> (EX (x::int). Q(x)) --> (EX (x::int). F(x)) " | |
| 14577 | 146 | by blast | 
| 13876 | 147 | |
| 14577 | 148 | text {*
 | 
| 149 | Theorem needed for proving at runtime divide properties using the | |
| 150 | arithmetic tactic (which knows only about modulo = 0). *} | |
| 13876 | 151 | |
| 152 | lemma zdvd_iff_zmod_eq_0: "(m dvd n) = (n mod m = (0::int))" | |
| 14577 | 153 | by(simp add:dvd_def zmod_eq_0_iff) | 
| 13876 | 154 | |
| 14577 | 155 | text {*
 | 
| 156 | \medskip Theorems used for the combination of proof for the | |
| 157 | backwards direction of Cooper's Theorem. They rely exclusively on | |
| 158 | Predicate calculus.*} | |
| 13876 | 159 | |
| 160 | lemma not_ast_p_disjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P1(x) --> P1(x + d))
 | |
| 161 | ==> | |
| 162 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P2(x) --> P2(x + d))
 | |
| 163 | ==> | |
| 164 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->(P1(x) \<or> P2(x)) --> (P1(x + d) \<or> P2(x + d))) "
 | |
| 14577 | 165 | by blast | 
| 13876 | 166 | |
| 167 | ||
| 168 | lemma not_ast_p_conjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a- j)) --> P1(x) --> P1(x + d))
 | |
| 169 | ==> | |
| 170 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> P2(x) --> P2(x + d))
 | |
| 171 | ==> | |
| 172 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->(P1(x) \<and> P2(x)) --> (P1(x + d)
 | |
| 173 | \<and> P2(x + d))) " | |
| 14577 | 174 | by blast | 
| 13876 | 175 | |
| 176 | lemma not_ast_p_Q_elim: " | |
| 177 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) -->P(x) --> P(x + d))
 | |
| 178 | ==> ( P = Q ) | |
| 179 | ==> (ALL x. ~(EX (j::int) : {1..d}. EX (a::int) : A. P(a - j)) -->P(x) --> P(x + d))"
 | |
| 14577 | 180 | by blast | 
| 13876 | 181 | |
| 14577 | 182 | text {*
 | 
| 183 | \medskip Theorems used for the combination of proof for the | |
| 184 | backwards direction of Cooper's Theorem. They rely exclusively on | |
| 185 | Predicate calculus.*} | |
| 13876 | 186 | |
| 187 | lemma not_bst_p_disjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P1(x) --> P1(x - d))
 | |
| 188 | ==> | |
| 189 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P2(x) --> P2(x - d))
 | |
| 190 | ==> | |
| 191 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->(P1(x) \<or> P2(x)) --> (P1(x - d)
 | |
| 192 | \<or> P2(x-d))) " | |
| 14577 | 193 | by blast | 
| 13876 | 194 | |
| 195 | lemma not_bst_p_conjI: "(ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P1(x) --> P1(x - d))
 | |
| 196 | ==> | |
| 197 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> P2(x) --> P2(x - d))
 | |
| 198 | ==> | |
| 199 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->(P1(x) \<and> P2(x)) --> (P1(x - d)
 | |
| 200 | \<and> P2(x-d))) " | |
| 14577 | 201 | by blast | 
| 13876 | 202 | |
| 203 | lemma not_bst_p_Q_elim: " | |
| 204 | (ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) -->P(x) --> P(x - d)) 
 | |
| 205 | ==> ( P = Q ) | |
| 206 | ==> (ALL x. ~(EX (j::int) : {1..d}. EX (b::int) : B. P(b+j)) -->P(x) --> P(x - d))"
 | |
| 14577 | 207 | by blast | 
| 13876 | 208 | |
| 14577 | 209 | text {* \medskip This is the first direction of Cooper's Theorem. *}
 | 
| 13876 | 210 | lemma cooper_thm: "(R --> (EX x::int. P x)) ==> (Q -->(EX x::int. P x )) ==> ((R|Q) --> (EX x::int. P x )) " | 
| 14577 | 211 | by blast | 
| 13876 | 212 | |
| 14577 | 213 | text {*
 | 
| 214 | \medskip The full Cooper's Theorem in its equivalence Form. Given | |
| 215 | the premises it is trivial too, it relies exclusively on prediacte calculus.*} | |
| 13876 | 216 | lemma cooper_eq_thm: "(R --> (EX x::int. P x)) ==> (Q -->(EX x::int. P x )) ==> ((~Q) | 
| 217 | --> (EX x::int. P x ) --> R) ==> (EX x::int. P x) = R|Q " | |
| 14577 | 218 | by blast | 
| 13876 | 219 | |
| 14577 | 220 | text {*
 | 
| 221 | \medskip Some of the atomic theorems generated each time the atom | |
| 222 |   does not depend on @{text x}, they are trivial.*}
 | |
| 13876 | 223 | |
| 224 | lemma fm_eq_minf: "EX z::int. ALL x. x < z --> (P = P) " | |
| 14577 | 225 | by blast | 
| 13876 | 226 | |
| 227 | lemma fm_modd_minf: "ALL (x::int). ALL (k::int). (P = P)" | |
| 14577 | 228 | by blast | 
| 13876 | 229 | |
| 230 | lemma not_bst_p_fm: "ALL (x::int). Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> fm --> fm"
 | |
| 14577 | 231 | by blast | 
| 13876 | 232 | |
| 233 | lemma fm_eq_pinf: "EX z::int. ALL x. z < x --> (P = P) " | |
| 14577 | 234 | by blast | 
| 13876 | 235 | |
| 14577 | 236 | text {* The next two thms are the same as the @{text minusinf} version. *}
 | 
| 237 | ||
| 13876 | 238 | lemma fm_modd_pinf: "ALL (x::int). ALL (k::int). (P = P)" | 
| 14577 | 239 | by blast | 
| 13876 | 240 | |
| 241 | lemma not_ast_p_fm: "ALL (x::int). Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> fm --> fm"
 | |
| 14577 | 242 | by blast | 
| 13876 | 243 | |
| 14577 | 244 | text {* Theorems to be deleted from simpset when proving simplified formulaes. *}
 | 
| 13876 | 245 | |
| 246 | lemma P_eqtrue: "(P=True) = P" | |
| 17589 | 247 | by iprover | 
| 13876 | 248 | |
| 249 | lemma P_eqfalse: "(P=False) = (~P)" | |
| 17589 | 250 | by iprover | 
| 13876 | 251 | |
| 14577 | 252 | text {*
 | 
| 253 | \medskip Theorems for the generation of the bachwards direction of | |
| 254 | Cooper's Theorem. | |
| 13876 | 255 | |
| 14577 | 256 | These are the 6 interesting atomic cases which have to be proved relying on the | 
| 257 | properties of B-set and the arithmetic and contradiction proofs. *} | |
| 13876 | 258 | |
| 259 | lemma not_bst_p_lt: "0 < (d::int) ==> | |
| 260 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ( 0 < -x + a) --> (0 < -(x - d) + a )"
 | |
| 14577 | 261 | by arith | 
| 13876 | 262 | |
| 263 | lemma not_bst_p_gt: "\<lbrakk> (g::int) \<in> B; g = -a \<rbrakk> \<Longrightarrow> | |
| 264 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> (0 < (x) + a) --> ( 0 < (x - d) + a)"
 | |
| 265 | apply clarsimp | |
| 266 | apply(rule ccontr) | |
| 267 | apply(drule_tac x = "x+a" in bspec) | |
| 268 | apply(simp add:atLeastAtMost_iff) | |
| 269 | apply(drule_tac x = "-a" in bspec) | |
| 270 | apply assumption | |
| 271 | apply(simp) | |
| 272 | done | |
| 273 | ||
| 274 | lemma not_bst_p_eq: "\<lbrakk> 0 < d; (g::int) \<in> B; g = -a - 1 \<rbrakk> \<Longrightarrow> | |
| 275 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> (0 = x + a) --> (0 = (x - d) + a )"
 | |
| 276 | apply clarsimp | |
| 277 | apply(subgoal_tac "x = -a") | |
| 278 | prefer 2 apply arith | |
| 279 | apply(drule_tac x = "1" in bspec) | |
| 280 | apply(simp add:atLeastAtMost_iff) | |
| 281 | apply(drule_tac x = "-a- 1" in bspec) | |
| 282 | apply assumption | |
| 283 | apply(simp) | |
| 284 | done | |
| 285 | ||
| 286 | ||
| 287 | lemma not_bst_p_ne: "\<lbrakk> 0 < d; (g::int) \<in> B; g = -a \<rbrakk> \<Longrightarrow> | |
| 288 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ~(0 = x + a) --> ~(0 = (x - d) + a)"
 | |
| 289 | apply clarsimp | |
| 290 | apply(subgoal_tac "x = -a+d") | |
| 291 | prefer 2 apply arith | |
| 292 | apply(drule_tac x = "d" in bspec) | |
| 293 | apply(simp add:atLeastAtMost_iff) | |
| 294 | apply(drule_tac x = "-a" in bspec) | |
| 295 | apply assumption | |
| 296 | apply(simp) | |
| 297 | done | |
| 298 | ||
| 299 | ||
| 300 | lemma not_bst_p_dvd: "(d1::int) dvd d ==> | |
| 301 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> d1 dvd (x + a) --> d1 dvd ((x - d) + a )"
 | |
| 302 | apply(clarsimp simp add:dvd_def) | |
| 303 | apply(rename_tac m) | |
| 304 | apply(rule_tac x = "m - k" in exI) | |
| 305 | apply(simp add:int_distrib) | |
| 306 | done | |
| 307 | ||
| 308 | lemma not_bst_p_ndvd: "(d1::int) dvd d ==> | |
| 309 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (b::int) : B. Q(b+j)) --> ~(d1 dvd (x + a)) --> ~(d1 dvd ((x - d) + a ))"
 | |
| 310 | apply(clarsimp simp add:dvd_def) | |
| 311 | apply(rename_tac m) | |
| 312 | apply(erule_tac x = "m + k" in allE) | |
| 313 | apply(simp add:int_distrib) | |
| 314 | done | |
| 315 | ||
| 14577 | 316 | text {*
 | 
| 317 | \medskip Theorems for the generation of the bachwards direction of | |
| 318 | Cooper's Theorem. | |
| 13876 | 319 | |
| 14577 | 320 | These are the 6 interesting atomic cases which have to be proved | 
| 321 | relying on the properties of A-set ant the arithmetic and | |
| 322 | contradiction proofs. *} | |
| 13876 | 323 | |
| 324 | lemma not_ast_p_gt: "0 < (d::int) ==> | |
| 325 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ( 0 < x + t) --> (0 < (x + d) + t )"
 | |
| 14577 | 326 | by arith | 
| 13876 | 327 | |
| 328 | lemma not_ast_p_lt: "\<lbrakk>0 < d ;(t::int) \<in> A \<rbrakk> \<Longrightarrow> | |
| 329 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> (0 < -x + t) --> ( 0 < -(x + d) + t)"
 | |
| 330 | apply clarsimp | |
| 331 | apply (rule ccontr) | |
| 332 | apply (drule_tac x = "t-x" in bspec) | |
| 333 | apply simp | |
| 334 | apply (drule_tac x = "t" in bspec) | |
| 335 | apply assumption | |
| 336 | apply simp | |
| 337 | done | |
| 338 | ||
| 339 | lemma not_ast_p_eq: "\<lbrakk> 0 < d; (g::int) \<in> A; g = -t + 1 \<rbrakk> \<Longrightarrow> | |
| 340 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> (0 = x + t) --> (0 = (x + d) + t )"
 | |
| 341 | apply clarsimp | |
| 342 | apply (drule_tac x="1" in bspec) | |
| 343 | apply simp | |
| 344 | apply (drule_tac x="- t + 1" in bspec) | |
| 345 | apply assumption | |
| 346 | apply(subgoal_tac "x = -t") | |
| 347 | prefer 2 apply arith | |
| 348 | apply simp | |
| 349 | done | |
| 350 | ||
| 351 | lemma not_ast_p_ne: "\<lbrakk> 0 < d; (g::int) \<in> A; g = -t \<rbrakk> \<Longrightarrow> | |
| 352 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ~(0 = x + t) --> ~(0 = (x + d) + t)"
 | |
| 353 | apply clarsimp | |
| 354 | apply (subgoal_tac "x = -t-d") | |
| 355 | prefer 2 apply arith | |
| 356 | apply (drule_tac x = "d" in bspec) | |
| 357 | apply simp | |
| 358 | apply (drule_tac x = "-t" in bspec) | |
| 359 | apply assumption | |
| 360 | apply simp | |
| 361 | done | |
| 362 | ||
| 363 | lemma not_ast_p_dvd: "(d1::int) dvd d ==> | |
| 364 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> d1 dvd (x + t) --> d1 dvd ((x + d) + t )"
 | |
| 365 | apply(clarsimp simp add:dvd_def) | |
| 366 | apply(rename_tac m) | |
| 367 | apply(rule_tac x = "m + k" in exI) | |
| 368 | apply(simp add:int_distrib) | |
| 369 | done | |
| 370 | ||
| 371 | lemma not_ast_p_ndvd: "(d1::int) dvd d ==> | |
| 372 |  ALL x. Q(x::int) \<and> ~(EX (j::int) : {1..d}. EX (a::int) : A. Q(a - j)) --> ~(d1 dvd (x + t)) --> ~(d1 dvd ((x + d) + t ))"
 | |
| 373 | apply(clarsimp simp add:dvd_def) | |
| 374 | apply(rename_tac m) | |
| 375 | apply(erule_tac x = "m - k" in allE) | |
| 376 | apply(simp add:int_distrib) | |
| 377 | done | |
| 378 | ||
| 14577 | 379 | text {*
 | 
| 380 | \medskip These are the atomic cases for the proof generation for the | |
| 381 |   modulo @{text D} property for @{text "P plusinfinity"}
 | |
| 13876 | 382 | |
| 14577 | 383 | They are fully based on arithmetics. *} | 
| 13876 | 384 | |
| 385 | lemma dvd_modd_pinf: "((d::int) dvd d1) ==> | |
| 386 | (ALL (x::int). ALL (k::int). (((d::int) dvd (x + t)) = (d dvd (x+k*d1 + t))))" | |
| 387 | apply(clarsimp simp add:dvd_def) | |
| 388 | apply(rule iffI) | |
| 389 | apply(clarsimp) | |
| 390 | apply(rename_tac n m) | |
| 391 | apply(rule_tac x = "m + n*k" in exI) | |
| 392 | apply(simp add:int_distrib) | |
| 393 | apply(clarsimp) | |
| 394 | apply(rename_tac n m) | |
| 395 | apply(rule_tac x = "m - n*k" in exI) | |
| 14271 | 396 | apply(simp add:int_distrib mult_ac) | 
| 13876 | 397 | done | 
| 398 | ||
| 399 | lemma not_dvd_modd_pinf: "((d::int) dvd d1) ==> | |
| 400 | (ALL (x::int). ALL k. (~((d::int) dvd (x + t))) = (~(d dvd (x+k*d1 + t))))" | |
| 401 | apply(clarsimp simp add:dvd_def) | |
| 402 | apply(rule iffI) | |
| 403 | apply(clarsimp) | |
| 404 | apply(rename_tac n m) | |
| 405 | apply(erule_tac x = "m - n*k" in allE) | |
| 14271 | 406 | apply(simp add:int_distrib mult_ac) | 
| 13876 | 407 | apply(clarsimp) | 
| 408 | apply(rename_tac n m) | |
| 409 | apply(erule_tac x = "m + n*k" in allE) | |
| 14271 | 410 | apply(simp add:int_distrib mult_ac) | 
| 13876 | 411 | done | 
| 412 | ||
| 14577 | 413 | text {*
 | 
| 414 | \medskip These are the atomic cases for the proof generation for the | |
| 415 |   equivalence of @{text P} and @{text "P plusinfinity"} for integers
 | |
| 416 |   @{text x} greater than some integer @{text z}.
 | |
| 417 | ||
| 418 | They are fully based on arithmetics. *} | |
| 13876 | 419 | |
| 420 | lemma eq_eq_pinf: "EX z::int. ALL x. z < x --> (( 0 = x +t ) = False )" | |
| 421 | apply(rule_tac x = "-t" in exI) | |
| 422 | apply simp | |
| 423 | done | |
| 424 | ||
| 425 | lemma neq_eq_pinf: "EX z::int. ALL x. z < x --> ((~( 0 = x +t )) = True )" | |
| 426 | apply(rule_tac x = "-t" in exI) | |
| 427 | apply simp | |
| 428 | done | |
| 429 | ||
| 430 | lemma le_eq_pinf: "EX z::int. ALL x. z < x --> ( 0 < x +t = True )" | |
| 431 | apply(rule_tac x = "-t" in exI) | |
| 432 | apply simp | |
| 433 | done | |
| 434 | ||
| 435 | lemma len_eq_pinf: "EX z::int. ALL x. z < x --> (0 < -x +t = False )" | |
| 436 | apply(rule_tac x = "t" in exI) | |
| 437 | apply simp | |
| 438 | done | |
| 439 | ||
| 440 | lemma dvd_eq_pinf: "EX z::int. ALL x. z < x --> ((d dvd (x + t)) = (d dvd (x + t))) " | |
| 14577 | 441 | by simp | 
| 13876 | 442 | |
| 443 | lemma not_dvd_eq_pinf: "EX z::int. ALL x. z < x --> ((~(d dvd (x + t))) = (~(d dvd (x + t)))) " | |
| 14577 | 444 | by simp | 
| 13876 | 445 | |
| 14577 | 446 | text {*
 | 
| 447 | \medskip These are the atomic cases for the proof generation for the | |
| 448 |   modulo @{text D} property for @{text "P minusinfinity"}.
 | |
| 449 | ||
| 450 | They are fully based on arithmetics. *} | |
| 13876 | 451 | |
| 452 | lemma dvd_modd_minf: "((d::int) dvd d1) ==> | |
| 453 | (ALL (x::int). ALL (k::int). (((d::int) dvd (x + t)) = (d dvd (x-k*d1 + t))))" | |
| 454 | apply(clarsimp simp add:dvd_def) | |
| 455 | apply(rule iffI) | |
| 456 | apply(clarsimp) | |
| 457 | apply(rename_tac n m) | |
| 458 | apply(rule_tac x = "m - n*k" in exI) | |
| 459 | apply(simp add:int_distrib) | |
| 460 | apply(clarsimp) | |
| 461 | apply(rename_tac n m) | |
| 462 | apply(rule_tac x = "m + n*k" in exI) | |
| 14271 | 463 | apply(simp add:int_distrib mult_ac) | 
| 13876 | 464 | done | 
| 465 | ||
| 466 | ||
| 467 | lemma not_dvd_modd_minf: "((d::int) dvd d1) ==> | |
| 468 | (ALL (x::int). ALL k. (~((d::int) dvd (x + t))) = (~(d dvd (x-k*d1 + t))))" | |
| 469 | apply(clarsimp simp add:dvd_def) | |
| 470 | apply(rule iffI) | |
| 471 | apply(clarsimp) | |
| 472 | apply(rename_tac n m) | |
| 473 | apply(erule_tac x = "m + n*k" in allE) | |
| 14271 | 474 | apply(simp add:int_distrib mult_ac) | 
| 13876 | 475 | apply(clarsimp) | 
| 476 | apply(rename_tac n m) | |
| 477 | apply(erule_tac x = "m - n*k" in allE) | |
| 14271 | 478 | apply(simp add:int_distrib mult_ac) | 
| 13876 | 479 | done | 
| 480 | ||
| 14577 | 481 | text {*
 | 
| 482 | \medskip These are the atomic cases for the proof generation for the | |
| 483 |   equivalence of @{text P} and @{text "P minusinfinity"} for integers
 | |
| 484 |   @{text x} less than some integer @{text z}.
 | |
| 13876 | 485 | |
| 14577 | 486 | They are fully based on arithmetics. *} | 
| 13876 | 487 | |
| 488 | lemma eq_eq_minf: "EX z::int. ALL x. x < z --> (( 0 = x +t ) = False )" | |
| 489 | apply(rule_tac x = "-t" in exI) | |
| 490 | apply simp | |
| 491 | done | |
| 492 | ||
| 493 | lemma neq_eq_minf: "EX z::int. ALL x. x < z --> ((~( 0 = x +t )) = True )" | |
| 494 | apply(rule_tac x = "-t" in exI) | |
| 495 | apply simp | |
| 496 | done | |
| 497 | ||
| 498 | lemma le_eq_minf: "EX z::int. ALL x. x < z --> ( 0 < x +t = False )" | |
| 499 | apply(rule_tac x = "-t" in exI) | |
| 500 | apply simp | |
| 501 | done | |
| 502 | ||
| 503 | ||
| 504 | lemma len_eq_minf: "EX z::int. ALL x. x < z --> (0 < -x +t = True )" | |
| 505 | apply(rule_tac x = "t" in exI) | |
| 506 | apply simp | |
| 507 | done | |
| 508 | ||
| 509 | lemma dvd_eq_minf: "EX z::int. ALL x. x < z --> ((d dvd (x + t)) = (d dvd (x + t))) " | |
| 14577 | 510 | by simp | 
| 13876 | 511 | |
| 512 | lemma not_dvd_eq_minf: "EX z::int. ALL x. x < z --> ((~(d dvd (x + t))) = (~(d dvd (x + t)))) " | |
| 14577 | 513 | by simp | 
| 13876 | 514 | |
| 14577 | 515 | text {*
 | 
| 516 |   \medskip This Theorem combines whithnesses about @{text "P
 | |
| 517 | minusinfinity"} to show one component of the equivalence proof for | |
| 518 | Cooper's Theorem. | |
| 13876 | 519 | |
| 14577 | 520 | FIXME: remove once they are part of the distribution. *} | 
| 521 | ||
| 13876 | 522 | theorem int_ge_induct[consumes 1,case_names base step]: | 
| 523 | assumes ge: "k \<le> (i::int)" and | |
| 524 | base: "P(k)" and | |
| 525 | step: "\<And>i. \<lbrakk>k \<le> i; P i\<rbrakk> \<Longrightarrow> P(i+1)" | |
| 526 | shows "P i" | |
| 527 | proof - | |
| 528 |   { fix n have "\<And>i::int. n = nat(i-k) \<Longrightarrow> k <= i \<Longrightarrow> P i"
 | |
| 529 | proof (induct n) | |
| 530 | case 0 | |
| 531 | hence "i = k" by arith | |
| 532 | thus "P i" using base by simp | |
| 533 | next | |
| 534 | case (Suc n) | |
| 535 | hence "n = nat((i - 1) - k)" by arith | |
| 536 | moreover | |
| 537 | have ki1: "k \<le> i - 1" using Suc.prems by arith | |
| 538 | ultimately | |
| 539 | have "P(i - 1)" by(rule Suc.hyps) | |
| 540 | from step[OF ki1 this] show ?case by simp | |
| 541 | qed | |
| 542 | } | |
| 543 | from this ge show ?thesis by fast | |
| 544 | qed | |
| 545 | ||
| 546 | theorem int_gr_induct[consumes 1,case_names base step]: | |
| 547 | assumes gr: "k < (i::int)" and | |
| 548 | base: "P(k+1)" and | |
| 549 | step: "\<And>i. \<lbrakk>k < i; P i\<rbrakk> \<Longrightarrow> P(i+1)" | |
| 550 | shows "P i" | |
| 551 | apply(rule int_ge_induct[of "k + 1"]) | |
| 552 | using gr apply arith | |
| 553 | apply(rule base) | |
| 554 | apply(rule step) | |
| 555 | apply simp+ | |
| 556 | done | |
| 557 | ||
| 558 | lemma decr_lemma: "0 < (d::int) \<Longrightarrow> x - (abs(x-z)+1) * d < z" | |
| 559 | apply(induct rule: int_gr_induct) | |
| 560 | apply simp | |
| 561 | apply arith | |
| 562 | apply (simp add:int_distrib) | |
| 563 | apply arith | |
| 564 | done | |
| 565 | ||
| 566 | lemma incr_lemma: "0 < (d::int) \<Longrightarrow> z < x + (abs(x-z)+1) * d" | |
| 567 | apply(induct rule: int_gr_induct) | |
| 568 | apply simp | |
| 569 | apply arith | |
| 570 | apply (simp add:int_distrib) | |
| 571 | apply arith | |
| 572 | done | |
| 573 | ||
| 574 | lemma minusinfinity: | |
| 575 | assumes "0 < d" and | |
| 576 | P1eqP1: "ALL x k. P1 x = P1(x - k*d)" and | |
| 577 | ePeqP1: "EX z::int. ALL x. x < z \<longrightarrow> (P x = P1 x)" | |
| 578 | shows "(EX x. P1 x) \<longrightarrow> (EX x. P x)" | |
| 579 | proof | |
| 580 | assume eP1: "EX x. P1 x" | |
| 581 | then obtain x where P1: "P1 x" .. | |
| 582 | from ePeqP1 obtain z where P1eqP: "ALL x. x < z \<longrightarrow> (P x = P1 x)" .. | |
| 583 | let ?w = "x - (abs(x-z)+1) * d" | |
| 584 | show "EX x. P x" | |
| 585 | proof | |
| 586 | have w: "?w < z" by(rule decr_lemma) | |
| 587 | have "P1 x = P1 ?w" using P1eqP1 by blast | |
| 588 | also have "\<dots> = P(?w)" using w P1eqP by blast | |
| 589 | finally show "P ?w" using P1 by blast | |
| 590 | qed | |
| 591 | qed | |
| 592 | ||
| 14577 | 593 | text {*
 | 
| 594 |   \medskip This Theorem combines whithnesses about @{text "P
 | |
| 595 | minusinfinity"} to show one component of the equivalence proof for | |
| 596 | Cooper's Theorem. *} | |
| 13876 | 597 | |
| 598 | lemma plusinfinity: | |
| 599 | assumes "0 < d" and | |
| 600 | P1eqP1: "ALL (x::int) (k::int). P1 x = P1 (x + k * d)" and | |
| 601 | ePeqP1: "EX z::int. ALL x. z < x --> (P x = P1 x)" | |
| 602 | shows "(EX x::int. P1 x) --> (EX x::int. P x)" | |
| 603 | proof | |
| 604 | assume eP1: "EX x. P1 x" | |
| 605 | then obtain x where P1: "P1 x" .. | |
| 606 | from ePeqP1 obtain z where P1eqP: "ALL x. z < x \<longrightarrow> (P x = P1 x)" .. | |
| 607 | let ?w = "x + (abs(x-z)+1) * d" | |
| 608 | show "EX x. P x" | |
| 609 | proof | |
| 610 | have w: "z < ?w" by(rule incr_lemma) | |
| 611 | have "P1 x = P1 ?w" using P1eqP1 by blast | |
| 612 | also have "\<dots> = P(?w)" using w P1eqP by blast | |
| 613 | finally show "P ?w" using P1 by blast | |
| 614 | qed | |
| 615 | qed | |
| 616 | ||
| 14577 | 617 | text {*
 | 
| 618 | \medskip Theorem for periodic function on discrete sets. *} | |
| 13876 | 619 | |
| 620 | lemma minf_vee: | |
| 621 | assumes dpos: "(0::int) < d" and modd: "ALL x k. P x = P(x - k*d)" | |
| 622 |   shows "(EX x. P x) = (EX j : {1..d}. P j)"
 | |
| 623 | (is "?LHS = ?RHS") | |
| 624 | proof | |
| 625 | assume ?LHS | |
| 626 | then obtain x where P: "P x" .. | |
| 627 | have "x mod d = x - (x div d)*d" | |
| 14271 | 628 | by(simp add:zmod_zdiv_equality mult_ac eq_diff_eq) | 
| 13876 | 629 | hence Pmod: "P x = P(x mod d)" using modd by simp | 
| 630 | show ?RHS | |
| 631 | proof (cases) | |
| 632 | assume "x mod d = 0" | |
| 633 | hence "P 0" using P Pmod by simp | |
| 634 | moreover have "P 0 = P(0 - (-1)*d)" using modd by blast | |
| 635 | ultimately have "P d" by simp | |
| 636 |     moreover have "d : {1..d}" using dpos by(simp add:atLeastAtMost_iff)
 | |
| 637 | ultimately show ?RHS .. | |
| 638 | next | |
| 639 | assume not0: "x mod d \<noteq> 0" | |
| 640 | have "P(x mod d)" using dpos P Pmod by(simp add:pos_mod_sign pos_mod_bound) | |
| 641 |     moreover have "x mod d : {1..d}"
 | |
| 642 | proof - | |
| 643 | have "0 \<le> x mod d" by(rule pos_mod_sign) | |
| 644 | moreover have "x mod d < d" by(rule pos_mod_bound) | |
| 645 | ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff) | |
| 646 | qed | |
| 647 | ultimately show ?RHS .. | |
| 648 | qed | |
| 649 | next | |
| 650 | assume ?RHS thus ?LHS by blast | |
| 651 | qed | |
| 652 | ||
| 14577 | 653 | text {*
 | 
| 654 | \medskip Theorem for periodic function on discrete sets. *} | |
| 655 | ||
| 13876 | 656 | lemma pinf_vee: | 
| 657 | assumes dpos: "0 < (d::int)" and modd: "ALL (x::int) (k::int). P x = P (x+k*d)" | |
| 658 |   shows "(EX x::int. P x) = (EX (j::int) : {1..d} . P j)"
 | |
| 659 | (is "?LHS = ?RHS") | |
| 660 | proof | |
| 661 | assume ?LHS | |
| 662 | then obtain x where P: "P x" .. | |
| 663 | have "x mod d = x + (-(x div d))*d" | |
| 14271 | 664 | by(simp add:zmod_zdiv_equality mult_ac eq_diff_eq) | 
| 13876 | 665 | hence Pmod: "P x = P(x mod d)" using modd by (simp only:) | 
| 666 | show ?RHS | |
| 667 | proof (cases) | |
| 668 | assume "x mod d = 0" | |
| 669 | hence "P 0" using P Pmod by simp | |
| 670 | moreover have "P 0 = P(0 + 1*d)" using modd by blast | |
| 671 | ultimately have "P d" by simp | |
| 672 |     moreover have "d : {1..d}" using dpos by(simp add:atLeastAtMost_iff)
 | |
| 673 | ultimately show ?RHS .. | |
| 674 | next | |
| 675 | assume not0: "x mod d \<noteq> 0" | |
| 676 | have "P(x mod d)" using dpos P Pmod by(simp add:pos_mod_sign pos_mod_bound) | |
| 677 |     moreover have "x mod d : {1..d}"
 | |
| 678 | proof - | |
| 679 | have "0 \<le> x mod d" by(rule pos_mod_sign) | |
| 680 | moreover have "x mod d < d" by(rule pos_mod_bound) | |
| 681 | ultimately show ?thesis using not0 by(simp add:atLeastAtMost_iff) | |
| 682 | qed | |
| 683 | ultimately show ?RHS .. | |
| 684 | qed | |
| 685 | next | |
| 686 | assume ?RHS thus ?LHS by blast | |
| 687 | qed | |
| 688 | ||
| 689 | lemma decr_mult_lemma: | |
| 690 | assumes dpos: "(0::int) < d" and | |
| 691 | minus: "ALL x::int. P x \<longrightarrow> P(x - d)" and | |
| 692 | knneg: "0 <= k" | |
| 693 | shows "ALL x. P x \<longrightarrow> P(x - k*d)" | |
| 694 | using knneg | |
| 695 | proof (induct rule:int_ge_induct) | |
| 696 | case base thus ?case by simp | |
| 697 | next | |
| 698 | case (step i) | |
| 699 | show ?case | |
| 700 | proof | |
| 701 | fix x | |
| 702 | have "P x \<longrightarrow> P (x - i * d)" using step.hyps by blast | |
| 703 | also have "\<dots> \<longrightarrow> P(x - (i + 1) * d)" | |
| 704 | using minus[THEN spec, of "x - i * d"] | |
| 14738 | 705 | by (simp add:int_distrib OrderedGroup.diff_diff_eq[symmetric]) | 
| 13876 | 706 | ultimately show "P x \<longrightarrow> P(x - (i + 1) * d)" by blast | 
| 707 | qed | |
| 708 | qed | |
| 709 | ||
| 710 | lemma incr_mult_lemma: | |
| 711 | assumes dpos: "(0::int) < d" and | |
| 712 | plus: "ALL x::int. P x \<longrightarrow> P(x + d)" and | |
| 713 | knneg: "0 <= k" | |
| 714 | shows "ALL x. P x \<longrightarrow> P(x + k*d)" | |
| 715 | using knneg | |
| 716 | proof (induct rule:int_ge_induct) | |
| 717 | case base thus ?case by simp | |
| 718 | next | |
| 719 | case (step i) | |
| 720 | show ?case | |
| 721 | proof | |
| 722 | fix x | |
| 723 | have "P x \<longrightarrow> P (x + i * d)" using step.hyps by blast | |
| 724 | also have "\<dots> \<longrightarrow> P(x + (i + 1) * d)" | |
| 725 | using plus[THEN spec, of "x + i * d"] | |
| 726 | by (simp add:int_distrib zadd_ac) | |
| 727 | ultimately show "P x \<longrightarrow> P(x + (i + 1) * d)" by blast | |
| 728 | qed | |
| 729 | qed | |
| 730 | ||
| 731 | lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x)) | |
| 732 | ==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) 
 | |
| 733 | ==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D)))) | |
| 734 | ==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))"
 | |
| 735 | apply(rule iffI) | |
| 736 | prefer 2 | |
| 737 | apply(drule minusinfinity) | |
| 738 | apply assumption+ | |
| 739 | apply(fastsimp) | |
| 740 | apply clarsimp | |
| 741 | apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)") | |
| 742 | apply(frule_tac x = x and z=z in decr_lemma) | |
| 743 | apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)") | |
| 744 | prefer 2 | |
| 745 | apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)") | |
| 746 | prefer 2 apply arith | |
| 747 | apply fastsimp | |
| 748 | apply(drule (1) minf_vee) | |
| 749 | apply blast | |
| 750 | apply(blast dest:decr_mult_lemma) | |
| 751 | done | |
| 752 | ||
| 14577 | 753 | text {* Cooper Theorem, plus infinity version. *}
 | 
| 13876 | 754 | lemma cppi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. z < x --> (P x = P1 x)) | 
| 755 | ==> ALL x.~(EX (j::int) : {1..D}. EX (a::int) : A. P(a - j)) --> P (x) --> P (x + D) 
 | |
| 756 | ==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x+k*D)))) | |
| 757 | ==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (a::int) : A. P (a - j)))"
 | |
| 758 | apply(rule iffI) | |
| 759 | prefer 2 | |
| 760 | apply(drule plusinfinity) | |
| 761 | apply assumption+ | |
| 762 | apply(fastsimp) | |
| 763 | apply clarsimp | |
| 764 | apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x + k*D)") | |
| 765 | apply(frule_tac x = x and z=z in incr_lemma) | |
| 766 | apply(subgoal_tac "P1(x + (\<bar>x - z\<bar> + 1) * D)") | |
| 767 | prefer 2 | |
| 768 | apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)") | |
| 769 | prefer 2 apply arith | |
| 770 | apply fastsimp | |
| 771 | apply(drule (1) pinf_vee) | |
| 772 | apply blast | |
| 773 | apply(blast dest:incr_mult_lemma) | |
| 774 | done | |
| 775 | ||
| 776 | ||
| 14577 | 777 | text {*
 | 
| 778 | \bigskip Theorems for the quantifier elminination Functions. *} | |
| 13876 | 779 | |
| 780 | lemma qe_ex_conj: "(EX (x::int). A x) = R | |
| 781 | ==> (EX (x::int). P x) = (Q & (EX x::int. A x)) | |
| 782 | ==> (EX (x::int). P x) = (Q & R)" | |
| 783 | by blast | |
| 784 | ||
| 785 | lemma qe_ex_nconj: "(EX (x::int). P x) = (True & Q) | |
| 786 | ==> (EX (x::int). P x) = Q" | |
| 787 | by blast | |
| 788 | ||
| 789 | lemma qe_conjI: "P1 = P2 ==> Q1 = Q2 ==> (P1 & Q1) = (P2 & Q2)" | |
| 790 | by blast | |
| 791 | ||
| 792 | lemma qe_disjI: "P1 = P2 ==> Q1 = Q2 ==> (P1 | Q1) = (P2 | Q2)" | |
| 793 | by blast | |
| 794 | ||
| 795 | lemma qe_impI: "P1 = P2 ==> Q1 = Q2 ==> (P1 --> Q1) = (P2 --> Q2)" | |
| 796 | by blast | |
| 797 | ||
| 798 | lemma qe_eqI: "P1 = P2 ==> Q1 = Q2 ==> (P1 = Q1) = (P2 = Q2)" | |
| 799 | by blast | |
| 800 | ||
| 801 | lemma qe_Not: "P = Q ==> (~P) = (~Q)" | |
| 802 | by blast | |
| 803 | ||
| 804 | lemma qe_ALL: "(EX x. ~P x) = R ==> (ALL x. P x) = (~R)" | |
| 805 | by blast | |
| 806 | ||
| 14577 | 807 | text {* \bigskip Theorems for proving NNF *}
 | 
| 13876 | 808 | |
| 809 | lemma nnf_im: "((~P) = P1) ==> (Q=Q1) ==> ((P --> Q) = (P1 | Q1))" | |
| 810 | by blast | |
| 811 | ||
| 812 | lemma nnf_eq: "((P & Q) = (P1 & Q1)) ==> (((~P) & (~Q)) = (P2 & Q2)) ==> ((P = Q) = ((P1 & Q1)|(P2 & Q2)))" | |
| 813 | by blast | |
| 814 | ||
| 815 | lemma nnf_nn: "(P = Q) ==> ((~~P) = Q)" | |
| 816 | by blast | |
| 817 | lemma nnf_ncj: "((~P) = P1) ==> ((~Q) = Q1) ==> ((~(P & Q)) = (P1 | Q1))" | |
| 818 | by blast | |
| 819 | ||
| 820 | lemma nnf_ndj: "((~P) = P1) ==> ((~Q) = Q1) ==> ((~(P | Q)) = (P1 & Q1))" | |
| 821 | by blast | |
| 822 | lemma nnf_nim: "(P = P1) ==> ((~Q) = Q1) ==> ((~(P --> Q)) = (P1 & Q1))" | |
| 823 | by blast | |
| 824 | lemma nnf_neq: "((P & (~Q)) = (P1 & Q1)) ==> (((~P) & Q) = (P2 & Q2)) ==> ((~(P = Q)) = ((P1 & Q1)|(P2 & Q2)))" | |
| 825 | by blast | |
| 826 | lemma nnf_sdj: "((A & (~B)) = (A1 & B1)) ==> ((C & (~D)) = (C1 & D1)) ==> (A = (~C)) ==> ((~((A & B) | (C & D))) = ((A1 & B1) | (C1 & D1)))" | |
| 827 | by blast | |
| 828 | ||
| 829 | ||
| 830 | lemma qe_exI2: "A = B ==> (EX (x::int). A(x)) = (EX (x::int). B(x))" | |
| 831 | by simp | |
| 832 | ||
| 833 | lemma qe_exI: "(!!x::int. A x = B x) ==> (EX (x::int). A(x)) = (EX (x::int). B(x))" | |
| 17589 | 834 | by iprover | 
| 13876 | 835 | |
| 836 | lemma qe_ALLI: "(!!x::int. A x = B x) ==> (ALL (x::int). A(x)) = (ALL (x::int). B(x))" | |
| 17589 | 837 | by iprover | 
| 13876 | 838 | |
| 839 | lemma cp_expand: "(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (b::int) : B. (P1 (j) | P(b+j)))
 | |
| 840 | ==>(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (b::int) : B. (P1 (j) | P(b+j))) "
 | |
| 841 | by blast | |
| 842 | ||
| 843 | lemma cppi_expand: "(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (a::int) : A. (P1 (j) | P(a - j)))
 | |
| 844 | ==>(EX (x::int). P (x)) = (EX (j::int) : {1..d}. EX (a::int) : A. (P1 (j) | P(a - j))) "
 | |
| 845 | by blast | |
| 846 | ||
| 847 | ||
| 848 | lemma simp_from_to: "{i..j::int} = (if j < i then {} else insert i {i+1..j})"
 | |
| 849 | apply(simp add:atLeastAtMost_def atLeast_def atMost_def) | |
| 850 | apply(fastsimp) | |
| 851 | done | |
| 852 | ||
| 14577 | 853 | text {* \bigskip Theorems required for the @{text adjustcoeffitienteq} *}
 | 
| 13876 | 854 | |
| 855 | lemma ac_dvd_eq: assumes not0: "0 ~= (k::int)" | |
| 856 | shows "((m::int) dvd (c*n+t)) = (k*m dvd ((k*c)*n+(k*t)))" (is "?P = ?Q") | |
| 857 | proof | |
| 858 | assume ?P | |
| 859 | thus ?Q | |
| 860 | apply(simp add:dvd_def) | |
| 861 | apply clarify | |
| 862 | apply(rename_tac d) | |
| 863 | apply(drule_tac f = "op * k" in arg_cong) | |
| 864 | apply(simp only:int_distrib) | |
| 865 | apply(rule_tac x = "d" in exI) | |
| 14271 | 866 | apply(simp only:mult_ac) | 
| 13876 | 867 | done | 
| 868 | next | |
| 869 | assume ?Q | |
| 870 | then obtain d where "k * c * n + k * t = (k*m)*d" by(fastsimp simp:dvd_def) | |
| 14271 | 871 | hence "(c * n + t) * k = (m*d) * k" by(simp add:int_distrib mult_ac) | 
| 13876 | 872 | hence "((c * n + t) * k) div k = ((m*d) * k) div k" by(rule arg_cong[of _ _ "%t. t div k"]) | 
| 873 | hence "c*n+t = m*d" by(simp add: zdiv_zmult_self1[OF not0[symmetric]]) | |
| 874 | thus ?P by(simp add:dvd_def) | |
| 875 | qed | |
| 876 | ||
| 877 | lemma ac_lt_eq: assumes gr0: "0 < (k::int)" | |
| 878 | shows "((m::int) < (c*n+t)) = (k*m <((k*c)*n+(k*t)))" (is "?P = ?Q") | |
| 879 | proof | |
| 880 | assume P: ?P | |
| 14271 | 881 | show ?Q using zmult_zless_mono2[OF P gr0] by(simp add: int_distrib mult_ac) | 
| 13876 | 882 | next | 
| 883 | assume ?Q | |
| 14271 | 884 | hence "0 < k*(c*n + t - m)" by(simp add: int_distrib mult_ac) | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14271diff
changeset | 885 | with gr0 have "0 < (c*n + t - m)" by(simp add: zero_less_mult_iff) | 
| 13876 | 886 | thus ?P by(simp) | 
| 887 | qed | |
| 888 | ||
| 889 | lemma ac_eq_eq : assumes not0: "0 ~= (k::int)" shows "((m::int) = (c*n+t)) = (k*m =((k*c)*n+(k*t)) )" (is "?P = ?Q") | |
| 890 | proof | |
| 891 | assume ?P | |
| 892 | thus ?Q | |
| 893 | apply(drule_tac f = "op * k" in arg_cong) | |
| 894 | apply(simp only:int_distrib) | |
| 895 | done | |
| 896 | next | |
| 897 | assume ?Q | |
| 14271 | 898 | hence "m * k = (c*n + t) * k" by(simp add:int_distrib mult_ac) | 
| 13876 | 899 | hence "((m) * k) div k = ((c*n + t) * k) div k" by(rule arg_cong[of _ _ "%t. t div k"]) | 
| 900 | thus ?P by(simp add: zdiv_zmult_self1[OF not0[symmetric]]) | |
| 901 | qed | |
| 902 | ||
| 903 | lemma ac_pi_eq: assumes gr0: "0 < (k::int)" shows "(~((0::int) < (c*n + t))) = (0 < ((-k)*c)*n + ((-k)*t + k))" | |
| 904 | proof - | |
| 905 | have "(~ (0::int) < (c*n + t)) = (0<1-(c*n + t))" by arith | |
| 14271 | 906 | also have "(1-(c*n + t)) = (-1*c)*n + (-t+1)" by(simp add: int_distrib mult_ac) | 
| 13876 | 907 | also have "0<(-1*c)*n + (-t+1) = (0 < (k*(-1*c)*n) + (k*(-t+1)))" by(rule ac_lt_eq[of _ 0,OF gr0,simplified]) | 
| 14271 | 908 | also have "(k*(-1*c)*n) + (k*(-t+1)) = ((-k)*c)*n + ((-k)*t + k)" by(simp add: int_distrib mult_ac) | 
| 13876 | 909 | finally show ?thesis . | 
| 910 | qed | |
| 911 | ||
| 912 | lemma binminus_uminus_conv: "(a::int) - b = a + (-b)" | |
| 913 | by arith | |
| 914 | ||
| 915 | lemma linearize_dvd: "(t::int) = t1 ==> (d dvd t) = (d dvd t1)" | |
| 916 | by simp | |
| 917 | ||
| 918 | lemma lf_lt: "(l::int) = ll ==> (r::int) = lr ==> (l < r) =(ll < lr)" | |
| 919 | by simp | |
| 920 | ||
| 921 | lemma lf_eq: "(l::int) = ll ==> (r::int) = lr ==> (l = r) =(ll = lr)" | |
| 922 | by simp | |
| 923 | ||
| 924 | lemma lf_dvd: "(l::int) = ll ==> (r::int) = lr ==> (l dvd r) =(ll dvd lr)" | |
| 925 | by simp | |
| 926 | ||
| 14577 | 927 | text {* \bigskip Theorems for transforming predicates on nat to predicates on @{text int}*}
 | 
| 13876 | 928 | |
| 929 | theorem all_nat: "(\<forall>x::nat. P x) = (\<forall>x::int. 0 <= x \<longrightarrow> P (nat x))" | |
| 930 | by (simp split add: split_nat) | |
| 931 | ||
| 932 | theorem ex_nat: "(\<exists>x::nat. P x) = (\<exists>x::int. 0 <= x \<and> P (nat x))" | |
| 933 | apply (simp split add: split_nat) | |
| 934 | apply (rule iffI) | |
| 935 | apply (erule exE) | |
| 936 | apply (rule_tac x = "int x" in exI) | |
| 937 | apply simp | |
| 938 | apply (erule exE) | |
| 939 | apply (rule_tac x = "nat x" in exI) | |
| 940 | apply (erule conjE) | |
| 941 | apply (erule_tac x = "nat x" in allE) | |
| 942 | apply simp | |
| 943 | done | |
| 944 | ||
| 945 | theorem zdiff_int_split: "P (int (x - y)) = | |
| 946 | ((y \<le> x \<longrightarrow> P (int x - int y)) \<and> (x < y \<longrightarrow> P 0))" | |
| 947 | apply (case_tac "y \<le> x") | |
| 948 | apply (simp_all add: zdiff_int) | |
| 949 | done | |
| 950 | ||
| 951 | theorem zdvd_int: "(x dvd y) = (int x dvd int y)" | |
| 952 | apply (simp only: dvd_def ex_nat int_int_eq [symmetric] zmult_int [symmetric] | |
| 953 | nat_0_le cong add: conj_cong) | |
| 954 | apply (rule iffI) | |
| 17589 | 955 | apply iprover | 
| 13876 | 956 | apply (erule exE) | 
| 957 | apply (case_tac "x=0") | |
| 958 | apply (rule_tac x=0 in exI) | |
| 959 | apply simp | |
| 960 | apply (case_tac "0 \<le> k") | |
| 17589 | 961 | apply iprover | 
| 13876 | 962 | apply (simp add: linorder_not_le) | 
| 14378 
69c4d5997669
generic of_nat and of_int functions, and generalization of iszero
 paulson parents: 
14353diff
changeset | 963 | apply (drule mult_strict_left_mono_neg [OF iffD2 [OF zero_less_int_conv]]) | 
| 13876 | 964 | apply assumption | 
| 14271 | 965 | apply (simp add: mult_ac) | 
| 13876 | 966 | done | 
| 967 | ||
| 968 | theorem number_of1: "(0::int) <= number_of n \<Longrightarrow> (0::int) <= number_of (n BIT b)" | |
| 969 | by simp | |
| 970 | ||
| 15013 | 971 | theorem number_of2: "(0::int) <= Numeral0" by simp | 
| 13876 | 972 | |
| 973 | theorem Suc_plus1: "Suc n = n + 1" by simp | |
| 974 | ||
| 14577 | 975 | text {*
 | 
| 976 | \medskip Specific instances of congruence rules, to prevent | |
| 977 | simplifier from looping. *} | |
| 13876 | 978 | |
| 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: 
14738diff
changeset | 979 | theorem imp_le_cong: "(0 <= x \<Longrightarrow> P = P') \<Longrightarrow> (0 <= (x::int) \<longrightarrow> P) = (0 <= x \<longrightarrow> P')" | 
| 13876 | 980 | by simp | 
| 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: 
14738diff
changeset | 982 | theorem conj_le_cong: "(0 <= x \<Longrightarrow> P = P') \<Longrightarrow> (0 <= (x::int) \<and> P) = (0 <= x \<and> 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: 
14738diff
changeset | 983 | by (simp cong: conj_cong) | 
| 13876 | 984 | |
| 18202 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 985 | (* Theorems used in presburger.ML for the computation simpset*) | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 986 | (* FIXME: They are present in Float.thy, so may be Float.thy should be lightened.*) | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 987 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 988 | lemma lift_bool: "x \<Longrightarrow> x=True" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 989 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 990 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 991 | lemma nlift_bool: "~x \<Longrightarrow> x=False" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 992 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 993 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 994 | lemma not_false_eq_true: "(~ False) = True" by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 995 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 996 | lemma not_true_eq_false: "(~ True) = False" by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 997 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 998 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 999 | lemma int_eq_number_of_eq: "(((number_of v)::int)=(number_of w)) = iszero ((number_of (bin_add v (bin_minus w)))::int)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1000 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1001 | lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1002 | by (simp only: iszero_number_of_Pls) | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1003 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1004 | lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1005 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1006 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1007 | lemma int_iszero_number_of_0: "iszero ((number_of (w BIT bit.B0))::int) = iszero ((number_of w)::int)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1008 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1009 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1010 | lemma int_iszero_number_of_1: "\<not> iszero ((number_of (w BIT bit.B1))::int)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1011 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1012 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1013 | lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (bin_add x (bin_minus y)))::int)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1014 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1015 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1016 | lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1017 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1018 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1019 | lemma int_neg_number_of_Min: "neg (-1::int)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1020 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1021 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1022 | lemma int_neg_number_of_BIT: "neg ((number_of (w BIT x))::int) = neg ((number_of w)::int)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1023 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1024 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1025 | lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (bin_add y (bin_minus x)))::int))" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1026 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1027 | lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (bin_add v w)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1028 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1029 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1030 | lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (bin_add v (bin_minus w))" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1031 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1032 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1033 | lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (bin_mult v w)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1034 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1035 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1036 | lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (bin_minus v)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1037 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1038 | lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1039 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1040 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1041 | lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1042 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1043 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1044 | lemma mult_left_one: "1 * a = (a::'a::semiring_1)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1045 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1046 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1047 | lemma mult_right_one: "a * 1 = (a::'a::semiring_1)" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1048 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1049 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1050 | lemma int_pow_0: "(a::int)^(Numeral0) = 1" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1051 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1052 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1053 | lemma int_pow_1: "(a::int)^(Numeral1) = a" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1054 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1055 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1056 | lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1057 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1058 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1059 | lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1060 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1061 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1062 | lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1063 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1064 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1065 | lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1066 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1067 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1068 | lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1069 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1070 | |
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1071 | lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1072 | proof - | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1073 | have 1:"((-1)::nat) = 0" | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1074 | by simp | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1075 | show ?thesis by (simp add: 1) | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1076 | qed | 
| 
46af82efd311
presburger method updated to deal better with mod and div, tweo lemmas added to Divides.thy
 chaieb parents: 
17589diff
changeset | 1077 | |
| 13876 | 1078 | use "cooper_dec.ML" | 
| 17378 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 chaieb parents: 
16836diff
changeset | 1079 | use "reflected_presburger.ML" | 
| 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 chaieb parents: 
16836diff
changeset | 1080 | use "reflected_cooper.ML" | 
| 14941 | 1081 | oracle | 
| 17378 
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
 chaieb parents: 
16836diff
changeset | 1082 |   presburger_oracle ("term") = ReflectedCooper.presburger_oracle
 | 
| 14941 | 1083 | |
| 13876 | 1084 | use "cooper_proof.ML" | 
| 1085 | use "qelim.ML" | |
| 1086 | use "presburger.ML" | |
| 1087 | ||
| 1088 | setup "Presburger.setup" | |
| 1089 | ||
| 1090 | end |