| author | webertj | 
| Wed, 03 Dec 2008 14:02:24 +0000 | |
| changeset 28949 | 610fe33ca358 | 
| parent 28290 | 4cc2b6046258 | 
| child 29265 | 5b4247055bd7 | 
| permissions | -rw-r--r-- | 
| 27456 | 1 | (* Title: HOL/ex/Reflected_Presburger.thy | 
| 2 | Author: Amine Chaieb | |
| 3 | *) | |
| 4 | ||
| 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: diff
changeset | 5 | theory Reflected_Presburger | 
| 25592 | 6 | imports Main GCD Efficient_Nat | 
| 27456 | 7 | uses ("coopertac.ML")
 | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 8 | begin | 
| 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: diff
changeset | 9 | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 10 | function | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 11 | iupt :: "int \<Rightarrow> int \<Rightarrow> int list" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 12 | where | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 13 | "iupt i j = (if j < i then [] else i # iupt (i+1) j)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 14 | by pat_completeness auto | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 15 | termination by (relation "measure (\<lambda> (i, j). nat (j-i+1))") auto | 
| 23274 | 16 | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 17 | lemma iupt_set: "set (iupt i j) = {i..j}"
 | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 18 | by (induct rule: iupt.induct) (simp add: simp_from_to) | 
| 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: diff
changeset | 19 | |
| 
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: diff
changeset | 20 | (* Periodicity of dvd *) | 
| 23315 | 21 | |
| 23274 | 22 | (*********************************************************************************) | 
| 23 | (**** SHADOW SYNTAX AND SEMANTICS ****) | |
| 24 | (*********************************************************************************) | |
| 25 | ||
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 26 | datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num | 
| 23274 | 27 | | Mul int num | 
| 28 | ||
| 29 | (* A size for num to make inductive proofs simpler*) | |
| 27456 | 30 | primrec num_size :: "num \<Rightarrow> nat" where | 
| 23274 | 31 | "num_size (C c) = 1" | 
| 27456 | 32 | | "num_size (Bound n) = 1" | 
| 33 | | "num_size (Neg a) = 1 + num_size a" | |
| 34 | | "num_size (Add a b) = 1 + num_size a + num_size b" | |
| 35 | | "num_size (Sub a b) = 3 + num_size a + num_size b" | |
| 36 | | "num_size (CN n c a) = 4 + num_size a" | |
| 37 | | "num_size (Mul c a) = 1 + num_size a" | |
| 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: diff
changeset | 38 | |
| 27456 | 39 | primrec Inum :: "int list \<Rightarrow> num \<Rightarrow> int" where | 
| 23274 | 40 | "Inum bs (C c) = c" | 
| 27456 | 41 | | "Inum bs (Bound n) = bs!n" | 
| 42 | | "Inum bs (CN n c a) = c * (bs!n) + (Inum bs a)" | |
| 43 | | "Inum bs (Neg a) = -(Inum bs a)" | |
| 44 | | "Inum bs (Add a b) = Inum bs a + Inum bs b" | |
| 45 | | "Inum bs (Sub a b) = Inum bs a - Inum bs b" | |
| 46 | | "Inum bs (Mul c a) = c* Inum bs a" | |
| 23274 | 47 | |
| 48 | datatype fm = | |
| 49 | T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num| | |
| 50 | NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm | |
| 51 | | Closed nat | NClosed nat | |
| 52 | ||
| 53 | ||
| 54 | (* A size for fm *) | |
| 55 | consts fmsize :: "fm \<Rightarrow> nat" | |
| 56 | recdef fmsize "measure size" | |
| 57 | "fmsize (NOT p) = 1 + fmsize p" | |
| 58 | "fmsize (And p q) = 1 + fmsize p + fmsize q" | |
| 59 | "fmsize (Or p q) = 1 + fmsize p + fmsize q" | |
| 60 | "fmsize (Imp p q) = 3 + fmsize p + fmsize q" | |
| 61 | "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)" | |
| 62 | "fmsize (E p) = 1 + fmsize p" | |
| 63 | "fmsize (A p) = 4+ fmsize p" | |
| 64 | "fmsize (Dvd i t) = 2" | |
| 65 | "fmsize (NDvd i t) = 2" | |
| 66 | "fmsize p = 1" | |
| 67 | (* several lemmas about fmsize *) | |
| 25162 | 68 | lemma fmsize_pos: "fmsize p > 0" | 
| 23274 | 69 | by (induct p rule: fmsize.induct) simp_all | 
| 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: diff
changeset | 70 | |
| 23274 | 71 | (* Semantics of formulae (fm) *) | 
| 72 | consts Ifm ::"bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool" | |
| 73 | primrec | |
| 74 | "Ifm bbs bs T = True" | |
| 75 | "Ifm bbs bs F = False" | |
| 76 | "Ifm bbs bs (Lt a) = (Inum bs a < 0)" | |
| 77 | "Ifm bbs bs (Gt a) = (Inum bs a > 0)" | |
| 78 | "Ifm bbs bs (Le a) = (Inum bs a \<le> 0)" | |
| 79 | "Ifm bbs bs (Ge a) = (Inum bs a \<ge> 0)" | |
| 80 | "Ifm bbs bs (Eq a) = (Inum bs a = 0)" | |
| 81 | "Ifm bbs bs (NEq a) = (Inum bs a \<noteq> 0)" | |
| 82 | "Ifm bbs bs (Dvd i b) = (i dvd Inum bs b)" | |
| 83 | "Ifm bbs bs (NDvd i b) = (\<not>(i dvd Inum bs b))" | |
| 84 | "Ifm bbs bs (NOT p) = (\<not> (Ifm bbs bs p))" | |
| 85 | "Ifm bbs bs (And p q) = (Ifm bbs bs p \<and> Ifm bbs bs q)" | |
| 86 | "Ifm bbs bs (Or p q) = (Ifm bbs bs p \<or> Ifm bbs bs q)" | |
| 87 | "Ifm bbs bs (Imp p q) = ((Ifm bbs bs p) \<longrightarrow> (Ifm bbs bs q))" | |
| 88 | "Ifm bbs bs (Iff p q) = (Ifm bbs bs p = Ifm bbs bs q)" | |
| 89 | "Ifm bbs bs (E p) = (\<exists> x. Ifm bbs (x#bs) p)" | |
| 90 | "Ifm bbs bs (A p) = (\<forall> x. Ifm bbs (x#bs) p)" | |
| 91 | "Ifm bbs bs (Closed n) = bbs!n" | |
| 92 | "Ifm bbs bs (NClosed n) = (\<not> bbs!n)" | |
| 93 | ||
| 94 | consts prep :: "fm \<Rightarrow> fm" | |
| 95 | recdef prep "measure fmsize" | |
| 96 | "prep (E T) = T" | |
| 97 | "prep (E F) = F" | |
| 98 | "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))" | |
| 99 | "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))" | |
| 100 | "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" | |
| 101 | "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))" | |
| 102 | "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))" | |
| 103 | "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" | |
| 104 | "prep (E p) = E (prep p)" | |
| 105 | "prep (A (And p q)) = And (prep (A p)) (prep (A q))" | |
| 106 | "prep (A p) = prep (NOT (E (NOT p)))" | |
| 107 | "prep (NOT (NOT p)) = prep p" | |
| 108 | "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))" | |
| 109 | "prep (NOT (A p)) = prep (E (NOT p))" | |
| 110 | "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))" | |
| 111 | "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))" | |
| 112 | "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))" | |
| 113 | "prep (NOT p) = NOT (prep p)" | |
| 114 | "prep (Or p q) = Or (prep p) (prep q)" | |
| 115 | "prep (And p q) = And (prep p) (prep q)" | |
| 116 | "prep (Imp p q) = prep (Or (NOT p) q)" | |
| 117 | "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))" | |
| 118 | "prep p = p" | |
| 25162 | 119 | (hints simp add: fmsize_pos) | 
| 23274 | 120 | lemma prep: "Ifm bbs bs (prep p) = Ifm bbs bs p" | 
| 121 | by (induct p arbitrary: bs rule: prep.induct, auto) | |
| 122 | ||
| 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: diff
changeset | 123 | |
| 23274 | 124 | (* Quantifier freeness *) | 
| 125 | consts qfree:: "fm \<Rightarrow> bool" | |
| 126 | recdef qfree "measure size" | |
| 127 | "qfree (E p) = False" | |
| 128 | "qfree (A p) = False" | |
| 129 | "qfree (NOT p) = qfree p" | |
| 130 | "qfree (And p q) = (qfree p \<and> qfree q)" | |
| 131 | "qfree (Or p q) = (qfree p \<and> qfree q)" | |
| 132 | "qfree (Imp p q) = (qfree p \<and> qfree q)" | |
| 133 | "qfree (Iff p q) = (qfree p \<and> qfree q)" | |
| 134 | "qfree p = True" | |
| 135 | ||
| 136 | (* Boundedness and substitution *) | |
| 137 | consts | |
| 138 | numbound0:: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *) | |
| 139 | bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *) | |
| 140 | subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *) | |
| 141 | primrec | |
| 142 | "numbound0 (C c) = True" | |
| 143 | "numbound0 (Bound n) = (n>0)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 144 | "numbound0 (CN n i a) = (n >0 \<and> numbound0 a)" | 
| 23274 | 145 | "numbound0 (Neg a) = numbound0 a" | 
| 146 | "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)" | |
| 147 | "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)" | |
| 148 | "numbound0 (Mul i a) = numbound0 a" | |
| 149 | ||
| 150 | lemma numbound0_I: | |
| 151 | assumes nb: "numbound0 a" | |
| 152 | shows "Inum (b#bs) a = Inum (b'#bs) a" | |
| 153 | using nb | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 154 | by (induct a rule: numbound0.induct) (auto simp add: gr0_conv_Suc) | 
| 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: diff
changeset | 155 | |
| 23274 | 156 | primrec | 
| 157 | "bound0 T = True" | |
| 158 | "bound0 F = True" | |
| 159 | "bound0 (Lt a) = numbound0 a" | |
| 160 | "bound0 (Le a) = numbound0 a" | |
| 161 | "bound0 (Gt a) = numbound0 a" | |
| 162 | "bound0 (Ge a) = numbound0 a" | |
| 163 | "bound0 (Eq a) = numbound0 a" | |
| 164 | "bound0 (NEq a) = numbound0 a" | |
| 165 | "bound0 (Dvd i a) = numbound0 a" | |
| 166 | "bound0 (NDvd i a) = numbound0 a" | |
| 167 | "bound0 (NOT p) = bound0 p" | |
| 168 | "bound0 (And p q) = (bound0 p \<and> bound0 q)" | |
| 169 | "bound0 (Or p q) = (bound0 p \<and> bound0 q)" | |
| 170 | "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))" | |
| 171 | "bound0 (Iff p q) = (bound0 p \<and> bound0 q)" | |
| 172 | "bound0 (E p) = False" | |
| 173 | "bound0 (A p) = False" | |
| 174 | "bound0 (Closed P) = True" | |
| 175 | "bound0 (NClosed P) = True" | |
| 176 | lemma bound0_I: | |
| 177 | assumes bp: "bound0 p" | |
| 178 | shows "Ifm bbs (b#bs) p = Ifm bbs (b'#bs) p" | |
| 179 | using bp numbound0_I[where b="b" and bs="bs" and b'="b'"] | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 180 | by (induct p rule: bound0.induct) (auto simp add: gr0_conv_Suc) | 
| 23274 | 181 | |
| 27456 | 182 | fun numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" where | 
| 23274 | 183 | "numsubst0 t (C c) = (C c)" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 184 | | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 185 | | "numsubst0 t (CN 0 i a) = Add (Mul i t) (numsubst0 t a)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 186 | | "numsubst0 t (CN n i a) = CN n i (numsubst0 t a)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 187 | | "numsubst0 t (Neg a) = Neg (numsubst0 t a)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 188 | | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 189 | | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 190 | | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)" | 
| 23274 | 191 | |
| 192 | lemma numsubst0_I: | |
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 193 | "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t" | 
| 25162 | 194 | by (induct t rule: numsubst0.induct,auto simp:nth_Cons') | 
| 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: diff
changeset | 195 | |
| 23274 | 196 | lemma numsubst0_I': | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 197 | "numbound0 a \<Longrightarrow> Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t" | 
| 25162 | 198 | by (induct t rule: numsubst0.induct, auto simp: nth_Cons' numbound0_I[where b="b" and b'="b'"]) | 
| 23274 | 199 | |
| 200 | primrec | |
| 201 | "subst0 t T = T" | |
| 202 | "subst0 t F = F" | |
| 203 | "subst0 t (Lt a) = Lt (numsubst0 t a)" | |
| 204 | "subst0 t (Le a) = Le (numsubst0 t a)" | |
| 205 | "subst0 t (Gt a) = Gt (numsubst0 t a)" | |
| 206 | "subst0 t (Ge a) = Ge (numsubst0 t a)" | |
| 207 | "subst0 t (Eq a) = Eq (numsubst0 t a)" | |
| 208 | "subst0 t (NEq a) = NEq (numsubst0 t a)" | |
| 209 | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)" | |
| 210 | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)" | |
| 211 | "subst0 t (NOT p) = NOT (subst0 t p)" | |
| 212 | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)" | |
| 213 | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)" | |
| 214 | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)" | |
| 215 | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)" | |
| 216 | "subst0 t (Closed P) = (Closed P)" | |
| 217 | "subst0 t (NClosed P) = (NClosed P)" | |
| 218 | ||
| 219 | lemma subst0_I: assumes qfp: "qfree p" | |
| 220 | shows "Ifm bbs (b#bs) (subst0 a p) = Ifm bbs ((Inum (b#bs) a)#bs) p" | |
| 221 | using qfp numsubst0_I[where b="b" and bs="bs" and a="a"] | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 222 | by (induct p) (simp_all add: gr0_conv_Suc) | 
| 23274 | 223 | |
| 224 | ||
| 225 | consts | |
| 226 | decrnum:: "num \<Rightarrow> num" | |
| 227 | decr :: "fm \<Rightarrow> fm" | |
| 228 | ||
| 229 | recdef decrnum "measure size" | |
| 230 | "decrnum (Bound n) = Bound (n - 1)" | |
| 231 | "decrnum (Neg a) = Neg (decrnum a)" | |
| 232 | "decrnum (Add a b) = Add (decrnum a) (decrnum b)" | |
| 233 | "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" | |
| 234 | "decrnum (Mul c a) = Mul c (decrnum a)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 235 | "decrnum (CN n i a) = (CN (n - 1) i (decrnum a))" | 
| 23274 | 236 | "decrnum a = a" | 
| 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: diff
changeset | 237 | |
| 23274 | 238 | recdef decr "measure size" | 
| 239 | "decr (Lt a) = Lt (decrnum a)" | |
| 240 | "decr (Le a) = Le (decrnum a)" | |
| 241 | "decr (Gt a) = Gt (decrnum a)" | |
| 242 | "decr (Ge a) = Ge (decrnum a)" | |
| 243 | "decr (Eq a) = Eq (decrnum a)" | |
| 244 | "decr (NEq a) = NEq (decrnum a)" | |
| 245 | "decr (Dvd i a) = Dvd i (decrnum a)" | |
| 246 | "decr (NDvd i a) = NDvd i (decrnum a)" | |
| 247 | "decr (NOT p) = NOT (decr p)" | |
| 248 | "decr (And p q) = And (decr p) (decr q)" | |
| 249 | "decr (Or p q) = Or (decr p) (decr q)" | |
| 250 | "decr (Imp p q) = Imp (decr p) (decr q)" | |
| 251 | "decr (Iff p q) = Iff (decr p) (decr q)" | |
| 252 | "decr p = p" | |
| 253 | ||
| 254 | lemma decrnum: assumes nb: "numbound0 t" | |
| 255 | shows "Inum (x#bs) t = Inum bs (decrnum t)" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 256 | using nb by (induct t rule: decrnum.induct, auto simp add: gr0_conv_Suc) | 
| 23274 | 257 | |
| 258 | lemma decr: assumes nb: "bound0 p" | |
| 259 | shows "Ifm bbs (x#bs) p = Ifm bbs bs (decr p)" | |
| 260 | using nb | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 261 | by (induct p rule: decr.induct, simp_all add: gr0_conv_Suc decrnum) | 
| 23274 | 262 | |
| 263 | lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)" | |
| 264 | by (induct p, simp_all) | |
| 265 | ||
| 266 | consts | |
| 267 | isatom :: "fm \<Rightarrow> bool" (* test for atomicity *) | |
| 268 | recdef isatom "measure size" | |
| 269 | "isatom T = True" | |
| 270 | "isatom F = True" | |
| 271 | "isatom (Lt a) = True" | |
| 272 | "isatom (Le a) = True" | |
| 273 | "isatom (Gt a) = True" | |
| 274 | "isatom (Ge a) = True" | |
| 275 | "isatom (Eq a) = True" | |
| 276 | "isatom (NEq a) = True" | |
| 277 | "isatom (Dvd i b) = True" | |
| 278 | "isatom (NDvd i b) = True" | |
| 279 | "isatom (Closed P) = True" | |
| 280 | "isatom (NClosed P) = True" | |
| 281 | "isatom p = False" | |
| 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: diff
changeset | 282 | |
| 23274 | 283 | lemma numsubst0_numbound0: assumes nb: "numbound0 t" | 
| 284 | shows "numbound0 (numsubst0 t a)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 285 | using nb apply (induct a rule: numbound0.induct) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 286 | apply simp_all | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 287 | apply (case_tac n, simp_all) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 288 | done | 
| 23274 | 289 | |
| 290 | lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t" | |
| 291 | shows "bound0 (subst0 t p)" | |
| 292 | using qf numsubst0_numbound0[OF nb] by (induct p rule: subst0.induct, auto) | |
| 293 | ||
| 294 | lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p" | |
| 295 | by (induct p, simp_all) | |
| 296 | ||
| 297 | ||
| 298 | constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
 | |
| 299 | "djf f p q \<equiv> (if q=T then T else if q=F then f p else | |
| 300 | (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))" | |
| 301 | constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
 | |
| 302 | "evaldjf f ps \<equiv> foldr (djf f) ps F" | |
| 303 | ||
| 304 | lemma djf_Or: "Ifm bbs bs (djf f p q) = Ifm bbs bs (Or (f p) q)" | |
| 305 | by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) | |
| 306 | (cases "f p", simp_all add: Let_def djf_def) | |
| 307 | ||
| 308 | lemma evaldjf_ex: "Ifm bbs bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm bbs bs (f p))" | |
| 309 | by(induct ps, simp_all add: evaldjf_def djf_Or) | |
| 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: diff
changeset | 310 | |
| 23274 | 311 | lemma evaldjf_bound0: | 
| 312 | assumes nb: "\<forall> x\<in> set xs. bound0 (f x)" | |
| 313 | shows "bound0 (evaldjf f xs)" | |
| 314 | using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) | |
| 315 | ||
| 316 | lemma evaldjf_qf: | |
| 317 | assumes nb: "\<forall> x\<in> set xs. qfree (f x)" | |
| 318 | shows "qfree (evaldjf f xs)" | |
| 319 | using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) | |
| 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: diff
changeset | 320 | |
| 23274 | 321 | consts disjuncts :: "fm \<Rightarrow> fm list" | 
| 322 | recdef disjuncts "measure size" | |
| 323 | "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)" | |
| 324 | "disjuncts F = []" | |
| 325 | "disjuncts p = [p]" | |
| 326 | ||
| 327 | lemma disjuncts: "(\<exists> q\<in> set (disjuncts p). Ifm bbs bs q) = Ifm bbs bs p" | |
| 328 | by(induct p rule: disjuncts.induct, auto) | |
| 329 | ||
| 330 | lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). bound0 q" | |
| 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: diff
changeset | 331 | proof- | 
| 23274 | 332 | assume nb: "bound0 p" | 
| 333 | hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto) | |
| 334 | thus ?thesis by (simp only: list_all_iff) | |
| 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: diff
changeset | 335 | qed | 
| 
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: diff
changeset | 336 | |
| 23274 | 337 | lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). qfree q" | 
| 338 | proof- | |
| 339 | assume qf: "qfree p" | |
| 340 | hence "list_all qfree (disjuncts p)" | |
| 341 | by (induct p rule: disjuncts.induct, auto) | |
| 342 | thus ?thesis by (simp only: list_all_iff) | |
| 343 | qed | |
| 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: diff
changeset | 344 | |
| 23274 | 345 | constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" | 
| 346 | "DJ f p \<equiv> evaldjf f (disjuncts p)" | |
| 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: diff
changeset | 347 | |
| 23274 | 348 | lemma DJ: assumes fdj: "\<forall> p q. f (Or p q) = Or (f p) (f q)" | 
| 349 | and fF: "f F = F" | |
| 350 | shows "Ifm bbs bs (DJ f p) = Ifm bbs bs (f p)" | |
| 351 | proof- | |
| 352 | have "Ifm bbs bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm bbs bs (f q))" | |
| 353 | by (simp add: DJ_def evaldjf_ex) | |
| 354 | also have "\<dots> = Ifm bbs bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) | |
| 355 | finally show ?thesis . | |
| 356 | qed | |
| 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: diff
changeset | 357 | |
| 23274 | 358 | lemma DJ_qf: assumes | 
| 359 | fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)" | |
| 360 | shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) " | |
| 361 | proof(clarify) | |
| 362 | fix p assume qf: "qfree p" | |
| 363 | have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def) | |
| 364 | from disjuncts_qf[OF qf] have "\<forall> q\<in> set (disjuncts p). qfree q" . | |
| 365 | with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast | |
| 366 | ||
| 367 | from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp | |
| 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: diff
changeset | 368 | qed | 
| 
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: diff
changeset | 369 | |
| 23274 | 370 | lemma DJ_qe: assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bbs bs (qe p) = Ifm bbs bs (E p))" | 
| 371 | shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm bbs bs ((DJ qe p)) = Ifm bbs bs (E p))" | |
| 372 | proof(clarify) | |
| 373 | fix p::fm and bs | |
| 374 | assume qf: "qfree p" | |
| 375 | from qe have qth: "\<forall> p. qfree p \<longrightarrow> qfree (qe p)" by blast | |
| 376 | from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto | |
| 377 | have "Ifm bbs bs (DJ qe p) = (\<exists> q\<in> set (disjuncts p). Ifm bbs bs (qe q))" | |
| 378 | by (simp add: DJ_def evaldjf_ex) | |
| 379 | also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm bbs bs (E q))" using qe disjuncts_qf[OF qf] by auto | |
| 380 | also have "\<dots> = Ifm bbs bs (E p)" by (induct p rule: disjuncts.induct, auto) | |
| 381 | finally show "qfree (DJ qe p) \<and> Ifm bbs bs (DJ qe p) = Ifm bbs bs (E p)" using qfth by blast | |
| 382 | qed | |
| 383 | (* Simplification *) | |
| 384 | ||
| 385 | (* Algebraic simplifications for nums *) | |
| 386 | consts bnds:: "num \<Rightarrow> nat list" | |
| 387 | lex_ns:: "nat list \<times> nat list \<Rightarrow> bool" | |
| 388 | recdef bnds "measure size" | |
| 389 | "bnds (Bound n) = [n]" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 390 | "bnds (CN n c a) = n#(bnds a)" | 
| 23274 | 391 | "bnds (Neg a) = bnds a" | 
| 392 | "bnds (Add a b) = (bnds a)@(bnds b)" | |
| 393 | "bnds (Sub a b) = (bnds a)@(bnds b)" | |
| 394 | "bnds (Mul i a) = bnds a" | |
| 395 | "bnds a = []" | |
| 396 | recdef lex_ns "measure (\<lambda> (xs,ys). length xs + length ys)" | |
| 397 | "lex_ns ([], ms) = True" | |
| 398 | "lex_ns (ns, []) = False" | |
| 399 | "lex_ns (n#ns, m#ms) = (n<m \<or> ((n = m) \<and> lex_ns (ns,ms))) " | |
| 400 | constdefs lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" | |
| 401 | "lex_bnd t s \<equiv> lex_ns (bnds t, bnds s)" | |
| 402 | ||
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 403 | consts | 
| 23274 | 404 | numadd:: "num \<times> num \<Rightarrow> num" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 405 | recdef numadd "measure (\<lambda> (t,s). num_size t + num_size s)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 406 | "numadd (CN n1 c1 r1 ,CN n2 c2 r2) = | 
| 23274 | 407 | (if n1=n2 then | 
| 408 | (let c = c1 + c2 | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 409 | in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2)))) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 410 | else if n1 \<le> n2 then CN n1 c1 (numadd (r1,Add (Mul c2 (Bound n2)) r2)) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 411 | else CN n2 c2 (numadd (Add (Mul c1 (Bound n1)) r1,r2)))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 412 | "numadd (CN n1 c1 r1, t) = CN n1 c1 (numadd (r1, t))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 413 | "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" | 
| 23274 | 414 | "numadd (C b1, C b2) = C (b1+b2)" | 
| 415 | "numadd (a,b) = Add a b" | |
| 416 | ||
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 417 | (*function (sequential) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 418 | numadd :: "num \<Rightarrow> num \<Rightarrow> num" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 419 | where | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 420 | "numadd (Add (Mul c1 (Bound n1)) r1) (Add (Mul c2 (Bound n2)) r2) = | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 421 | (if n1 = n2 then (let c = c1 + c2 | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 422 | in (if c = 0 then numadd r1 r2 else | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 423 | Add (Mul c (Bound n1)) (numadd r1 r2))) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 424 | else if n1 \<le> n2 then | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 425 | Add (Mul c1 (Bound n1)) (numadd r1 (Add (Mul c2 (Bound n2)) r2)) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 426 | else | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 427 | Add (Mul c2 (Bound n2)) (numadd (Add (Mul c1 (Bound n1)) r1) r2))" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 428 | | "numadd (Add (Mul c1 (Bound n1)) r1) t = | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 429 | Add (Mul c1 (Bound n1)) (numadd r1 t)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 430 | | "numadd t (Add (Mul c2 (Bound n2)) r2) = | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 431 | Add (Mul c2 (Bound n2)) (numadd t r2)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 432 | | "numadd (C b1) (C b2) = C (b1 + b2)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 433 | | "numadd a b = Add a b" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 434 | apply pat_completeness apply auto*) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 435 | |
| 23274 | 436 | lemma numadd: "Inum bs (numadd (t,s)) = Inum bs (Add t s)" | 
| 437 | apply (induct t s rule: numadd.induct, simp_all add: Let_def) | |
| 438 | apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all) | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 439 | apply (case_tac "n1 = n2") | 
| 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 440 | apply(simp_all add: ring_simps) | 
| 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 441 | apply(simp add: left_distrib[symmetric]) | 
| 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 442 | done | 
| 23274 | 443 | |
| 444 | lemma numadd_nb: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))" | |
| 445 | by (induct t s rule: numadd.induct, auto simp add: Let_def) | |
| 446 | ||
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 447 | fun | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 448 | nummul :: "int \<Rightarrow> num \<Rightarrow> num" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 449 | where | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 450 | "nummul i (C j) = C (i * j)" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 451 | | "nummul i (CN n c t) = CN n (c*i) (nummul i t)" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 452 | | "nummul i t = Mul i t" | 
| 23274 | 453 | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 454 | lemma nummul: "\<And> i. Inum bs (nummul i t) = Inum bs (Mul i t)" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 455 | by (induct t rule: nummul.induct, auto simp add: ring_simps numadd) | 
| 23274 | 456 | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 457 | lemma nummul_nb: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul i t)" | 
| 23274 | 458 | by (induct t rule: nummul.induct, auto simp add: numadd_nb) | 
| 459 | ||
| 460 | constdefs numneg :: "num \<Rightarrow> num" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 461 | "numneg t \<equiv> nummul (- 1) t" | 
| 23274 | 462 | |
| 463 | constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 464 | "numsub s t \<equiv> (if s = t then C 0 else numadd (s, numneg t))" | 
| 23274 | 465 | |
| 466 | lemma numneg: "Inum bs (numneg t) = Inum bs (Neg t)" | |
| 467 | using numneg_def nummul by simp | |
| 468 | ||
| 469 | lemma numneg_nb: "numbound0 t \<Longrightarrow> numbound0 (numneg t)" | |
| 470 | using numneg_def nummul_nb by simp | |
| 471 | ||
| 472 | lemma numsub: "Inum bs (numsub a b) = Inum bs (Sub a b)" | |
| 473 | using numneg numadd numsub_def by simp | |
| 474 | ||
| 475 | lemma numsub_nb: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numsub t s)" | |
| 476 | using numsub_def numadd_nb numneg_nb by simp | |
| 477 | ||
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 478 | fun | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 479 | simpnum :: "num \<Rightarrow> num" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 480 | where | 
| 23274 | 481 | "simpnum (C j) = C j" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 482 | | "simpnum (Bound n) = CN n 1 (C 0)" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 483 | | "simpnum (Neg t) = numneg (simpnum t)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 484 | | "simpnum (Add t s) = numadd (simpnum t, simpnum s)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 485 | | "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 486 | | "simpnum (Mul i t) = (if i = 0 then C 0 else nummul i (simpnum t))" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 487 | | "simpnum t = t" | 
| 23274 | 488 | |
| 489 | lemma simpnum_ci: "Inum bs (simpnum t) = Inum bs t" | |
| 490 | by (induct t rule: simpnum.induct, auto simp add: numneg numadd numsub nummul) | |
| 491 | ||
| 492 | lemma simpnum_numbound0: | |
| 493 | "numbound0 t \<Longrightarrow> numbound0 (simpnum t)" | |
| 494 | by (induct t rule: simpnum.induct, auto simp add: numadd_nb numsub_nb nummul_nb numneg_nb) | |
| 495 | ||
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 496 | fun | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 497 | not :: "fm \<Rightarrow> fm" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 498 | where | 
| 23274 | 499 | "not (NOT p) = p" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 500 | | "not T = F" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 501 | | "not F = T" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 502 | | "not p = NOT p" | 
| 23274 | 503 | lemma not: "Ifm bbs bs (not p) = Ifm bbs bs (NOT p)" | 
| 504 | by (cases p) auto | |
| 505 | lemma not_qf: "qfree p \<Longrightarrow> qfree (not p)" | |
| 506 | by (cases p, auto) | |
| 507 | lemma not_bn: "bound0 p \<Longrightarrow> bound0 (not p)" | |
| 508 | by (cases p, auto) | |
| 509 | ||
| 510 | constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" | |
| 511 | "conj p q \<equiv> (if (p = F \<or> q=F) then F else if p=T then q else if q=T then p else And p q)" | |
| 512 | lemma conj: "Ifm bbs bs (conj p q) = Ifm bbs bs (And p q)" | |
| 513 | by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all) | |
| 514 | ||
| 515 | lemma conj_qf: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)" | |
| 516 | using conj_def by auto | |
| 517 | lemma conj_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)" | |
| 518 | using conj_def by auto | |
| 519 | ||
| 520 | constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" | |
| 521 | "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p else Or p q)" | |
| 522 | ||
| 523 | lemma disj: "Ifm bbs bs (disj p q) = Ifm bbs bs (Or p q)" | |
| 524 | by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all) | |
| 525 | lemma disj_qf: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)" | |
| 526 | using disj_def by auto | |
| 527 | lemma disj_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)" | |
| 528 | using disj_def by auto | |
| 529 | ||
| 530 | constdefs imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" | |
| 531 | "imp p q \<equiv> (if (p = F \<or> q=T) then T else if p=T then q else if q=F then not p else Imp p q)" | |
| 532 | lemma imp: "Ifm bbs bs (imp p q) = Ifm bbs bs (Imp p q)" | |
| 533 | by (cases "p=F \<or> q=T",simp_all add: imp_def,cases p) (simp_all add: not) | |
| 534 | lemma imp_qf: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)" | |
| 535 | using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def,cases p) (simp_all add: not_qf) | |
| 536 | lemma imp_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)" | |
| 537 | using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def,cases p) simp_all | |
| 538 | ||
| 539 | constdefs iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" | |
| 540 | "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else | |
| 541 | if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else | |
| 542 | Iff p q)" | |
| 543 | lemma iff: "Ifm bbs bs (iff p q) = Ifm bbs bs (Iff p q)" | |
| 544 | by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp add:not) | |
| 545 | (cases "not p= q", auto simp add:not) | |
| 546 | lemma iff_qf: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)" | |
| 547 | by (unfold iff_def,cases "p=q", auto simp add: not_qf) | |
| 548 | lemma iff_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (iff p q)" | |
| 549 | using iff_def by (unfold iff_def,cases "p=q", auto simp add: not_bn) | |
| 550 | ||
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 551 | function (sequential) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 552 | simpfm :: "fm \<Rightarrow> fm" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 553 | where | 
| 23274 | 554 | "simpfm (And p q) = conj (simpfm p) (simpfm q)" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 555 | | "simpfm (Or p q) = disj (simpfm p) (simpfm q)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 556 | | "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 557 | | "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 558 | | "simpfm (NOT p) = not (simpfm p)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 559 | | "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 560 | | _ \<Rightarrow> Lt a')" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 561 | | "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<le> 0) then T else F | _ \<Rightarrow> Le a')" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 562 | | "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v > 0) then T else F | _ \<Rightarrow> Gt a')" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 563 | | "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<ge> 0) then T else F | _ \<Rightarrow> Ge a')" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 564 | | "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v = 0) then T else F | _ \<Rightarrow> Eq a')" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 565 | | "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<noteq> 0) then T else F | _ \<Rightarrow> NEq a')" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 566 | | "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a) | 
| 23274 | 567 | else if (abs i = 1) then T | 
| 568 | else let a' = simpnum a in case a' of C v \<Rightarrow> if (i dvd v) then T else F | _ \<Rightarrow> Dvd i a')" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 569 | | "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) | 
| 23274 | 570 | else if (abs i = 1) then F | 
| 571 | else let a' = simpnum a in case a' of C v \<Rightarrow> if (\<not>(i dvd v)) then T else F | _ \<Rightarrow> NDvd i a')" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 572 | | "simpfm p = p" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 573 | by pat_completeness auto | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 574 | termination by (relation "measure fmsize") auto | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 575 | |
| 23274 | 576 | lemma simpfm: "Ifm bbs bs (simpfm p) = Ifm bbs bs p" | 
| 577 | proof(induct p rule: simpfm.induct) | |
| 578 | case (6 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp | |
| 579 |   {fix v assume "?sa = C v" hence ?case using sa by simp }
 | |
| 580 |   moreover {assume "\<not> (\<exists> v. ?sa = C v)" hence ?case using sa 
 | |
| 581 | by (cases ?sa, simp_all add: Let_def)} | |
| 582 | ultimately show ?case by blast | |
| 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: diff
changeset | 583 | next | 
| 23274 | 584 | case (7 a) let ?sa = "simpnum a" | 
| 585 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp | |
| 586 |   {fix v assume "?sa = C v" hence ?case using sa by simp }
 | |
| 587 |   moreover {assume "\<not> (\<exists> v. ?sa = C v)" hence ?case using sa 
 | |
| 588 | by (cases ?sa, simp_all add: Let_def)} | |
| 589 | ultimately show ?case by blast | |
| 590 | next | |
| 591 | case (8 a) let ?sa = "simpnum a" | |
| 592 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp | |
| 593 |   {fix v assume "?sa = C v" hence ?case using sa by simp }
 | |
| 594 |   moreover {assume "\<not> (\<exists> v. ?sa = C v)" hence ?case using sa 
 | |
| 595 | by (cases ?sa, simp_all add: Let_def)} | |
| 596 | ultimately show ?case by blast | |
| 597 | next | |
| 598 | case (9 a) let ?sa = "simpnum a" | |
| 599 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp | |
| 600 |   {fix v assume "?sa = C v" hence ?case using sa by simp }
 | |
| 601 |   moreover {assume "\<not> (\<exists> v. ?sa = C v)" hence ?case using sa 
 | |
| 602 | by (cases ?sa, simp_all add: Let_def)} | |
| 603 | ultimately show ?case by blast | |
| 604 | next | |
| 605 | case (10 a) let ?sa = "simpnum a" | |
| 606 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp | |
| 607 |   {fix v assume "?sa = C v" hence ?case using sa by simp }
 | |
| 608 |   moreover {assume "\<not> (\<exists> v. ?sa = C v)" hence ?case using sa 
 | |
| 609 | by (cases ?sa, simp_all add: Let_def)} | |
| 610 | ultimately show ?case by blast | |
| 611 | next | |
| 612 | case (11 a) let ?sa = "simpnum a" | |
| 613 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp | |
| 614 |   {fix v assume "?sa = C v" hence ?case using sa by simp }
 | |
| 615 |   moreover {assume "\<not> (\<exists> v. ?sa = C v)" hence ?case using sa 
 | |
| 616 | by (cases ?sa, simp_all add: Let_def)} | |
| 617 | ultimately show ?case by blast | |
| 618 | next | |
| 619 | case (12 i a) let ?sa = "simpnum a" from simpnum_ci | |
| 620 | have sa: "Inum bs ?sa = Inum bs a" by simp | |
| 621 | have "i=0 \<or> abs i = 1 \<or> (i\<noteq>0 \<and> (abs i \<noteq> 1))" by auto | |
| 622 |   {assume "i=0" hence ?case using "12.hyps" by (simp add: dvd_def Let_def)}
 | |
| 623 | moreover | |
| 624 |   {assume i1: "abs i = 1"
 | |
| 625 | from zdvd_1_left[where m = "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"] | |
| 23315 | 626 | have ?case using i1 apply (cases "i=0", simp_all add: Let_def) | 
| 627 | by (cases "i > 0", simp_all)} | |
| 23274 | 628 | moreover | 
| 629 |   {assume inz: "i\<noteq>0" and cond: "abs i \<noteq> 1"
 | |
| 630 |     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
 | |
| 631 | by (cases "abs i = 1", auto) } | |
| 632 |     moreover {assume "\<not> (\<exists> v. ?sa = C v)" 
 | |
| 633 | hence "simpfm (Dvd i a) = Dvd i ?sa" using inz cond | |
| 634 | by (cases ?sa, auto simp add: Let_def) | |
| 635 | hence ?case using sa by simp} | |
| 636 | ultimately have ?case by blast} | |
| 637 | ultimately show ?case by blast | |
| 638 | next | |
| 639 | case (13 i a) let ?sa = "simpnum a" from simpnum_ci | |
| 640 | have sa: "Inum bs ?sa = Inum bs a" by simp | |
| 641 | have "i=0 \<or> abs i = 1 \<or> (i\<noteq>0 \<and> (abs i \<noteq> 1))" by auto | |
| 642 |   {assume "i=0" hence ?case using "13.hyps" by (simp add: dvd_def Let_def)}
 | |
| 643 | moreover | |
| 644 |   {assume i1: "abs i = 1"
 | |
| 645 | from zdvd_1_left[where m = "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"] | |
| 23315 | 646 | have ?case using i1 apply (cases "i=0", simp_all add: Let_def) | 
| 647 | apply (cases "i > 0", simp_all) done} | |
| 23274 | 648 | moreover | 
| 649 |   {assume inz: "i\<noteq>0" and cond: "abs i \<noteq> 1"
 | |
| 650 |     {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
 | |
| 651 | by (cases "abs i = 1", auto) } | |
| 652 |     moreover {assume "\<not> (\<exists> v. ?sa = C v)" 
 | |
| 653 | hence "simpfm (NDvd i a) = NDvd i ?sa" using inz cond | |
| 654 | by (cases ?sa, auto simp add: Let_def) | |
| 655 | hence ?case using sa by simp} | |
| 656 | ultimately have ?case by blast} | |
| 657 | ultimately show ?case by blast | |
| 658 | qed (induct p rule: simpfm.induct, simp_all add: conj disj imp iff not) | |
| 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: diff
changeset | 659 | |
| 23274 | 660 | lemma simpfm_bound0: "bound0 p \<Longrightarrow> bound0 (simpfm p)" | 
| 661 | proof(induct p rule: simpfm.induct) | |
| 662 | case (6 a) hence nb: "numbound0 a" by simp | |
| 663 | hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | |
| 664 | thus ?case by (cases "simpnum a", auto simp add: Let_def) | |
| 665 | next | |
| 666 | case (7 a) hence nb: "numbound0 a" by simp | |
| 667 | hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | |
| 668 | thus ?case by (cases "simpnum a", auto simp add: Let_def) | |
| 669 | next | |
| 670 | case (8 a) hence nb: "numbound0 a" by simp | |
| 671 | hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | |
| 672 | thus ?case by (cases "simpnum a", auto simp add: Let_def) | |
| 673 | next | |
| 674 | case (9 a) hence nb: "numbound0 a" by simp | |
| 675 | hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | |
| 676 | thus ?case by (cases "simpnum a", auto simp add: Let_def) | |
| 677 | next | |
| 678 | case (10 a) hence nb: "numbound0 a" by simp | |
| 679 | hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | |
| 680 | thus ?case by (cases "simpnum a", auto simp add: Let_def) | |
| 681 | next | |
| 682 | case (11 a) hence nb: "numbound0 a" by simp | |
| 683 | hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | |
| 684 | thus ?case by (cases "simpnum a", auto simp add: Let_def) | |
| 685 | next | |
| 686 | case (12 i a) hence nb: "numbound0 a" by simp | |
| 687 | hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | |
| 688 | thus ?case by (cases "simpnum a", auto simp add: Let_def) | |
| 689 | next | |
| 690 | case (13 i a) hence nb: "numbound0 a" by simp | |
| 691 | hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | |
| 692 | thus ?case by (cases "simpnum a", auto simp add: Let_def) | |
| 693 | qed(auto simp add: disj_def imp_def iff_def conj_def not_bn) | |
| 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: diff
changeset | 694 | |
| 23274 | 695 | lemma simpfm_qf: "qfree p \<Longrightarrow> qfree (simpfm p)" | 
| 696 | by (induct p rule: simpfm.induct, auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def) | |
| 697 | (case_tac "simpnum a",auto)+ | |
| 698 | ||
| 699 | (* Generic quantifier elimination *) | |
| 700 | consts qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm" | |
| 701 | recdef qelim "measure fmsize" | |
| 702 | "qelim (E p) = (\<lambda> qe. DJ qe (qelim p qe))" | |
| 703 | "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))" | |
| 704 | "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))" | |
| 705 | "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))" | |
| 706 | "qelim (Or p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))" | |
| 707 | "qelim (Imp p q) = (\<lambda> qe. imp (qelim p qe) (qelim q qe))" | |
| 708 | "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))" | |
| 709 | "qelim p = (\<lambda> y. simpfm p)" | |
| 710 | ||
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 711 | (*function (sequential) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 712 | qelim :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 713 | where | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 714 | "qelim qe (E p) = DJ qe (qelim qe p)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 715 | | "qelim qe (A p) = not (qe ((qelim qe (NOT p))))" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 716 | | "qelim qe (NOT p) = not (qelim qe p)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 717 | | "qelim qe (And p q) = conj (qelim qe p) (qelim qe q)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 718 | | "qelim qe (Or p q) = disj (qelim qe p) (qelim qe q)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 719 | | "qelim qe (Imp p q) = imp (qelim qe p) (qelim qe q)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 720 | | "qelim qe (Iff p q) = iff (qelim qe p) (qelim qe q)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 721 | | "qelim qe p = simpfm p" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 722 | by pat_completeness auto | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 723 | termination by (relation "measure (fmsize o snd)") auto*) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 724 | |
| 23274 | 725 | lemma qelim_ci: | 
| 726 | assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bbs bs (qe p) = Ifm bbs bs (E p))" | |
| 727 | shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm bbs bs (qelim p qe) = Ifm bbs bs p)" | |
| 728 | using qe_inv DJ_qe[OF qe_inv] | |
| 729 | by(induct p rule: qelim.induct) | |
| 730 | (auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf | |
| 731 | simpfm simpfm_qf simp del: simpfm.simps) | |
| 732 | (* Linearity for fm where Bound 0 ranges over \<int> *) | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 733 | |
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 734 | fun | 
| 23274 | 735 | zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*) | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 736 | where | 
| 23274 | 737 | "zsplit0 (C c) = (0,C c)" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 738 | | "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 739 | | "zsplit0 (CN n i a) = | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 740 | (let (i',a') = zsplit0 a | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 741 | in if n=0 then (i+i', a') else (i',CN n i a'))" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 742 | | "zsplit0 (Neg a) = (let (i',a') = zsplit0 a in (-i', Neg a'))" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 743 | | "zsplit0 (Add a b) = (let (ia,a') = zsplit0 a ; | 
| 23274 | 744 | (ib,b') = zsplit0 b | 
| 745 | in (ia+ib, Add a' b'))" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 746 | | "zsplit0 (Sub a b) = (let (ia,a') = zsplit0 a ; | 
| 23274 | 747 | (ib,b') = zsplit0 b | 
| 748 | in (ia-ib, Sub a' b'))" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 749 | | "zsplit0 (Mul i a) = (let (i',a') = zsplit0 a in (i*i', Mul i a'))" | 
| 23274 | 750 | |
| 751 | lemma zsplit0_I: | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 752 | shows "\<And> n a. zsplit0 t = (n,a) \<Longrightarrow> (Inum ((x::int) #bs) (CN 0 n a) = Inum (x #bs) t) \<and> numbound0 a" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 753 | (is "\<And> n a. ?S t = (n,a) \<Longrightarrow> (?I x (CN 0 n a) = ?I x t) \<and> ?N a") | 
| 23274 | 754 | proof(induct t rule: zsplit0.induct) | 
| 755 | case (1 c n a) thus ?case by auto | |
| 756 | next | |
| 757 | case (2 m n a) thus ?case by (cases "m=0") auto | |
| 758 | next | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 759 | case (3 m i a n a') | 
| 23274 | 760 | let ?j = "fst (zsplit0 a)" | 
| 761 | let ?b = "snd (zsplit0 a)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 762 | have abj: "zsplit0 a = (?j,?b)" by simp | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 763 |   {assume "m\<noteq>0" 
 | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 764 | with prems(1)[OF abj] prems(2) have ?case by (auto simp add: Let_def split_def)} | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 765 | moreover | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 766 |   {assume m0: "m =0"
 | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 767 | from abj have th: "a'=?b \<and> n=i+?j" using prems | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 768 | by (simp add: Let_def split_def) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 769 | from abj prems have th2: "(?I x (CN 0 ?j ?b) = ?I x a) \<and> ?N ?b" by blast | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 770 | from th have "?I x (CN 0 n a') = ?I x (CN 0 (i+?j) ?b)" by simp | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 771 | also from th2 have "\<dots> = ?I x (CN 0 i (CN 0 ?j ?b))" by (simp add: left_distrib) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 772 | finally have "?I x (CN 0 n a') = ?I x (CN 0 i a)" using th2 by simp | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 773 | with th2 th have ?case using m0 by blast} | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 774 | ultimately show ?case by blast | 
| 23274 | 775 | next | 
| 776 | case (4 t n a) | |
| 777 | let ?nt = "fst (zsplit0 t)" | |
| 778 | let ?at = "snd (zsplit0 t)" | |
| 779 | have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \<and> n=-?nt" using prems | |
| 780 | by (simp add: Let_def split_def) | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 781 | from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast | 
| 23274 | 782 | from th2[simplified] th[simplified] show ?case by simp | 
| 783 | next | |
| 784 | case (5 s t n a) | |
| 785 | let ?ns = "fst (zsplit0 s)" | |
| 786 | let ?as = "snd (zsplit0 s)" | |
| 787 | let ?nt = "fst (zsplit0 t)" | |
| 788 | let ?at = "snd (zsplit0 t)" | |
| 789 | have abjs: "zsplit0 s = (?ns,?as)" by simp | |
| 790 | moreover have abjt: "zsplit0 t = (?nt,?at)" by simp | |
| 791 | ultimately have th: "a=Add ?as ?at \<and> n=?ns + ?nt" using prems | |
| 792 | by (simp add: Let_def split_def) | |
| 793 | from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 794 | from prems have "(\<exists> x y. (x,y) = zsplit0 s) \<longrightarrow> (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (x # bs) (CN 0 xa xb) = Inum (x # bs) t \<and> numbound0 xb)" by auto | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 795 | with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 796 | from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast | 
| 23274 | 797 | from th3[simplified] th2[simplified] th[simplified] show ?case | 
| 798 | by (simp add: left_distrib) | |
| 799 | next | |
| 800 | case (6 s t n a) | |
| 801 | let ?ns = "fst (zsplit0 s)" | |
| 802 | let ?as = "snd (zsplit0 s)" | |
| 803 | let ?nt = "fst (zsplit0 t)" | |
| 804 | let ?at = "snd (zsplit0 t)" | |
| 805 | have abjs: "zsplit0 s = (?ns,?as)" by simp | |
| 806 | moreover have abjt: "zsplit0 t = (?nt,?at)" by simp | |
| 807 | ultimately have th: "a=Sub ?as ?at \<and> n=?ns - ?nt" using prems | |
| 808 | by (simp add: Let_def split_def) | |
| 809 | from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 810 | from prems have "(\<exists> x y. (x,y) = zsplit0 s) \<longrightarrow> (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (x # bs) (CN 0 xa xb) = Inum (x # bs) t \<and> numbound0 xb)" by auto | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 811 | with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 812 | from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast | 
| 23274 | 813 | from th3[simplified] th2[simplified] th[simplified] show ?case | 
| 814 | by (simp add: left_diff_distrib) | |
| 815 | next | |
| 816 | case (7 i t n a) | |
| 817 | let ?nt = "fst (zsplit0 t)" | |
| 818 | let ?at = "snd (zsplit0 t)" | |
| 819 | have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \<and> n=i*?nt" using prems | |
| 820 | by (simp add: Let_def split_def) | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 821 | from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 822 | hence " ?I x (Mul i t) = i * ?I x (CN 0 ?nt ?at)" by simp | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 823 | also have "\<dots> = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: right_distrib) | 
| 23274 | 824 | finally show ?case using th th2 by simp | 
| 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: diff
changeset | 825 | qed | 
| 
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: diff
changeset | 826 | |
| 23274 | 827 | consts | 
| 828 | iszlfm :: "fm \<Rightarrow> bool" (* Linearity test for fm *) | |
| 829 | recdef iszlfm "measure size" | |
| 830 | "iszlfm (And p q) = (iszlfm p \<and> iszlfm q)" | |
| 831 | "iszlfm (Or p q) = (iszlfm p \<and> iszlfm q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 832 | "iszlfm (Eq (CN 0 c e)) = (c>0 \<and> numbound0 e)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 833 | "iszlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 834 | "iszlfm (Lt (CN 0 c e)) = (c>0 \<and> numbound0 e)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 835 | "iszlfm (Le (CN 0 c e)) = (c>0 \<and> numbound0 e)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 836 | "iszlfm (Gt (CN 0 c e)) = (c>0 \<and> numbound0 e)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 837 | "iszlfm (Ge (CN 0 c e)) = ( c>0 \<and> numbound0 e)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 838 | "iszlfm (Dvd i (CN 0 c e)) = | 
| 23274 | 839 | (c>0 \<and> i>0 \<and> numbound0 e)" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 840 | "iszlfm (NDvd i (CN 0 c e))= | 
| 23274 | 841 | (c>0 \<and> i>0 \<and> numbound0 e)" | 
| 842 | "iszlfm p = (isatom p \<and> (bound0 p))" | |
| 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: diff
changeset | 843 | |
| 23274 | 844 | lemma zlin_qfree: "iszlfm p \<Longrightarrow> qfree p" | 
| 845 | by (induct p rule: iszlfm.induct) auto | |
| 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: diff
changeset | 846 | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 847 | consts | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 848 | zlfm :: "fm \<Rightarrow> fm" (* Linearity transformation for fm *) | 
| 23274 | 849 | recdef zlfm "measure fmsize" | 
| 850 | "zlfm (And p q) = And (zlfm p) (zlfm q)" | |
| 851 | "zlfm (Or p q) = Or (zlfm p) (zlfm q)" | |
| 852 | "zlfm (Imp p q) = Or (zlfm (NOT p)) (zlfm q)" | |
| 853 | "zlfm (Iff p q) = Or (And (zlfm p) (zlfm q)) (And (zlfm (NOT p)) (zlfm (NOT q)))" | |
| 854 | "zlfm (Lt a) = (let (c,r) = zsplit0 a in | |
| 855 | if c=0 then Lt r else | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 856 | if c>0 then (Lt (CN 0 c r)) else (Gt (CN 0 (- c) (Neg r))))" | 
| 23274 | 857 | "zlfm (Le a) = (let (c,r) = zsplit0 a in | 
| 858 | if c=0 then Le r else | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 859 | if c>0 then (Le (CN 0 c r)) else (Ge (CN 0 (- c) (Neg r))))" | 
| 23274 | 860 | "zlfm (Gt a) = (let (c,r) = zsplit0 a in | 
| 861 | if c=0 then Gt r else | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 862 | if c>0 then (Gt (CN 0 c r)) else (Lt (CN 0 (- c) (Neg r))))" | 
| 23274 | 863 | "zlfm (Ge a) = (let (c,r) = zsplit0 a in | 
| 864 | if c=0 then Ge r else | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 865 | if c>0 then (Ge (CN 0 c r)) else (Le (CN 0 (- c) (Neg r))))" | 
| 23274 | 866 | "zlfm (Eq a) = (let (c,r) = zsplit0 a in | 
| 867 | if c=0 then Eq r else | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 868 | if c>0 then (Eq (CN 0 c r)) else (Eq (CN 0 (- c) (Neg r))))" | 
| 23274 | 869 | "zlfm (NEq a) = (let (c,r) = zsplit0 a in | 
| 870 | if c=0 then NEq r else | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 871 | if c>0 then (NEq (CN 0 c r)) else (NEq (CN 0 (- c) (Neg r))))" | 
| 23274 | 872 | "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) | 
| 873 | else (let (c,r) = zsplit0 a in | |
| 874 | if c=0 then (Dvd (abs i) r) else | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 875 | if c>0 then (Dvd (abs i) (CN 0 c r)) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 876 | else (Dvd (abs i) (CN 0 (- c) (Neg r)))))" | 
| 23274 | 877 | "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) | 
| 878 | else (let (c,r) = zsplit0 a in | |
| 879 | if c=0 then (NDvd (abs i) r) else | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 880 | if c>0 then (NDvd (abs i) (CN 0 c r)) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 881 | else (NDvd (abs i) (CN 0 (- c) (Neg r)))))" | 
| 23274 | 882 | "zlfm (NOT (And p q)) = Or (zlfm (NOT p)) (zlfm (NOT q))" | 
| 883 | "zlfm (NOT (Or p q)) = And (zlfm (NOT p)) (zlfm (NOT q))" | |
| 884 | "zlfm (NOT (Imp p q)) = And (zlfm p) (zlfm (NOT q))" | |
| 885 | "zlfm (NOT (Iff p q)) = Or (And(zlfm p) (zlfm(NOT q))) (And (zlfm(NOT p)) (zlfm q))" | |
| 886 | "zlfm (NOT (NOT p)) = zlfm p" | |
| 887 | "zlfm (NOT T) = F" | |
| 888 | "zlfm (NOT F) = T" | |
| 889 | "zlfm (NOT (Lt a)) = zlfm (Ge a)" | |
| 890 | "zlfm (NOT (Le a)) = zlfm (Gt a)" | |
| 891 | "zlfm (NOT (Gt a)) = zlfm (Le a)" | |
| 892 | "zlfm (NOT (Ge a)) = zlfm (Lt a)" | |
| 893 | "zlfm (NOT (Eq a)) = zlfm (NEq a)" | |
| 894 | "zlfm (NOT (NEq a)) = zlfm (Eq a)" | |
| 895 | "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)" | |
| 896 | "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)" | |
| 897 | "zlfm (NOT (Closed P)) = NClosed P" | |
| 898 | "zlfm (NOT (NClosed P)) = Closed P" | |
| 899 | "zlfm p = p" (hints simp add: fmsize_pos) | |
| 900 | ||
| 901 | lemma zlfm_I: | |
| 902 | assumes qfp: "qfree p" | |
| 903 | shows "(Ifm bbs (i#bs) (zlfm p) = Ifm bbs (i# bs) p) \<and> iszlfm (zlfm p)" | |
| 904 | (is "(?I (?l p) = ?I p) \<and> ?L (?l p)") | |
| 905 | using qfp | |
| 906 | proof(induct p rule: zlfm.induct) | |
| 907 | case (5 a) | |
| 908 | let ?c = "fst (zsplit0 a)" | |
| 909 | let ?r = "snd (zsplit0 a)" | |
| 910 | have spl: "zsplit0 a = (?c,?r)" by simp | |
| 911 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 912 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto | 
| 23274 | 913 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 914 | from prems Ia nb show ?case | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 915 | apply (auto simp add: Let_def split_def ring_simps) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 916 | apply (cases "?r",auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 917 | apply (case_tac nat, auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 918 | done | 
| 23274 | 919 | next | 
| 920 | case (6 a) | |
| 921 | let ?c = "fst (zsplit0 a)" | |
| 922 | let ?r = "snd (zsplit0 a)" | |
| 923 | have spl: "zsplit0 a = (?c,?r)" by simp | |
| 924 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 925 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto | 
| 23274 | 926 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 927 | from prems Ia nb show ?case | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 928 | apply (auto simp add: Let_def split_def ring_simps) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 929 | apply (cases "?r",auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 930 | apply (case_tac nat, auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 931 | done | 
| 23274 | 932 | next | 
| 933 | case (7 a) | |
| 934 | let ?c = "fst (zsplit0 a)" | |
| 935 | let ?r = "snd (zsplit0 a)" | |
| 936 | have spl: "zsplit0 a = (?c,?r)" by simp | |
| 937 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 938 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto | 
| 23274 | 939 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 940 | from prems Ia nb show ?case | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 941 | apply (auto simp add: Let_def split_def ring_simps) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 942 | apply (cases "?r",auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 943 | apply (case_tac nat, auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 944 | done | 
| 23274 | 945 | next | 
| 946 | case (8 a) | |
| 947 | let ?c = "fst (zsplit0 a)" | |
| 948 | let ?r = "snd (zsplit0 a)" | |
| 949 | have spl: "zsplit0 a = (?c,?r)" by simp | |
| 950 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 951 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto | 
| 23274 | 952 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 953 | from prems Ia nb show ?case | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 954 | apply (auto simp add: Let_def split_def ring_simps) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 955 | apply (cases "?r",auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 956 | apply (case_tac nat, auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 957 | done | 
| 23274 | 958 | next | 
| 959 | case (9 a) | |
| 960 | let ?c = "fst (zsplit0 a)" | |
| 961 | let ?r = "snd (zsplit0 a)" | |
| 962 | have spl: "zsplit0 a = (?c,?r)" by simp | |
| 963 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 964 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto | 
| 23274 | 965 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 966 | from prems Ia nb show ?case | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 967 | apply (auto simp add: Let_def split_def ring_simps) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 968 | apply (cases "?r",auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 969 | apply (case_tac nat, auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 970 | done | 
| 23274 | 971 | next | 
| 972 | case (10 a) | |
| 973 | let ?c = "fst (zsplit0 a)" | |
| 974 | let ?r = "snd (zsplit0 a)" | |
| 975 | have spl: "zsplit0 a = (?c,?r)" by simp | |
| 976 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 977 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto | 
| 23274 | 978 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 979 | from prems Ia nb show ?case | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 980 | apply (auto simp add: Let_def split_def ring_simps) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 981 | apply (cases "?r",auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 982 | apply (case_tac nat, auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 983 | done | 
| 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: diff
changeset | 984 | next | 
| 23274 | 985 | case (11 j a) | 
| 986 | let ?c = "fst (zsplit0 a)" | |
| 987 | let ?r = "snd (zsplit0 a)" | |
| 988 | have spl: "zsplit0 a = (?c,?r)" by simp | |
| 989 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 990 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto | 
| 23274 | 991 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 992 | have "j=0 \<or> (j\<noteq>0 \<and> ?c = 0) \<or> (j\<noteq>0 \<and> ?c >0) \<or> (j\<noteq> 0 \<and> ?c<0)" by arith | |
| 993 | moreover | |
| 994 |   {assume "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) 
 | |
| 995 | hence ?case using prems by (simp del: zlfm.simps add: zdvd_0_left)} | |
| 996 | moreover | |
| 997 |   {assume "?c=0" and "j\<noteq>0" hence ?case 
 | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 998 | using zsplit0_I[OF spl, where x="i" and bs="bs"] zdvd_abs1[where i="j"] | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 999 | apply (auto simp add: Let_def split_def ring_simps) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1000 | apply (cases "?r",auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1001 | apply (case_tac nat, auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1002 | done} | 
| 23274 | 1003 | moreover | 
| 1004 |   {assume cp: "?c > 0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
 | |
| 1005 | by (simp add: nb Let_def split_def) | |
| 1006 | hence ?case using Ia cp jnz by (simp add: Let_def split_def | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1007 | zdvd_abs1[where i="j" and j="(?c*i) + ?N ?r", symmetric])} | 
| 23274 | 1008 | moreover | 
| 1009 |   {assume cn: "?c < 0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
 | |
| 1010 | by (simp add: nb Let_def split_def) | |
| 1011 | hence ?case using Ia cn jnz zdvd_zminus_iff[where m="abs j" and n="?c*i + ?N ?r" ] | |
| 1012 | by (simp add: Let_def split_def | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1013 | zdvd_abs1[where i="j" and j="(?c*i) + ?N ?r", symmetric])} | 
| 23274 | 1014 | ultimately show ?case by blast | 
| 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: diff
changeset | 1015 | next | 
| 23274 | 1016 | case (12 j a) | 
| 1017 | let ?c = "fst (zsplit0 a)" | |
| 1018 | let ?r = "snd (zsplit0 a)" | |
| 1019 | have spl: "zsplit0 a = (?c,?r)" by simp | |
| 1020 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1021 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto | 
| 23274 | 1022 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 1023 | have "j=0 \<or> (j\<noteq>0 \<and> ?c = 0) \<or> (j\<noteq>0 \<and> ?c >0) \<or> (j\<noteq> 0 \<and> ?c<0)" by arith | |
| 1024 | moreover | |
| 1025 |   {assume "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) 
 | |
| 1026 | hence ?case using prems by (simp del: zlfm.simps add: zdvd_0_left)} | |
| 1027 | moreover | |
| 1028 |   {assume "?c=0" and "j\<noteq>0" hence ?case 
 | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1029 | using zsplit0_I[OF spl, where x="i" and bs="bs"] zdvd_abs1[where i="j"] | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1030 | apply (auto simp add: Let_def split_def ring_simps) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1031 | apply (cases "?r",auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1032 | apply (case_tac nat, auto) | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1033 | done} | 
| 23274 | 1034 | moreover | 
| 1035 |   {assume cp: "?c > 0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
 | |
| 1036 | by (simp add: nb Let_def split_def) | |
| 1037 | hence ?case using Ia cp jnz by (simp add: Let_def split_def | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1038 | zdvd_abs1[where i="j" and j="(?c*i) + ?N ?r", symmetric])} | 
| 23274 | 1039 | moreover | 
| 1040 |   {assume cn: "?c < 0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" 
 | |
| 1041 | by (simp add: nb Let_def split_def) | |
| 1042 | hence ?case using Ia cn jnz zdvd_zminus_iff[where m="abs j" and n="?c*i + ?N ?r" ] | |
| 1043 | by (simp add: Let_def split_def | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1044 | zdvd_abs1[where i="j" and j="(?c*i) + ?N ?r", symmetric])} | 
| 23274 | 1045 | ultimately show ?case by blast | 
| 1046 | qed auto | |
| 1047 | ||
| 1048 | consts | |
| 1049 | plusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of +\<infinity>*) | |
| 1050 | minusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of -\<infinity>*) | |
| 1051 |   \<delta> :: "fm \<Rightarrow> int" (* Compute lcm {d| N\<^isup>?\<^isup> Dvd c*x+t \<in> p}*)
 | |
| 1052 | d\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* checks if a given l divides all the ds above*) | |
| 1053 | ||
| 1054 | recdef minusinf "measure size" | |
| 1055 | "minusinf (And p q) = And (minusinf p) (minusinf q)" | |
| 1056 | "minusinf (Or p q) = Or (minusinf p) (minusinf q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1057 | "minusinf (Eq (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1058 | "minusinf (NEq (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1059 | "minusinf (Lt (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1060 | "minusinf (Le (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1061 | "minusinf (Gt (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1062 | "minusinf (Ge (CN 0 c e)) = F" | 
| 23274 | 1063 | "minusinf p = p" | 
| 1064 | ||
| 1065 | lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)" | |
| 1066 | by (induct p rule: minusinf.induct, auto) | |
| 1067 | ||
| 1068 | recdef plusinf "measure size" | |
| 1069 | "plusinf (And p q) = And (plusinf p) (plusinf q)" | |
| 1070 | "plusinf (Or p q) = Or (plusinf p) (plusinf q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1071 | "plusinf (Eq (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1072 | "plusinf (NEq (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1073 | "plusinf (Lt (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1074 | "plusinf (Le (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1075 | "plusinf (Gt (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1076 | "plusinf (Ge (CN 0 c e)) = T" | 
| 23274 | 1077 | "plusinf p = p" | 
| 1078 | ||
| 1079 | recdef \<delta> "measure size" | |
| 27556 | 1080 | "\<delta> (And p q) = zlcm (\<delta> p) (\<delta> q)" | 
| 1081 | "\<delta> (Or p q) = zlcm (\<delta> p) (\<delta> q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1082 | "\<delta> (Dvd i (CN 0 c e)) = i" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1083 | "\<delta> (NDvd i (CN 0 c e)) = i" | 
| 23274 | 1084 | "\<delta> p = 1" | 
| 1085 | ||
| 1086 | recdef d\<delta> "measure size" | |
| 1087 | "d\<delta> (And p q) = (\<lambda> d. d\<delta> p d \<and> d\<delta> q d)" | |
| 1088 | "d\<delta> (Or p q) = (\<lambda> d. d\<delta> p d \<and> d\<delta> q d)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1089 | "d\<delta> (Dvd i (CN 0 c e)) = (\<lambda> d. i dvd d)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1090 | "d\<delta> (NDvd i (CN 0 c e)) = (\<lambda> d. i dvd d)" | 
| 23274 | 1091 | "d\<delta> p = (\<lambda> d. True)" | 
| 1092 | ||
| 1093 | lemma delta_mono: | |
| 1094 | assumes lin: "iszlfm p" | |
| 1095 | and d: "d dvd d'" | |
| 1096 | and ad: "d\<delta> p d" | |
| 1097 | shows "d\<delta> p d'" | |
| 1098 | using lin ad d | |
| 1099 | proof(induct p rule: iszlfm.induct) | |
| 1100 | case (9 i c e) thus ?case using d | |
| 1101 | by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) | |
| 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: diff
changeset | 1102 | next | 
| 23274 | 1103 | case (10 i c e) thus ?case using d | 
| 1104 | by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) | |
| 1105 | qed simp_all | |
| 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: diff
changeset | 1106 | |
| 23274 | 1107 | lemma \<delta> : assumes lin:"iszlfm p" | 
| 1108 | shows "d\<delta> p (\<delta> p) \<and> \<delta> p >0" | |
| 1109 | using lin | |
| 1110 | proof (induct p rule: iszlfm.induct) | |
| 1111 | case (1 p q) | |
| 1112 | let ?d = "\<delta> (And p q)" | |
| 27556 | 1113 | from prems zlcm_pos have dp: "?d >0" by simp | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1114 | have d1: "\<delta> p dvd \<delta> (And p q)" using prems by simp | 
| 27556 | 1115 | hence th: "d\<delta> p ?d" using delta_mono prems(3-4) by(simp del:dvd_zlcm_self1) | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1116 | have "\<delta> q dvd \<delta> (And p q)" using prems by simp | 
| 27556 | 1117 | hence th': "d\<delta> q ?d" using delta_mono prems by(simp del:dvd_zlcm_self2) | 
| 23984 | 1118 | from th th' dp show ?case by simp | 
| 23274 | 1119 | next | 
| 1120 | case (2 p q) | |
| 1121 | let ?d = "\<delta> (And p q)" | |
| 27556 | 1122 | from prems zlcm_pos have dp: "?d >0" by simp | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1123 | have "\<delta> p dvd \<delta> (And p q)" using prems by simp | 
| 27556 | 1124 | hence th: "d\<delta> p ?d" using delta_mono prems by(simp del:dvd_zlcm_self1) | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1125 | have "\<delta> q dvd \<delta> (And p q)" using prems by simp | 
| 27556 | 1126 | hence th': "d\<delta> q ?d" using delta_mono prems by(simp del:dvd_zlcm_self2) | 
| 23984 | 1127 | from th th' dp show ?case by simp | 
| 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: diff
changeset | 1128 | qed simp_all | 
| 
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: diff
changeset | 1129 | |
| 
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: diff
changeset | 1130 | |
| 23274 | 1131 | consts | 
| 1132 | a\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coeffitients of a formula *) | |
| 1133 | d\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*) | |
| 1134 | \<zeta> :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*) | |
| 1135 | \<beta> :: "fm \<Rightarrow> num list" | |
| 1136 | \<alpha> :: "fm \<Rightarrow> num list" | |
| 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: diff
changeset | 1137 | |
| 23274 | 1138 | recdef a\<beta> "measure size" | 
| 1139 | "a\<beta> (And p q) = (\<lambda> k. And (a\<beta> p k) (a\<beta> q k))" | |
| 1140 | "a\<beta> (Or p q) = (\<lambda> k. Or (a\<beta> p k) (a\<beta> q k))" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1141 | "a\<beta> (Eq (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1142 | "a\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1143 | "a\<beta> (Lt (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1144 | "a\<beta> (Le (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1145 | "a\<beta> (Gt (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1146 | "a\<beta> (Ge (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1147 | "a\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1148 | "a\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" | 
| 23274 | 1149 | "a\<beta> p = (\<lambda> k. p)" | 
| 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: diff
changeset | 1150 | |
| 23274 | 1151 | recdef d\<beta> "measure size" | 
| 1152 | "d\<beta> (And p q) = (\<lambda> k. (d\<beta> p k) \<and> (d\<beta> q k))" | |
| 1153 | "d\<beta> (Or p q) = (\<lambda> k. (d\<beta> p k) \<and> (d\<beta> q k))" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1154 | "d\<beta> (Eq (CN 0 c e)) = (\<lambda> k. c dvd k)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1155 | "d\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1156 | "d\<beta> (Lt (CN 0 c e)) = (\<lambda> k. c dvd k)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1157 | "d\<beta> (Le (CN 0 c e)) = (\<lambda> k. c dvd k)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1158 | "d\<beta> (Gt (CN 0 c e)) = (\<lambda> k. c dvd k)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1159 | "d\<beta> (Ge (CN 0 c e)) = (\<lambda> k. c dvd k)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1160 | "d\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1161 | "d\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)" | 
| 23274 | 1162 | "d\<beta> p = (\<lambda> k. True)" | 
| 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: diff
changeset | 1163 | |
| 23274 | 1164 | recdef \<zeta> "measure size" | 
| 27556 | 1165 | "\<zeta> (And p q) = zlcm (\<zeta> p) (\<zeta> q)" | 
| 1166 | "\<zeta> (Or p q) = zlcm (\<zeta> p) (\<zeta> q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1167 | "\<zeta> (Eq (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1168 | "\<zeta> (NEq (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1169 | "\<zeta> (Lt (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1170 | "\<zeta> (Le (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1171 | "\<zeta> (Gt (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1172 | "\<zeta> (Ge (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1173 | "\<zeta> (Dvd i (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1174 | "\<zeta> (NDvd i (CN 0 c e))= c" | 
| 23274 | 1175 | "\<zeta> p = 1" | 
| 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: diff
changeset | 1176 | |
| 23274 | 1177 | recdef \<beta> "measure size" | 
| 1178 | "\<beta> (And p q) = (\<beta> p @ \<beta> q)" | |
| 1179 | "\<beta> (Or p q) = (\<beta> p @ \<beta> q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1180 | "\<beta> (Eq (CN 0 c e)) = [Sub (C -1) e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1181 | "\<beta> (NEq (CN 0 c e)) = [Neg e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1182 | "\<beta> (Lt (CN 0 c e)) = []" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1183 | "\<beta> (Le (CN 0 c e)) = []" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1184 | "\<beta> (Gt (CN 0 c e)) = [Neg e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1185 | "\<beta> (Ge (CN 0 c e)) = [Sub (C -1) e]" | 
| 23274 | 1186 | "\<beta> p = []" | 
| 19736 | 1187 | |
| 23274 | 1188 | recdef \<alpha> "measure size" | 
| 1189 | "\<alpha> (And p q) = (\<alpha> p @ \<alpha> q)" | |
| 1190 | "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1191 | "\<alpha> (Eq (CN 0 c e)) = [Add (C -1) e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1192 | "\<alpha> (NEq (CN 0 c e)) = [e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1193 | "\<alpha> (Lt (CN 0 c e)) = [e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1194 | "\<alpha> (Le (CN 0 c e)) = [Add (C -1) e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1195 | "\<alpha> (Gt (CN 0 c e)) = []" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1196 | "\<alpha> (Ge (CN 0 c e)) = []" | 
| 23274 | 1197 | "\<alpha> p = []" | 
| 1198 | consts mirror :: "fm \<Rightarrow> fm" | |
| 1199 | recdef mirror "measure size" | |
| 1200 | "mirror (And p q) = And (mirror p) (mirror q)" | |
| 1201 | "mirror (Or p q) = Or (mirror p) (mirror q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1202 | "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1203 | "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1204 | "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1205 | "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1206 | "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1207 | "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1208 | "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1209 | "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))" | 
| 23274 | 1210 | "mirror p = p" | 
| 1211 | (* Lemmas for the correctness of \<sigma>\<rho> *) | |
| 1212 | lemma dvd1_eq1: "x >0 \<Longrightarrow> (x::int) dvd 1 = (x = 1)" | |
| 1213 | by auto | |
| 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: diff
changeset | 1214 | |
| 23274 | 1215 | lemma minusinf_inf: | 
| 1216 | assumes linp: "iszlfm p" | |
| 1217 | and u: "d\<beta> p 1" | |
| 1218 | shows "\<exists> (z::int). \<forall> x < z. Ifm bbs (x#bs) (minusinf p) = Ifm bbs (x#bs) p" | |
| 1219 | (is "?P p" is "\<exists> (z::int). \<forall> x < z. ?I x (?M p) = ?I x p") | |
| 1220 | using linp u | |
| 1221 | proof (induct p rule: minusinf.induct) | |
| 1222 | case (1 p q) thus ?case | |
| 1223 | by (auto simp add: dvd1_eq1) (rule_tac x="min z za" in exI,simp) | |
| 1224 | next | |
| 1225 | case (2 p q) thus ?case | |
| 1226 | by (auto simp add: dvd1_eq1) (rule_tac x="min z za" in exI,simp) | |
| 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: diff
changeset | 1227 | next | 
| 23274 | 1228 | case (3 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ | 
| 26934 | 1229 | fix a | 
| 1230 | from 3 have "\<forall> x<(- Inum (a#bs) e). c*x + Inum (x#bs) e \<noteq> 0" | |
| 23274 | 1231 | proof(clarsimp) | 
| 1232 | fix x assume "x < (- Inum (a#bs) e)" and"x + Inum (x#bs) e = 0" | |
| 1233 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | |
| 1234 | show "False" by simp | |
| 1235 | qed | |
| 1236 | thus ?case by auto | |
| 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: diff
changeset | 1237 | next | 
| 23274 | 1238 | case (4 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ | 
| 26934 | 1239 | fix a | 
| 1240 | from 4 have "\<forall> x<(- Inum (a#bs) e). c*x + Inum (x#bs) e \<noteq> 0" | |
| 23274 | 1241 | proof(clarsimp) | 
| 1242 | fix x assume "x < (- Inum (a#bs) e)" and"x + Inum (x#bs) e = 0" | |
| 1243 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | |
| 1244 | show "False" by simp | |
| 1245 | qed | |
| 1246 | thus ?case by auto | |
| 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: diff
changeset | 1247 | next | 
| 23274 | 1248 | case (5 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ | 
| 26934 | 1249 | fix a | 
| 1250 | from 5 have "\<forall> x<(- Inum (a#bs) e). c*x + Inum (x#bs) e < 0" | |
| 23274 | 1251 | proof(clarsimp) | 
| 1252 | fix x assume "x < (- Inum (a#bs) e)" | |
| 1253 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | |
| 1254 | show "x + Inum (x#bs) e < 0" by simp | |
| 1255 | qed | |
| 1256 | thus ?case by auto | |
| 1257 | next | |
| 1258 | case (6 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ | |
| 26934 | 1259 | fix a | 
| 1260 | from 6 have "\<forall> x<(- Inum (a#bs) e). c*x + Inum (x#bs) e \<le> 0" | |
| 23274 | 1261 | proof(clarsimp) | 
| 1262 | fix x assume "x < (- Inum (a#bs) e)" | |
| 1263 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | |
| 1264 | show "x + Inum (x#bs) e \<le> 0" by simp | |
| 1265 | qed | |
| 1266 | thus ?case by auto | |
| 1267 | next | |
| 1268 | case (7 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ | |
| 26934 | 1269 | fix a | 
| 1270 | from 7 have "\<forall> x<(- Inum (a#bs) e). \<not> (c*x + Inum (x#bs) e > 0)" | |
| 23274 | 1271 | proof(clarsimp) | 
| 1272 | fix x assume "x < (- Inum (a#bs) e)" and"x + Inum (x#bs) e > 0" | |
| 1273 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | |
| 1274 | show "False" by simp | |
| 1275 | qed | |
| 1276 | thus ?case by auto | |
| 1277 | next | |
| 1278 | case (8 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ | |
| 26934 | 1279 | fix a | 
| 1280 | from 8 have "\<forall> x<(- Inum (a#bs) e). \<not> (c*x + Inum (x#bs) e \<ge> 0)" | |
| 23274 | 1281 | proof(clarsimp) | 
| 1282 | fix x assume "x < (- Inum (a#bs) e)" and"x + Inum (x#bs) e \<ge> 0" | |
| 1283 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | |
| 1284 | show "False" by simp | |
| 1285 | qed | |
| 1286 | thus ?case by auto | |
| 1287 | qed auto | |
| 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: diff
changeset | 1288 | |
| 23274 | 1289 | lemma minusinf_repeats: | 
| 1290 | assumes d: "d\<delta> p d" and linp: "iszlfm p" | |
| 1291 | shows "Ifm bbs ((x - k*d)#bs) (minusinf p) = Ifm bbs (x #bs) (minusinf p)" | |
| 1292 | using linp d | |
| 1293 | proof(induct p rule: iszlfm.induct) | |
| 1294 | case (9 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ | |
| 1295 | hence "\<exists> k. d=i*k" by (simp add: dvd_def) | |
| 1296 | then obtain "di" where di_def: "d=i*di" by blast | |
| 1297 | show ?case | |
| 1298 | proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib, rule iffI) | |
| 1299 | assume | |
| 1300 | "i dvd c * x - c*(k*d) + Inum (x # bs) e" | |
| 1301 | (is "?ri dvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri dvd ?rt") | |
| 1302 | hence "\<exists> (l::int). ?rt = i * l" by (simp add: dvd_def) | |
| 1303 | hence "\<exists> (l::int). c*x+ ?I x e = i*l+c*(k * i*di)" | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1304 | by (simp add: ring_simps di_def) | 
| 23274 | 1305 | hence "\<exists> (l::int). c*x+ ?I x e = i*(l + c*k*di)" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1306 | by (simp add: ring_simps) | 
| 23274 | 1307 | hence "\<exists> (l::int). c*x+ ?I x e = i*l" by blast | 
| 1308 | thus "i dvd c*x + Inum (x # bs) e" by (simp add: dvd_def) | |
| 1309 | next | |
| 1310 | assume | |
| 1311 | "i dvd c*x + Inum (x # bs) e" (is "?ri dvd ?rc*?rx+?e") | |
| 1312 | hence "\<exists> (l::int). c*x+?e = i*l" by (simp add: dvd_def) | |
| 1313 | hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*d)" by simp | |
| 1314 | hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*i*di)" by (simp add: di_def) | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1315 | hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: ring_simps) | 
| 23274 | 1316 | hence "\<exists> (l::int). c*x - c * (k*d) +?e = i*l" | 
| 1317 | by blast | |
| 1318 | thus "i dvd c*x - c*(k*d) + Inum (x # bs) e" by (simp add: dvd_def) | |
| 1319 | qed | |
| 1320 | next | |
| 1321 | case (10 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ | |
| 1322 | hence "\<exists> k. d=i*k" by (simp add: dvd_def) | |
| 1323 | then obtain "di" where di_def: "d=i*di" by blast | |
| 1324 | show ?case | |
| 1325 | proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib, rule iffI) | |
| 1326 | assume | |
| 1327 | "i dvd c * x - c*(k*d) + Inum (x # bs) e" | |
| 1328 | (is "?ri dvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri dvd ?rt") | |
| 1329 | hence "\<exists> (l::int). ?rt = i * l" by (simp add: dvd_def) | |
| 1330 | hence "\<exists> (l::int). c*x+ ?I x e = i*l+c*(k * i*di)" | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1331 | by (simp add: ring_simps di_def) | 
| 23274 | 1332 | hence "\<exists> (l::int). c*x+ ?I x e = i*(l + c*k*di)" | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1333 | by (simp add: ring_simps) | 
| 23274 | 1334 | hence "\<exists> (l::int). c*x+ ?I x e = i*l" by blast | 
| 1335 | thus "i dvd c*x + Inum (x # bs) e" by (simp add: dvd_def) | |
| 1336 | next | |
| 1337 | assume | |
| 1338 | "i dvd c*x + Inum (x # bs) e" (is "?ri dvd ?rc*?rx+?e") | |
| 1339 | hence "\<exists> (l::int). c*x+?e = i*l" by (simp add: dvd_def) | |
| 1340 | hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*d)" by simp | |
| 1341 | hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*l - c*(k*i*di)" by (simp add: di_def) | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1342 | hence "\<exists> (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: ring_simps) | 
| 23274 | 1343 | hence "\<exists> (l::int). c*x - c * (k*d) +?e = i*l" | 
| 1344 | by blast | |
| 1345 | thus "i dvd c*x - c*(k*d) + Inum (x # bs) e" by (simp add: dvd_def) | |
| 1346 | qed | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1347 | qed (auto simp add: gr0_conv_Suc numbound0_I[where bs="bs" and b="x - k*d" and b'="x"]) | 
| 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: diff
changeset | 1348 | |
| 23274 | 1349 | lemma mirror\<alpha>\<beta>: | 
| 1350 | assumes lp: "iszlfm p" | |
| 1351 | shows "(Inum (i#bs)) ` set (\<alpha> p) = (Inum (i#bs)) ` set (\<beta> (mirror p))" | |
| 1352 | using lp | |
| 1353 | by (induct p rule: mirror.induct, auto) | |
| 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: diff
changeset | 1354 | |
| 23274 | 1355 | lemma mirror: | 
| 1356 | assumes lp: "iszlfm p" | |
| 1357 | shows "Ifm bbs (x#bs) (mirror p) = Ifm bbs ((- x)#bs) p" | |
| 1358 | using lp | |
| 1359 | proof(induct p rule: iszlfm.induct) | |
| 1360 | case (9 j c e) hence nb: "numbound0 e" by simp | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1361 | have "Ifm bbs (x#bs) (mirror (Dvd j (CN 0 c e))) = (j dvd c*x - Inum (x#bs) e)" (is "_ = (j dvd c*x - ?e)") by simp | 
| 23274 | 1362 | also have "\<dots> = (j dvd (- (c*x - ?e)))" | 
| 1363 | by (simp only: zdvd_zminus_iff) | |
| 1364 | also have "\<dots> = (j dvd (c* (- x)) + ?e)" | |
| 1365 | apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_def zadd_ac zminus_zadd_distrib) | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1366 | by (simp add: ring_simps) | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1367 | also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))" | 
| 23274 | 1368 | using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] | 
| 1369 | by simp | |
| 1370 | finally show ?case . | |
| 1371 | next | |
| 1372 | case (10 j c e) hence nb: "numbound0 e" by simp | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1373 | have "Ifm bbs (x#bs) (mirror (Dvd j (CN 0 c e))) = (j dvd c*x - Inum (x#bs) e)" (is "_ = (j dvd c*x - ?e)") by simp | 
| 23274 | 1374 | also have "\<dots> = (j dvd (- (c*x - ?e)))" | 
| 1375 | by (simp only: zdvd_zminus_iff) | |
| 1376 | also have "\<dots> = (j dvd (c* (- x)) + ?e)" | |
| 1377 | apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_def zadd_ac zminus_zadd_distrib) | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1378 | by (simp add: ring_simps) | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1379 | also have "\<dots> = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))" | 
| 23274 | 1380 | using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] | 
| 1381 | by simp | |
| 1382 | finally show ?case by simp | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1383 | qed (auto simp add: numbound0_I[where bs="bs" and b="x" and b'="- x"] gr0_conv_Suc) | 
| 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: diff
changeset | 1384 | |
| 23274 | 1385 | lemma mirror_l: "iszlfm p \<and> d\<beta> p 1 | 
| 1386 | \<Longrightarrow> iszlfm (mirror p) \<and> d\<beta> (mirror p) 1" | |
| 1387 | by (induct p rule: mirror.induct, auto) | |
| 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: diff
changeset | 1388 | |
| 23274 | 1389 | lemma mirror_\<delta>: "iszlfm p \<Longrightarrow> \<delta> (mirror p) = \<delta> p" | 
| 1390 | by (induct p rule: mirror.induct,auto) | |
| 1391 | ||
| 1392 | lemma \<beta>_numbound0: assumes lp: "iszlfm p" | |
| 1393 | shows "\<forall> b\<in> set (\<beta> p). numbound0 b" | |
| 1394 | using lp by (induct p rule: \<beta>.induct,auto) | |
| 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: diff
changeset | 1395 | |
| 23274 | 1396 | lemma d\<beta>_mono: | 
| 1397 | assumes linp: "iszlfm p" | |
| 1398 | and dr: "d\<beta> p l" | |
| 1399 | and d: "l dvd l'" | |
| 1400 | shows "d\<beta> p l'" | |
| 1401 | using dr linp zdvd_trans[where n="l" and k="l'", simplified d] | |
| 1402 | by (induct p rule: iszlfm.induct) simp_all | |
| 1403 | ||
| 1404 | lemma \<alpha>_l: assumes lp: "iszlfm p" | |
| 1405 | shows "\<forall> b\<in> set (\<alpha> p). numbound0 b" | |
| 1406 | using lp | |
| 1407 | by(induct p rule: \<alpha>.induct, auto) | |
| 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: diff
changeset | 1408 | |
| 23274 | 1409 | lemma \<zeta>: | 
| 1410 | assumes linp: "iszlfm p" | |
| 1411 | shows "\<zeta> p > 0 \<and> d\<beta> p (\<zeta> p)" | |
| 1412 | using linp | |
| 1413 | proof(induct p rule: iszlfm.induct) | |
| 1414 | case (1 p q) | |
| 27556 | 1415 | from prems have dl1: "\<zeta> p dvd zlcm (\<zeta> p) (\<zeta> q)" by simp | 
| 1416 | from prems have dl2: "\<zeta> q dvd zlcm (\<zeta> p) (\<zeta> q)" by simp | |
| 1417 | from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="zlcm (\<zeta> p) (\<zeta> q)"] | |
| 1418 | d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="zlcm (\<zeta> p) (\<zeta> q)"] | |
| 1419 | dl1 dl2 show ?case by (auto simp add: zlcm_pos) | |
| 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: diff
changeset | 1420 | next | 
| 23274 | 1421 | case (2 p q) | 
| 27556 | 1422 | from prems have dl1: "\<zeta> p dvd zlcm (\<zeta> p) (\<zeta> q)" by simp | 
| 1423 | from prems have dl2: "\<zeta> q dvd zlcm (\<zeta> p) (\<zeta> q)" by simp | |
| 1424 | from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="zlcm (\<zeta> p) (\<zeta> q)"] | |
| 1425 | d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="zlcm (\<zeta> p) (\<zeta> q)"] | |
| 1426 | dl1 dl2 show ?case by (auto simp add: zlcm_pos) | |
| 1427 | qed (auto simp add: zlcm_pos) | |
| 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: diff
changeset | 1428 | |
| 23274 | 1429 | lemma a\<beta>: assumes linp: "iszlfm p" and d: "d\<beta> p l" and lp: "l > 0" | 
| 1430 | shows "iszlfm (a\<beta> p l) \<and> d\<beta> (a\<beta> p l) 1 \<and> (Ifm bbs (l*x #bs) (a\<beta> p l) = Ifm bbs (x#bs) p)" | |
| 1431 | using linp d | |
| 1432 | proof (induct p rule: iszlfm.induct) | |
| 1433 | case (5 c e) hence cp: "c>0" and be: "numbound0 e" and d': "c dvd l" by simp+ | |
| 1434 | from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1435 | from cp have cnz: "c \<noteq> 0" by simp | |
| 1436 | have "c div c\<le> l div c" | |
| 1437 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1438 | then have ldcp:"0 < l div c" | |
| 1439 | by (simp add: zdiv_self[OF cnz]) | |
| 1440 | have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp | |
| 1441 | hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 1442 | by simp | |
| 1443 | hence "(l*x + (l div c) * Inum (x # bs) e < 0) = | |
| 1444 | ((c * (l div c)) * x + (l div c) * Inum (x # bs) e < 0)" | |
| 1445 | by simp | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1446 | also have "\<dots> = ((l div c) * (c*x + Inum (x # bs) e) < (l div c) * 0)" by (simp add: ring_simps) | 
| 23274 | 1447 | also have "\<dots> = (c*x + Inum (x # bs) e < 0)" | 
| 1448 | using mult_less_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp by simp | |
| 1449 | finally show ?case using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] be by simp | |
| 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: diff
changeset | 1450 | next | 
| 23274 | 1451 | case (6 c e) hence cp: "c>0" and be: "numbound0 e" and d': "c dvd l" by simp+ | 
| 1452 | from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1453 | from cp have cnz: "c \<noteq> 0" by simp | |
| 1454 | have "c div c\<le> l div c" | |
| 1455 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1456 | then have ldcp:"0 < l div c" | |
| 1457 | by (simp add: zdiv_self[OF cnz]) | |
| 1458 | have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp | |
| 1459 | hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 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: diff
changeset | 1460 | by simp | 
| 23274 | 1461 | hence "(l*x + (l div c) * Inum (x# bs) e \<le> 0) = | 
| 1462 | ((c * (l div c)) * x + (l div c) * Inum (x # bs) e \<le> 0)" | |
| 1463 | by simp | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1464 | also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) \<le> ((l div c)) * 0)" by (simp add: ring_simps) | 
| 23274 | 1465 | also have "\<dots> = (c*x + Inum (x # bs) e \<le> 0)" | 
| 1466 | using mult_le_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp by simp | |
| 1467 | finally show ?case using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] be by simp | |
| 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: diff
changeset | 1468 | next | 
| 23274 | 1469 | case (7 c e) hence cp: "c>0" and be: "numbound0 e" and d': "c dvd l" by simp+ | 
| 1470 | from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1471 | from cp have cnz: "c \<noteq> 0" by simp | |
| 1472 | have "c div c\<le> l div c" | |
| 1473 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1474 | then have ldcp:"0 < l div c" | |
| 1475 | by (simp add: zdiv_self[OF cnz]) | |
| 1476 | have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp | |
| 1477 | hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 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: diff
changeset | 1478 | by simp | 
| 23274 | 1479 | hence "(l*x + (l div c)* Inum (x # bs) e > 0) = | 
| 1480 | ((c * (l div c)) * x + (l div c) * Inum (x # bs) e > 0)" | |
| 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: diff
changeset | 1481 | by simp | 
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1482 | also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) > ((l div c)) * 0)" by (simp add: ring_simps) | 
| 23274 | 1483 | also have "\<dots> = (c * x + Inum (x # bs) e > 0)" | 
| 1484 | using zero_less_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp | |
| 1485 | finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be by simp | |
| 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: diff
changeset | 1486 | next | 
| 23274 | 1487 | case (8 c e) hence cp: "c>0" and be: "numbound0 e" and d': "c dvd l" by simp+ | 
| 1488 | from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1489 | from cp have cnz: "c \<noteq> 0" by simp | |
| 1490 | have "c div c\<le> l div c" | |
| 1491 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1492 | then have ldcp:"0 < l div c" | |
| 1493 | by (simp add: zdiv_self[OF cnz]) | |
| 1494 | have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp | |
| 1495 | hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 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: diff
changeset | 1496 | by simp | 
| 23274 | 1497 | hence "(l*x + (l div c)* Inum (x # bs) e \<ge> 0) = | 
| 1498 | ((c*(l div c))*x + (l div c)* Inum (x # bs) e \<ge> 0)" | |
| 1499 | by simp | |
| 1500 | also have "\<dots> = ((l div c)*(c*x + Inum (x # bs) e) \<ge> ((l div c)) * 0)" | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1501 | by (simp add: ring_simps) | 
| 23274 | 1502 | also have "\<dots> = (c*x + Inum (x # bs) e \<ge> 0)" using ldcp | 
| 1503 | zero_le_mult_iff [where a="l div c" and b="c*x + Inum (x # bs) e"] by simp | |
| 1504 | finally show ?case using be numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] | |
| 1505 | by simp | |
| 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: diff
changeset | 1506 | next | 
| 23274 | 1507 | case (3 c e) hence cp: "c>0" and be: "numbound0 e" and d': "c dvd l" by simp+ | 
| 1508 | from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1509 | from cp have cnz: "c \<noteq> 0" by simp | |
| 1510 | have "c div c\<le> l div c" | |
| 1511 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1512 | then have ldcp:"0 < l div c" | |
| 1513 | by (simp add: zdiv_self[OF cnz]) | |
| 1514 | have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp | |
| 1515 | hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 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: diff
changeset | 1516 | by simp | 
| 23274 | 1517 | hence "(l * x + (l div c) * Inum (x # bs) e = 0) = | 
| 1518 | ((c * (l div c)) * x + (l div c) * Inum (x # bs) e = 0)" | |
| 1519 | by simp | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1520 | also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) = ((l div c)) * 0)" by (simp add: ring_simps) | 
| 23274 | 1521 | also have "\<dots> = (c * x + Inum (x # bs) e = 0)" | 
| 1522 | using mult_eq_0_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp | |
| 1523 | finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be by simp | |
| 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: diff
changeset | 1524 | next | 
| 23274 | 1525 | case (4 c e) hence cp: "c>0" and be: "numbound0 e" and d': "c dvd l" by simp+ | 
| 1526 | from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1527 | from cp have cnz: "c \<noteq> 0" by simp | |
| 1528 | have "c div c\<le> l div c" | |
| 1529 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1530 | then have ldcp:"0 < l div c" | |
| 1531 | by (simp add: zdiv_self[OF cnz]) | |
| 1532 | have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp | |
| 1533 | hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 1534 | by simp | |
| 1535 | hence "(l * x + (l div c) * Inum (x # bs) e \<noteq> 0) = | |
| 1536 | ((c * (l div c)) * x + (l div c) * Inum (x # bs) e \<noteq> 0)" | |
| 1537 | by simp | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1538 | also have "\<dots> = ((l div c) * (c * x + Inum (x # bs) e) \<noteq> ((l div c)) * 0)" by (simp add: ring_simps) | 
| 23274 | 1539 | also have "\<dots> = (c * x + Inum (x # bs) e \<noteq> 0)" | 
| 1540 | using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp | |
| 1541 | finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be by simp | |
| 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: diff
changeset | 1542 | next | 
| 23274 | 1543 | case (9 j c e) hence cp: "c>0" and be: "numbound0 e" and jp: "j > 0" and d': "c dvd l" by simp+ | 
| 1544 | from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1545 | from cp have cnz: "c \<noteq> 0" by simp | |
| 1546 | have "c div c\<le> l div c" | |
| 1547 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1548 | then have ldcp:"0 < l div c" | |
| 1549 | by (simp add: zdiv_self[OF cnz]) | |
| 1550 | have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp | |
| 1551 | hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 1552 | by simp | |
| 1553 | hence "(\<exists> (k::int). l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) = (\<exists> (k::int). (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)" by simp | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1554 | also have "\<dots> = (\<exists> (k::int). (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c)*0)" by (simp add: ring_simps) | 
| 26934 | 1555 | also fix k have "\<dots> = (\<exists> (k::int). c * x + Inum (x # bs) e - j * k = 0)" | 
| 23274 | 1556 | using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k"] ldcp by simp | 
| 1557 | also have "\<dots> = (\<exists> (k::int). c * x + Inum (x # bs) e = j * k)" by simp | |
| 1558 | finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be mult_strict_mono[OF ldcp jp ldcp ] by (simp add: dvd_def) | |
| 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: diff
changeset | 1559 | next | 
| 23274 | 1560 | case (10 j c e) hence cp: "c>0" and be: "numbound0 e" and jp: "j > 0" and d': "c dvd l" by simp+ | 
| 1561 | from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1562 | from cp have cnz: "c \<noteq> 0" by simp | |
| 1563 | have "c div c\<le> l div c" | |
| 1564 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1565 | then have ldcp:"0 < l div c" | |
| 1566 | by (simp add: zdiv_self[OF cnz]) | |
| 1567 | have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp | |
| 1568 | hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 1569 | by simp | |
| 1570 | hence "(\<exists> (k::int). l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) = (\<exists> (k::int). (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)" by simp | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1571 | also have "\<dots> = (\<exists> (k::int). (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c)*0)" by (simp add: ring_simps) | 
| 26934 | 1572 | also fix k have "\<dots> = (\<exists> (k::int). c * x + Inum (x # bs) e - j * k = 0)" | 
| 23274 | 1573 | using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k"] ldcp by simp | 
| 1574 | also have "\<dots> = (\<exists> (k::int). c * x + Inum (x # bs) e = j * k)" by simp | |
| 1575 | finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be mult_strict_mono[OF ldcp jp ldcp ] by (simp add: dvd_def) | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1576 | qed (auto simp add: gr0_conv_Suc numbound0_I[where bs="bs" and b="(l * x)" and b'="x"]) | 
| 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: diff
changeset | 1577 | |
| 23274 | 1578 | lemma a\<beta>_ex: assumes linp: "iszlfm p" and d: "d\<beta> p l" and lp: "l>0" | 
| 1579 | shows "(\<exists> x. l dvd x \<and> Ifm bbs (x #bs) (a\<beta> p l)) = (\<exists> (x::int). Ifm bbs (x#bs) p)" | |
| 1580 | (is "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> x. ?P' x)") | |
| 1581 | proof- | |
| 1582 | have "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> (x::int). ?P (l*x))" | |
| 1583 | using unity_coeff_ex[where l="l" and P="?P", simplified] by simp | |
| 1584 | also have "\<dots> = (\<exists> (x::int). ?P' x)" using a\<beta>[OF linp d lp] by simp | |
| 1585 | finally show ?thesis . | |
| 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: diff
changeset | 1586 | qed | 
| 
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: diff
changeset | 1587 | |
| 23274 | 1588 | lemma \<beta>: | 
| 1589 | assumes lp: "iszlfm p" | |
| 1590 | and u: "d\<beta> p 1" | |
| 1591 | and d: "d\<delta> p d" | |
| 1592 | and dp: "d > 0" | |
| 1593 |   and nob: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). x = b + j)"
 | |
| 1594 | and p: "Ifm bbs (x#bs) p" (is "?P x") | |
| 1595 | shows "?P (x - d)" | |
| 1596 | using lp u d dp nob p | |
| 1597 | proof(induct p rule: iszlfm.induct) | |
| 1598 | case (5 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ | |
| 1599 | with dp p c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] prems | |
| 1600 | show ?case by simp | |
| 1601 | next | |
| 1602 | case (6 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ | |
| 1603 | with dp p c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] prems | |
| 1604 | show ?case by simp | |
| 1605 | next | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1606 | case (7 c e) hence p: "Ifm bbs (x #bs) (Gt (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ | 
| 23274 | 1607 | let ?e = "Inum (x # bs) e" | 
| 1608 |     {assume "(x-d) +?e > 0" hence ?case using c1 
 | |
| 1609 | numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] by simp} | |
| 1610 | moreover | |
| 1611 |     {assume H: "\<not> (x-d) + ?e > 0" 
 | |
| 1612 | let ?v="Neg e" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1613 | have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" by simp | 
| 23274 | 1614 | from prems(11)[simplified simp_thms Inum.simps \<beta>.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]] | 
| 1615 |       have nob: "\<not> (\<exists> j\<in> {1 ..d}. x =  - ?e + j)" by auto 
 | |
| 1616 | from H p have "x + ?e > 0 \<and> x + ?e \<le> d" by (simp add: c1) | |
| 1617 | hence "x + ?e \<ge> 1 \<and> x + ?e \<le> d" by simp | |
| 1618 |       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + ?e" by simp
 | |
| 1619 |       hence "\<exists> (j::int) \<in> {1 .. d}. x = (- ?e + j)" 
 | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1620 | by (simp add: ring_simps) | 
| 23274 | 1621 | with nob have ?case by auto} | 
| 1622 | ultimately show ?case by blast | |
| 1623 | next | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1624 | case (8 c e) hence p: "Ifm bbs (x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" | 
| 23274 | 1625 | using dvd1_eq1[where x="c"] by simp+ | 
| 1626 | let ?e = "Inum (x # bs) e" | |
| 1627 |     {assume "(x-d) +?e \<ge> 0" hence ?case using  c1 
 | |
| 1628 | numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] | |
| 1629 | by simp} | |
| 1630 | moreover | |
| 1631 |     {assume H: "\<not> (x-d) + ?e \<ge> 0" 
 | |
| 1632 | let ?v="Sub (C -1) e" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1633 | have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" by simp | 
| 23274 | 1634 | from prems(11)[simplified simp_thms Inum.simps \<beta>.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]] | 
| 1635 |       have nob: "\<not> (\<exists> j\<in> {1 ..d}. x =  - ?e - 1 + j)" by auto 
 | |
| 1636 | from H p have "x + ?e \<ge> 0 \<and> x + ?e < d" by (simp add: c1) | |
| 1637 | hence "x + ?e +1 \<ge> 1 \<and> x + ?e + 1 \<le> d" by simp | |
| 1638 |       hence "\<exists> (j::int) \<in> {1 .. d}. j = x + ?e + 1" by simp
 | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1639 |       hence "\<exists> (j::int) \<in> {1 .. d}. x= - ?e - 1 + j" by (simp add: ring_simps)
 | 
| 23274 | 1640 | with nob have ?case by simp } | 
| 1641 | ultimately show ?case by blast | |
| 1642 | next | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1643 | case (3 c e) hence p: "Ifm bbs (x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ | 
| 23274 | 1644 | let ?e = "Inum (x # bs) e" | 
| 1645 | let ?v="(Sub (C -1) e)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1646 | have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp | 
| 23274 | 1647 | from p have "x= - ?e" by (simp add: c1) with prems(11) show ?case using dp | 
| 1648 | by simp (erule ballE[where x="1"], | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 1649 | simp_all add:ring_simps numbound0_I[OF bn,where b="x"and b'="a"and bs="bs"]) | 
| 23274 | 1650 | next | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1651 | case (4 c e)hence p: "Ifm bbs (x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ | 
| 23274 | 1652 | let ?e = "Inum (x # bs) e" | 
| 1653 | let ?v="Neg e" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1654 | have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))" by simp | 
| 23274 | 1655 |     {assume "x - d + Inum (((x -d)) # bs) e \<noteq> 0" 
 | 
| 1656 | hence ?case by (simp add: c1)} | |
| 1657 | moreover | |
| 1658 |     {assume H: "x - d + Inum (((x -d)) # bs) e = 0"
 | |
| 1659 | hence "x = - Inum (((x -d)) # bs) e + d" by simp | |
| 1660 | hence "x = - Inum (a # bs) e + d" | |
| 1661 | by (simp add: numbound0_I[OF bn,where b="x - d"and b'="a"and bs="bs"]) | |
| 1662 | with prems(11) have ?case using dp by simp} | |
| 1663 | ultimately show ?case by blast | |
| 1664 | next | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1665 | case (9 j c e) hence p: "Ifm bbs (x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ | 
| 23274 | 1666 | let ?e = "Inum (x # bs) e" | 
| 1667 | from prems have id: "j dvd d" by simp | |
| 1668 | from c1 have "?p x = (j dvd (x+ ?e))" by simp | |
| 1669 | also have "\<dots> = (j dvd x - d + ?e)" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1670 | using zdvd_period[OF id, where x="x" and c="-1" and t="?e"] by simp | 
| 23274 | 1671 | finally show ?case | 
| 1672 | using numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] c1 p by simp | |
| 1673 | next | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1674 | case (10 j c e) hence p: "Ifm bbs (x #bs) (NDvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ | 
| 23274 | 1675 | let ?e = "Inum (x # bs) e" | 
| 1676 | from prems have id: "j dvd d" by simp | |
| 1677 | from c1 have "?p x = (\<not> j dvd (x+ ?e))" by simp | |
| 1678 | also have "\<dots> = (\<not> j dvd x - d + ?e)" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1679 | using zdvd_period[OF id, where x="x" and c="-1" and t="?e"] by simp | 
| 23274 | 1680 | finally show ?case using numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] c1 p by simp | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1681 | qed (auto simp add: numbound0_I[where bs="bs" and b="(x - d)" and b'="x"] gr0_conv_Suc) | 
| 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: diff
changeset | 1682 | |
| 23274 | 1683 | lemma \<beta>': | 
| 1684 | assumes lp: "iszlfm p" | |
| 1685 | and u: "d\<beta> p 1" | |
| 1686 | and d: "d\<delta> p d" | |
| 1687 | and dp: "d > 0" | |
| 1688 |   shows "\<forall> x. \<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> set(\<beta> p). Ifm bbs ((Inum (a#bs) b + j) #bs) p) \<longrightarrow> Ifm bbs (x#bs) p \<longrightarrow> Ifm bbs ((x - d)#bs) p" (is "\<forall> x. ?b \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)")
 | |
| 1689 | proof(clarify) | |
| 1690 | fix x | |
| 1691 | assume nb:"?b" and px: "?P x" | |
| 1692 |   hence nb2: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). x = b + j)"
 | |
| 1693 | by auto | |
| 1694 | from \<beta>[OF lp u d dp nb2 px] show "?P (x -d )" . | |
| 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: diff
changeset | 1695 | qed | 
| 23315 | 1696 | lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x)) | 
| 1697 | ==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) 
 | |
| 1698 | ==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D)))) | |
| 1699 | ==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))"
 | |
| 1700 | apply(rule iffI) | |
| 1701 | prefer 2 | |
| 1702 | apply(drule minusinfinity) | |
| 1703 | apply assumption+ | |
| 1704 | apply(fastsimp) | |
| 1705 | apply clarsimp | |
| 1706 | apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)") | |
| 1707 | apply(frule_tac x = x and z=z in decr_lemma) | |
| 1708 | apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)") | |
| 1709 | prefer 2 | |
| 1710 | apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)") | |
| 1711 | prefer 2 apply arith | |
| 1712 | apply fastsimp | |
| 1713 | apply(drule (1) periodic_finite_ex) | |
| 1714 | apply blast | |
| 1715 | apply(blast dest:decr_mult_lemma) | |
| 1716 | done | |
| 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: diff
changeset | 1717 | |
| 23274 | 1718 | theorem cp_thm: | 
| 1719 | assumes lp: "iszlfm p" | |
| 1720 | and u: "d\<beta> p 1" | |
| 1721 | and d: "d\<delta> p d" | |
| 1722 | and dp: "d > 0" | |
| 1723 |   shows "(\<exists> (x::int). Ifm bbs (x #bs) p) = (\<exists> j\<in> {1.. d}. Ifm bbs (j #bs) (minusinf p) \<or> (\<exists> b \<in> set (\<beta> p). Ifm bbs ((Inum (i#bs) b + j) #bs) p))"
 | |
| 1724 | (is "(\<exists> (x::int). ?P (x)) = (\<exists> j\<in> ?D. ?M j \<or> (\<exists> b\<in> ?B. ?P (?I b + j)))") | |
| 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: diff
changeset | 1725 | proof- | 
| 23274 | 1726 | from minusinf_inf[OF lp u] | 
| 1727 | have th: "\<exists>(z::int). \<forall>x<z. ?P (x) = ?M x" by blast | |
| 1728 |   let ?B' = "{?I b | b. b\<in> ?B}"
 | |
| 1729 | have BB': "(\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (?I b +j)) = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (b + j))" by auto | |
| 1730 | hence th2: "\<forall> x. \<not> (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P ((b + j))) \<longrightarrow> ?P (x) \<longrightarrow> ?P ((x - d))" | |
| 1731 | using \<beta>'[OF lp u d dp, where a="i" and bbs = "bbs"] by blast | |
| 1732 | from minusinf_repeats[OF d lp] | |
| 1733 | have th3: "\<forall> x k. ?M x = ?M (x-k*d)" by simp | |
| 1734 | from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast | |
| 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: diff
changeset | 1735 | qed | 
| 
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: diff
changeset | 1736 | |
| 23274 | 1737 | (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *) | 
| 1738 | lemma mirror_ex: | |
| 1739 | assumes lp: "iszlfm p" | |
| 1740 | shows "(\<exists> x. Ifm bbs (x#bs) (mirror p)) = (\<exists> x. Ifm bbs (x#bs) p)" | |
| 1741 | (is "(\<exists> x. ?I x ?mp) = (\<exists> x. ?I x p)") | |
| 1742 | proof(auto) | |
| 1743 | fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast | |
| 1744 | thus "\<exists> x. ?I x p" by blast | |
| 1745 | next | |
| 1746 | fix x assume "?I x p" hence "?I (- x) ?mp" | |
| 1747 | using mirror[OF lp, where x="- x", symmetric] by auto | |
| 1748 | thus "\<exists> x. ?I x ?mp" by blast | |
| 1749 | qed | |
| 24349 | 1750 | |
| 1751 | ||
| 23274 | 1752 | lemma cp_thm': | 
| 1753 | assumes lp: "iszlfm p" | |
| 1754 | and up: "d\<beta> p 1" and dd: "d\<delta> p d" and dp: "d > 0" | |
| 1755 |   shows "(\<exists> x. Ifm bbs (x#bs) p) = ((\<exists> j\<in> {1 .. d}. Ifm bbs (j#bs) (minusinf p)) \<or> (\<exists> j\<in> {1.. d}. \<exists> b\<in> (Inum (i#bs)) ` set (\<beta> p). Ifm bbs ((b+j)#bs) p))"
 | |
| 1756 | using cp_thm[OF lp up dd dp,where i="i"] by auto | |
| 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: diff
changeset | 1757 | |
| 23274 | 1758 | constdefs unit:: "fm \<Rightarrow> fm \<times> num list \<times> int" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1759 | "unit p \<equiv> (let p' = zlfm p ; l = \<zeta> p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\<beta> p' l); d = \<delta> q; | 
| 23274 | 1760 | B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q)) | 
| 1761 | in if length B \<le> length a then (q,B,d) else (mirror q, a,d))" | |
| 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: diff
changeset | 1762 | |
| 23274 | 1763 | lemma unit: assumes qf: "qfree p" | 
| 1764 | shows "\<And> q B d. unit p = (q,B,d) \<Longrightarrow> ((\<exists> x. Ifm bbs (x#bs) p) = (\<exists> x. Ifm bbs (x#bs) q)) \<and> (Inum (i#bs)) ` set B = (Inum (i#bs)) ` set (\<beta> q) \<and> d\<beta> q 1 \<and> d\<delta> q d \<and> d >0 \<and> iszlfm q \<and> (\<forall> b\<in> set B. numbound0 b)" | |
| 1765 | proof- | |
| 1766 | fix q B d | |
| 1767 | assume qBd: "unit p = (q,B,d)" | |
| 1768 | let ?thes = "((\<exists> x. Ifm bbs (x#bs) p) = (\<exists> x. Ifm bbs (x#bs) q)) \<and> | |
| 1769 | Inum (i#bs) ` set B = Inum (i#bs) ` set (\<beta> q) \<and> | |
| 1770 | d\<beta> q 1 \<and> d\<delta> q d \<and> 0 < d \<and> iszlfm q \<and> (\<forall> b\<in> set B. numbound0 b)" | |
| 1771 | let ?I = "\<lambda> x p. Ifm bbs (x#bs) p" | |
| 1772 | let ?p' = "zlfm p" | |
| 1773 | let ?l = "\<zeta> ?p'" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1774 | let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a\<beta> ?p' ?l)" | 
| 23274 | 1775 | let ?d = "\<delta> ?q" | 
| 1776 | let ?B = "set (\<beta> ?q)" | |
| 1777 | let ?B'= "remdups (map simpnum (\<beta> ?q))" | |
| 1778 | let ?A = "set (\<alpha> ?q)" | |
| 1779 | let ?A'= "remdups (map simpnum (\<alpha> ?q))" | |
| 1780 | from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] | |
| 1781 | have pp': "\<forall> i. ?I i ?p' = ?I i p" by auto | |
| 1782 | from conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]] | |
| 1783 | have lp': "iszlfm ?p'" . | |
| 1784 | from lp' \<zeta>[where p="?p'"] have lp: "?l >0" and dl: "d\<beta> ?p' ?l" by auto | |
| 1785 | from a\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp' dl lp] pp' | |
| 1786 | have pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by simp | |
| 1787 | from lp' lp a\<beta>[OF lp' dl lp] have lq:"iszlfm ?q" and uq: "d\<beta> ?q 1" by auto | |
| 1788 | from \<delta>[OF lq] have dp:"?d >0" and dd: "d\<delta> ?q ?d" by blast+ | |
| 1789 | let ?N = "\<lambda> t. Inum (i#bs) t" | |
| 1790 | have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by auto | |
| 1791 | also have "\<dots> = ?N ` ?B" using simpnum_ci[where bs="i#bs"] by auto | |
| 1792 | finally have BB': "?N ` set ?B' = ?N ` ?B" . | |
| 1793 | have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by auto | |
| 1794 | also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="i#bs"] by auto | |
| 1795 | finally have AA': "?N ` set ?A' = ?N ` ?A" . | |
| 1796 | from \<beta>_numbound0[OF lq] have B_nb:"\<forall> b\<in> set ?B'. numbound0 b" | |
| 1797 | by (simp add: simpnum_numbound0) | |
| 1798 | from \<alpha>_l[OF lq] have A_nb: "\<forall> b\<in> set ?A'. numbound0 b" | |
| 1799 | by (simp add: simpnum_numbound0) | |
| 1800 |     {assume "length ?B' \<le> length ?A'"
 | |
| 1801 | hence q:"q=?q" and "B = ?B'" and d:"d = ?d" | |
| 1802 | using qBd by (auto simp add: Let_def unit_def) | |
| 1803 | with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<beta> q)" | |
| 1804 | and bn: "\<forall>b\<in> set B. numbound0 b" by simp+ | |
| 1805 | with pq_ex dp uq dd lq q d have ?thes by simp} | |
| 1806 | moreover | |
| 1807 |   {assume "\<not> (length ?B' \<le> length ?A')"
 | |
| 1808 | hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" | |
| 1809 | using qBd by (auto simp add: Let_def unit_def) | |
| 1810 | with AA' mirror\<alpha>\<beta>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<beta> q)" | |
| 1811 | and bn: "\<forall>b\<in> set B. numbound0 b" by simp+ | |
| 1812 | from mirror_ex[OF lq] pq_ex q | |
| 1813 | have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp | |
| 1814 | from lq uq q mirror_l[where p="?q"] | |
| 1815 | have lq': "iszlfm q" and uq: "d\<beta> q 1" by auto | |
| 1816 | from \<delta>[OF lq'] mirror_\<delta>[OF lq] q d have dq:"d\<delta> q d " by auto | |
| 1817 | from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp | |
| 1818 | } | |
| 1819 | ultimately show ?thes by blast | |
| 1820 | qed | |
| 1821 | (* Cooper's Algorithm *) | |
| 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: diff
changeset | 1822 | |
| 23274 | 1823 | constdefs cooper :: "fm \<Rightarrow> fm" | 
| 1824 | "cooper p \<equiv> | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1825 | (let (q,B,d) = unit p; js = iupt 1 d; | 
| 23274 | 1826 | mq = simpfm (minusinf q); | 
| 1827 | md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) js | |
| 1828 | in if md = T then T else | |
| 1829 | (let qd = evaldjf (\<lambda> (b,j). simpfm (subst0 (Add b (C j)) q)) | |
| 24336 | 1830 | [(b,j). b\<leftarrow>B,j\<leftarrow>js] | 
| 23274 | 1831 | in decr (disj md qd)))" | 
| 1832 | lemma cooper: assumes qf: "qfree p" | |
| 1833 | shows "((\<exists> x. Ifm bbs (x#bs) p) = (Ifm bbs bs (cooper p))) \<and> qfree (cooper p)" | |
| 1834 | (is "(?lhs = ?rhs) \<and> _") | |
| 1835 | proof- | |
| 1836 | let ?I = "\<lambda> x p. Ifm bbs (x#bs) p" | |
| 1837 | let ?q = "fst (unit p)" | |
| 1838 | let ?B = "fst (snd(unit p))" | |
| 1839 | let ?d = "snd (snd (unit p))" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1840 | let ?js = "iupt 1 ?d" | 
| 23274 | 1841 | let ?mq = "minusinf ?q" | 
| 1842 | let ?smq = "simpfm ?mq" | |
| 1843 | let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js" | |
| 26934 | 1844 | fix i | 
| 23274 | 1845 | let ?N = "\<lambda> t. Inum (i#bs) t" | 
| 24336 | 1846 | let ?Bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]" | 
| 1847 | let ?qd = "evaldjf (\<lambda> (b,j). simpfm (subst0 (Add b (C j)) ?q)) ?Bjs" | |
| 23274 | 1848 | have qbf:"unit p = (?q,?B,?d)" by simp | 
| 1849 | from unit[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and | |
| 1850 | B:"?N ` set ?B = ?N ` set (\<beta> ?q)" and | |
| 1851 | uq:"d\<beta> ?q 1" and dd: "d\<delta> ?q ?d" and dp: "?d > 0" and | |
| 1852 | lq: "iszlfm ?q" and | |
| 1853 | Bn: "\<forall> b\<in> set ?B. numbound0 b" by auto | |
| 1854 | from zlin_qfree[OF lq] have qfq: "qfree ?q" . | |
| 1855 | from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". | |
| 1856 | have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp | |
| 1857 | hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" | |
| 1858 | by (auto simp only: subst0_bound0[OF qfmq]) | |
| 1859 | hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))" | |
| 1860 | by (auto simp add: simpfm_bound0) | |
| 1861 | from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp | |
| 24336 | 1862 | from Bn jsnb have "\<forall> (b,j) \<in> set ?Bjs. numbound0 (Add b (C j))" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1863 | by simp | 
| 24336 | 1864 | hence "\<forall> (b,j) \<in> set ?Bjs. bound0 (subst0 (Add b (C j)) ?q)" | 
| 23274 | 1865 | using subst0_bound0[OF qfq] by blast | 
| 24336 | 1866 | hence "\<forall> (b,j) \<in> set ?Bjs. bound0 (simpfm (subst0 (Add b (C j)) ?q))" | 
| 23274 | 1867 | using simpfm_bound0 by blast | 
| 24336 | 1868 | hence th': "\<forall> x \<in> set ?Bjs. bound0 ((\<lambda> (b,j). simpfm (subst0 (Add b (C j)) ?q)) x)" | 
| 23274 | 1869 | by auto | 
| 1870 | from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp | |
| 1871 | from mdb qdb | |
| 1872 | have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all) | |
| 1873 | from trans [OF pq_ex cp_thm'[OF lq uq dd dp,where i="i"]] B | |
| 1874 |   have "?lhs = (\<exists> j\<in> {1.. ?d}. ?I j ?mq \<or> (\<exists> b\<in> ?N ` set ?B. Ifm bbs ((b+ j)#bs) ?q))" by auto
 | |
| 1875 |   also have "\<dots> = (\<exists> j\<in> {1.. ?d}. ?I j ?mq \<or> (\<exists> b\<in> set ?B. Ifm bbs ((?N b+ j)#bs) ?q))" by simp
 | |
| 1876 |   also have "\<dots> = ((\<exists> j\<in> {1.. ?d}. ?I j ?mq ) \<or> (\<exists> j\<in> {1.. ?d}. \<exists> b\<in> set ?B. Ifm bbs ((?N (Add b (C j)))#bs) ?q))" by (simp only: Inum.simps) blast
 | |
| 1877 |   also have "\<dots> = ((\<exists> j\<in> {1.. ?d}. ?I j ?smq ) \<or> (\<exists> j\<in> {1.. ?d}. \<exists> b\<in> set ?B. Ifm bbs ((?N (Add b (C j)))#bs) ?q))" by (simp add: simpfm) 
 | |
| 1878 | also have "\<dots> = ((\<exists> j\<in> set ?js. (\<lambda> j. ?I i (simpfm (subst0 (C j) ?smq))) j) \<or> (\<exists> j\<in> set ?js. \<exists> b\<in> set ?B. Ifm bbs ((?N (Add b (C j)))#bs) ?q))" | |
| 1879 | by (simp only: simpfm subst0_I[OF qfmq] iupt_set) auto | |
| 1880 | also have "\<dots> = (?I i (evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js) \<or> (\<exists> j\<in> set ?js. \<exists> b\<in> set ?B. ?I i (subst0 (Add b (C j)) ?q)))" | |
| 1881 | by (simp only: evaldjf_ex subst0_I[OF qfq]) | |
| 24336 | 1882 | also have "\<dots>= (?I i ?md \<or> (\<exists> (b,j) \<in> set ?Bjs. (\<lambda> (b,j). ?I i (simpfm (subst0 (Add b (C j)) ?q))) (b,j)))" | 
| 24349 | 1883 | by (simp only: simpfm set_concat set_map concat_map_singleton UN_simps) blast | 
| 24336 | 1884 | also have "\<dots> = (?I i ?md \<or> (?I i (evaldjf (\<lambda> (b,j). simpfm (subst0 (Add b (C j)) ?q)) ?Bjs)))" | 
| 1885 | by (simp only: evaldjf_ex[where bs="i#bs" and f="\<lambda> (b,j). simpfm (subst0 (Add b (C j)) ?q)" and ps="?Bjs"]) (auto simp add: split_def) | |
| 23274 | 1886 | finally have mdqd: "?lhs = (?I i ?md \<or> ?I i ?qd)" by simp | 
| 1887 | also have "\<dots> = (?I i (disj ?md ?qd))" by (simp add: disj) | |
| 1888 | also have "\<dots> = (Ifm bbs bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) | |
| 1889 | finally have mdqd2: "?lhs = (Ifm bbs bs (decr (disj ?md ?qd)))" . | |
| 1890 |   {assume mdT: "?md = T"
 | |
| 1891 | hence cT:"cooper p = T" | |
| 1892 | by (simp only: cooper_def unit_def split_def Let_def if_True) simp | |
| 1893 | from mdT have lhs:"?lhs" using mdqd by simp | |
| 1894 | from mdT have "?rhs" by (simp add: cooper_def unit_def split_def) | |
| 1895 | with lhs cT have ?thesis by simp } | |
| 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: diff
changeset | 1896 | moreover | 
| 23274 | 1897 |   {assume mdT: "?md \<noteq> T" hence "cooper p = decr (disj ?md ?qd)" 
 | 
| 1898 | by (simp only: cooper_def unit_def split_def Let_def if_False) | |
| 1899 | with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } | |
| 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: diff
changeset | 1900 | ultimately show ?thesis by blast | 
| 
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: diff
changeset | 1901 | qed | 
| 
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: diff
changeset | 1902 | |
| 27456 | 1903 | definition pa :: "fm \<Rightarrow> fm" where | 
| 1904 | "pa p = qelim (prep p) cooper" | |
| 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: diff
changeset | 1905 | |
| 23274 | 1906 | theorem mirqe: "(Ifm bbs bs (pa p) = Ifm bbs bs p) \<and> qfree (pa p)" | 
| 1907 | using qelim_ci cooper prep by (auto simp add: pa_def) | |
| 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: diff
changeset | 1908 | |
| 23515 | 1909 | definition | 
| 1910 | cooper_test :: "unit \<Rightarrow> fm" | |
| 1911 | where | |
| 1912 | "cooper_test u = pa (E (A (Imp (Ge (Sub (Bound 0) (Bound 1))) | |
| 1913 | (E (E (Eq (Sub (Add (Mul 3 (Bound 1)) (Mul 5 (Bound 0))) | |
| 1914 | (Bound 2))))))))" | |
| 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: diff
changeset | 1915 | |
| 27456 | 1916 | ML {* @{code cooper_test} () *}
 | 
| 1917 | ||
| 1918 | (* | |
| 23808 | 1919 | code_reserved SML oo | 
| 27456 | 1920 | export_code pa in SML module_name GeneratedCooper file "~~/src/HOL/Tools/Qelim/raw_generated_cooper.ML" | 
| 1921 | *) | |
| 1922 | ||
| 28290 | 1923 | oracle linzqe_oracle = {*
 | 
| 27456 | 1924 | let | 
| 1925 | ||
| 1926 | fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t | |
| 1927 | of NONE => error "Variable not found in the list!" | |
| 1928 |       | SOME n => @{code Bound} n)
 | |
| 1929 |   | num_of_term vs @{term "0::int"} = @{code C} 0
 | |
| 1930 |   | num_of_term vs @{term "1::int"} = @{code C} 1
 | |
| 1931 |   | num_of_term vs (@{term "number_of :: int \<Rightarrow> int"} $ t) = @{code C} (HOLogic.dest_numeral t)
 | |
| 1932 |   | num_of_term vs (Bound i) = @{code Bound} i
 | |
| 1933 |   | num_of_term vs (@{term "uminus :: int \<Rightarrow> int"} $ t') = @{code Neg} (num_of_term vs t')
 | |
| 1934 |   | num_of_term vs (@{term "op + :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
 | |
| 1935 |       @{code Add} (num_of_term vs t1, num_of_term vs t2)
 | |
| 1936 |   | num_of_term vs (@{term "op - :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
 | |
| 1937 |       @{code Sub} (num_of_term vs t1, num_of_term vs t2)
 | |
| 1938 |   | num_of_term vs (@{term "op * :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
 | |
| 1939 | (case try HOLogic.dest_number t1 | |
| 1940 |        of SOME (_, i) => @{code Mul} (i, num_of_term vs t2)
 | |
| 1941 | | NONE => (case try HOLogic.dest_number t2 | |
| 1942 |                 of SOME (_, i) => @{code Mul} (i, num_of_term vs t1)
 | |
| 1943 | | NONE => error "num_of_term: unsupported multiplication")) | |
| 28264 | 1944 |   | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t);
 | 
| 27456 | 1945 | |
| 1946 | fun fm_of_term ps vs @{term True} = @{code T}
 | |
| 1947 |   | fm_of_term ps vs @{term False} = @{code F}
 | |
| 1948 |   | fm_of_term ps vs (@{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
 | |
| 1949 |       @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
 | |
| 1950 |   | fm_of_term ps vs (@{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
 | |
| 1951 |       @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
 | |
| 1952 |   | fm_of_term ps vs (@{term "op = :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
 | |
| 1953 |       @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) 
 | |
| 1954 |   | fm_of_term ps vs (@{term "op dvd :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
 | |
| 1955 | (case try HOLogic.dest_number t1 | |
| 1956 |        of SOME (_, i) => @{code Dvd} (i, num_of_term vs t2)
 | |
| 1957 | | NONE => error "num_of_term: unsupported dvd") | |
| 1958 |   | fm_of_term ps vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
 | |
| 1959 |       @{code Iff} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | |
| 1960 |   | fm_of_term ps vs (@{term "op &"} $ t1 $ t2) =
 | |
| 1961 |       @{code And} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | |
| 1962 |   | fm_of_term ps vs (@{term "op |"} $ t1 $ t2) =
 | |
| 1963 |       @{code Or} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | |
| 1964 |   | fm_of_term ps vs (@{term "op -->"} $ t1 $ t2) =
 | |
| 1965 |       @{code Imp} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | |
| 1966 |   | fm_of_term ps vs (@{term "Not"} $ t') =
 | |
| 1967 |       @{code NOT} (fm_of_term ps vs t')
 | |
| 1968 |   | fm_of_term ps vs (Const ("Ex", _) $ Abs (xn, xT, p)) =
 | |
| 1969 | let | |
| 1970 | val (xn', p') = variant_abs (xn, xT, p); | |
| 1971 | val vs' = (Free (xn', xT), 0) :: map (fn (v, n) => (v, n + 1)) vs; | |
| 1972 |       in @{code E} (fm_of_term ps vs' p) end
 | |
| 1973 |   | fm_of_term ps vs (Const ("All", _) $ Abs (xn, xT, p)) =
 | |
| 1974 | let | |
| 1975 | val (xn', p') = variant_abs (xn, xT, p); | |
| 1976 | val vs' = (Free (xn', xT), 0) :: map (fn (v, n) => (v, n + 1)) vs; | |
| 1977 |       in @{code A} (fm_of_term ps vs' p) end
 | |
| 28264 | 1978 |   | fm_of_term ps vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t);
 | 
| 23515 | 1979 | |
| 27456 | 1980 | fun term_of_num vs (@{code C} i) = HOLogic.mk_number HOLogic.intT i
 | 
| 1981 |   | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs))
 | |
| 1982 |   | term_of_num vs (@{code Neg} t') = @{term "uminus :: int \<Rightarrow> int"} $ term_of_num vs t'
 | |
| 1983 |   | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: int \<Rightarrow> int \<Rightarrow> int"} $
 | |
| 1984 | term_of_num vs t1 $ term_of_num vs t2 | |
| 1985 |   | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: int \<Rightarrow> int \<Rightarrow> int"} $
 | |
| 1986 | term_of_num vs t1 $ term_of_num vs t2 | |
| 1987 |   | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: int \<Rightarrow> int \<Rightarrow> int"} $
 | |
| 1988 |       term_of_num vs (@{code C} i) $ term_of_num vs t2
 | |
| 1989 |   | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t))
 | |
| 1990 | ||
| 1991 | fun term_of_fm ps vs @{code T} = HOLogic.true_const 
 | |
| 1992 |   | term_of_fm ps vs @{code F} = HOLogic.false_const
 | |
| 1993 |   | term_of_fm ps vs (@{code Lt} t) =
 | |
| 1994 |       @{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::int"}
 | |
| 1995 |   | term_of_fm ps vs (@{code Le} t) =
 | |
| 1996 |       @{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::int"}
 | |
| 1997 |   | term_of_fm ps vs (@{code Gt} t) =
 | |
| 1998 |       @{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"} $ @{term "0::int"} $ term_of_num vs t
 | |
| 1999 |   | term_of_fm ps vs (@{code Ge} t) =
 | |
| 2000 |       @{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"} $ @{term "0::int"} $ term_of_num vs t
 | |
| 2001 |   | term_of_fm ps vs (@{code Eq} t) =
 | |
| 2002 |       @{term "op = :: int \<Rightarrow> int \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::int"}
 | |
| 2003 |   | term_of_fm ps vs (@{code NEq} t) =
 | |
| 2004 |       term_of_fm ps vs (@{code NOT} (@{code Eq} t))
 | |
| 2005 |   | term_of_fm ps vs (@{code Dvd} (i, t)) =
 | |
| 2006 |       @{term "op dvd :: int \<Rightarrow> int \<Rightarrow> bool"} $ term_of_num vs (@{code C} i) $ term_of_num vs t
 | |
| 2007 |   | term_of_fm ps vs (@{code NDvd} (i, t)) =
 | |
| 2008 |       term_of_fm ps vs (@{code NOT} (@{code Dvd} (i, t)))
 | |
| 2009 |   | term_of_fm ps vs (@{code NOT} t') =
 | |
| 2010 | HOLogic.Not $ term_of_fm ps vs t' | |
| 2011 |   | term_of_fm ps vs (@{code And} (t1, t2)) =
 | |
| 2012 | HOLogic.conj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | |
| 2013 |   | term_of_fm ps vs (@{code Or} (t1, t2)) =
 | |
| 2014 | HOLogic.disj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | |
| 2015 |   | term_of_fm ps vs (@{code Imp} (t1, t2)) =
 | |
| 2016 | HOLogic.imp $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | |
| 2017 |   | term_of_fm ps vs (@{code Iff} (t1, t2)) =
 | |
| 2018 |       @{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ term_of_fm ps vs t1 $ term_of_fm ps vs t2
 | |
| 2019 |   | term_of_fm ps vs (@{code Closed} n) = (fst o the) (find_first (fn (_, m) => m = n) ps)
 | |
| 2020 |   | term_of_fm ps vs (@{code NClosed} n) = term_of_fm ps vs (@{code NOT} (@{code Closed} n))
 | |
| 2021 | ||
| 2022 | fun term_bools acc t = | |
| 2023 | let | |
| 2024 |     val is_op = member (op =) [@{term "op &"}, @{term "op |"}, @{term "op -->"}, @{term "op = :: bool => _"},
 | |
| 2025 |       @{term "op = :: int => _"}, @{term "op < :: int => _"},
 | |
| 2026 |       @{term "op <= :: int => _"}, @{term "Not"}, @{term "All :: (int => _) => _"},
 | |
| 2027 |       @{term "Ex :: (int => _) => _"}, @{term "True"}, @{term "False"}]
 | |
| 2028 | fun is_ty t = not (fastype_of t = HOLogic.boolT) | |
| 2029 | in case t | |
| 2030 | of (l as f $ a) $ b => if is_ty t orelse is_op t then term_bools (term_bools acc l)b | |
| 2031 | else insert (op aconv) t acc | |
| 2032 | | f $ a => if is_ty t orelse is_op t then term_bools (term_bools acc f) a | |
| 2033 | else insert (op aconv) t acc | |
| 2034 | | Abs p => term_bools acc (snd (variant_abs p)) | |
| 2035 | | _ => if is_ty t orelse is_op t then acc else insert (op aconv) t acc | |
| 2036 | end; | |
| 2037 | ||
| 28290 | 2038 | in fn ct => | 
| 2039 | let | |
| 2040 | val thy = Thm.theory_of_cterm ct; | |
| 2041 | val t = Thm.term_of ct; | |
| 27456 | 2042 | val fs = term_frees t; | 
| 2043 | val bs = term_bools [] t; | |
| 2044 | val vs = fs ~~ (0 upto (length fs - 1)) | |
| 2045 | val ps = bs ~~ (0 upto (length bs - 1)) | |
| 2046 |     val t' = (term_of_fm ps vs o @{code pa} o fm_of_term ps vs) t;
 | |
| 28290 | 2047 | in (Thm.cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end | 
| 27456 | 2048 | end; | 
| 2049 | *} | |
| 2050 | ||
| 23515 | 2051 | use "coopertac.ML" | 
| 23274 | 2052 | setup "LinZTac.setup" | 
| 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: diff
changeset | 2053 | |
| 27456 | 2054 | text {* Tests *}
 | 
| 2055 | ||
| 23274 | 2056 | lemma "\<exists> (j::int). \<forall> x\<ge>j. (\<exists> a b. x = 3*a+5*b)" | 
| 27456 | 2057 | by cooper | 
| 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: diff
changeset | 2058 | |
| 27456 | 2059 | lemma "ALL (x::int) >=8. EX i j. 5*i + 3*j = x" | 
| 2060 | by cooper | |
| 2061 | ||
| 23274 | 2062 | theorem "(\<forall>(y::int). 3 dvd y) ==> \<forall>(x::int). b < x --> a \<le> x" | 
| 2063 | by cooper | |
| 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: diff
changeset | 2064 | |
| 23274 | 2065 | theorem "!! (y::int) (z::int) (n::int). 3 dvd z ==> 2 dvd (y::int) ==> | 
| 2066 | (\<exists>(x::int). 2*x = y) & (\<exists>(k::int). 3*k = z)" | |
| 2067 | by cooper | |
| 2068 | ||
| 2069 | theorem "!! (y::int) (z::int) n. Suc(n::nat) < 6 ==> 3 dvd z ==> | |
| 2070 | 2 dvd (y::int) ==> (\<exists>(x::int). 2*x = y) & (\<exists>(k::int). 3*k = z)" | |
| 2071 | by cooper | |
| 2072 | ||
| 2073 | theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 --> y = 5 + x " | |
| 2074 | by cooper | |
| 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: diff
changeset | 2075 | |
| 27456 | 2076 | lemma "ALL (x::int) >=8. EX i j. 5*i + 3*j = x" | 
| 2077 | by cooper | |
| 2078 | ||
| 2079 | lemma "ALL (y::int) (z::int) (n::int). 3 dvd z --> 2 dvd (y::int) --> (EX (x::int). 2*x = y) & (EX (k::int). 3*k = z)" | |
| 2080 | by cooper | |
| 2081 | ||
| 2082 | lemma "ALL(x::int) y. x < y --> 2 * x + 1 < 2 * y" | |
| 2083 | by cooper | |
| 2084 | ||
| 2085 | lemma "ALL(x::int) y. 2 * x + 1 ~= 2 * y" | |
| 2086 | by cooper | |
| 2087 | ||
| 2088 | lemma "EX(x::int) y. 0 < x & 0 <= y & 3 * x - 5 * y = 1" | |
| 2089 | by cooper | |
| 2090 | ||
| 2091 | lemma "~ (EX(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)" | |
| 2092 | by cooper | |
| 2093 | ||
| 2094 | lemma "ALL(x::int). (2 dvd x) --> (EX(y::int). x = 2*y)" | |
| 2095 | by cooper | |
| 2096 | ||
| 2097 | lemma "ALL(x::int). (2 dvd x) = (EX(y::int). x = 2*y)" | |
| 2098 | by cooper | |
| 2099 | ||
| 2100 | lemma "ALL(x::int). ((2 dvd x) = (ALL(y::int). x ~= 2*y + 1))" | |
| 2101 | by cooper | |
| 2102 | ||
| 2103 | lemma "~ (ALL(x::int). ((2 dvd x) = (ALL(y::int). x ~= 2*y+1) | (EX(q::int) (u::int) i. 3*i + 2*q - u < 17) --> 0 < x | ((~ 3 dvd x) &(x + 8 = 0))))" | |
| 2104 | by cooper | |
| 2105 | ||
| 23274 | 2106 | lemma "~ (ALL(i::int). 4 <= i --> (EX x y. 0 <= x & 0 <= y & 3 * x + 5 * y = i))" | 
| 2107 | by cooper | |
| 27456 | 2108 | |
| 2109 | lemma "EX j. ALL (x::int) >= j. EX i j. 5*i + 3*j = x" | |
| 2110 | by cooper | |
| 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: diff
changeset | 2111 | |
| 23274 | 2112 | theorem "(\<forall>(y::int). 3 dvd y) ==> \<forall>(x::int). b < x --> a \<le> x" | 
| 2113 | by cooper | |
| 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: diff
changeset | 2114 | |
| 23274 | 2115 | theorem "!! (y::int) (z::int) (n::int). 3 dvd z ==> 2 dvd (y::int) ==> | 
| 2116 | (\<exists>(x::int). 2*x = y) & (\<exists>(k::int). 3*k = z)" | |
| 2117 | by cooper | |
| 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: diff
changeset | 2118 | |
| 23274 | 2119 | theorem "!! (y::int) (z::int) n. Suc(n::nat) < 6 ==> 3 dvd z ==> | 
| 2120 | 2 dvd (y::int) ==> (\<exists>(x::int). 2*x = y) & (\<exists>(k::int). 3*k = z)" | |
| 2121 | by cooper | |
| 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: diff
changeset | 2122 | |
| 23274 | 2123 | theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 --> y = 5 + x " | 
| 2124 | by cooper | |
| 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: diff
changeset | 2125 | |
| 23274 | 2126 | theorem "\<forall>(x::nat). \<exists>(y::nat). y = 5 + x | x div 6 + 1= 2" | 
| 2127 | by cooper | |
| 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: diff
changeset | 2128 | |
| 23274 | 2129 | theorem "\<exists>(x::int). 0 < x" | 
| 2130 | by cooper | |
| 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: diff
changeset | 2131 | |
| 23274 | 2132 | theorem "\<forall>(x::int) y. x < y --> 2 * x + 1 < 2 * y" | 
| 2133 | by cooper | |
| 2134 | ||
| 2135 | theorem "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y" | |
| 2136 | by cooper | |
| 2137 | ||
| 2138 | theorem "\<exists>(x::int) y. 0 < x & 0 \<le> y & 3 * x - 5 * y = 1" | |
| 2139 | by cooper | |
| 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: diff
changeset | 2140 | |
| 23274 | 2141 | theorem "~ (\<exists>(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)" | 
| 2142 | by cooper | |
| 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: diff
changeset | 2143 | |
| 23274 | 2144 | theorem "~ (\<exists>(x::int). False)" | 
| 2145 | by cooper | |
| 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: diff
changeset | 2146 | |
| 23274 | 2147 | theorem "\<forall>(x::int). (2 dvd x) --> (\<exists>(y::int). x = 2*y)" | 
| 2148 | by cooper | |
| 2149 | ||
| 2150 | theorem "\<forall>(x::int). (2 dvd x) --> (\<exists>(y::int). x = 2*y)" | |
| 2151 | by cooper | |
| 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: diff
changeset | 2152 | |
| 23274 | 2153 | theorem "\<forall>(x::int). (2 dvd x) = (\<exists>(y::int). x = 2*y)" | 
| 2154 | by cooper | |
| 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: diff
changeset | 2155 | |
| 23274 | 2156 | theorem "\<forall>(x::int). ((2 dvd x) = (\<forall>(y::int). x \<noteq> 2*y + 1))" | 
| 2157 | by cooper | |
| 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: diff
changeset | 2158 | |
| 23274 | 2159 | theorem "~ (\<forall>(x::int). | 
| 2160 | ((2 dvd x) = (\<forall>(y::int). x \<noteq> 2*y+1) | | |
| 2161 | (\<exists>(q::int) (u::int) i. 3*i + 2*q - u < 17) | |
| 2162 | --> 0 < x | ((~ 3 dvd x) &(x + 8 = 0))))" | |
| 2163 | by cooper | |
| 2164 | ||
| 2165 | theorem "~ (\<forall>(i::int). 4 \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i))" | |
| 2166 | by cooper | |
| 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: diff
changeset | 2167 | |
| 23274 | 2168 | theorem "\<forall>(i::int). 8 \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i)" | 
| 2169 | by cooper | |
| 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: diff
changeset | 2170 | |
| 23274 | 2171 | theorem "\<exists>(j::int). \<forall>i. j \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i)" | 
| 2172 | by cooper | |
| 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: diff
changeset | 2173 | |
| 23274 | 2174 | theorem "~ (\<forall>j (i::int). j \<le> i --> (\<exists>x y. 0 \<le> x & 0 \<le> y & 3 * x + 5 * y = i))" | 
| 2175 | by cooper | |
| 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: diff
changeset | 2176 | |
| 23274 | 2177 | theorem "(\<exists>m::nat. n = 2 * m) --> (n + 1) div 2 = n div 2" | 
| 2178 | by cooper | |
| 17388 | 2179 | |
| 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: diff
changeset | 2180 | end |