| author | blanchet | 
| Wed, 04 Nov 2015 15:07:23 +0100 | |
| changeset 61569 | 947ce60a06e1 | 
| parent 60708 | f425e80a3eb0 | 
| child 61586 | 5197a2ecb658 | 
| permissions | -rw-r--r-- | 
| 30439 | 1 | (* Title: HOL/Decision_Procs/Cooper.thy | 
| 27456 | 2 | Author: Amine Chaieb | 
| 3 | *) | |
| 4 | ||
| 29788 | 5 | theory Cooper | 
| 55885 | 6 | imports | 
| 7 | Complex_Main | |
| 8 | "~~/src/HOL/Library/Code_Target_Numeral" | |
| 9 | "~~/src/HOL/Library/Old_Recdef" | |
| 23477 
f4b83f03cac9
tuned and renamed group_eq_simps and ring_eq_simps
 nipkow parents: 
23315diff
changeset | 10 | 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 | 11 | |
| 
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 | 12 | (* Periodicity of dvd *) | 
| 23315 | 13 | |
| 55885 | 14 | (*********************************************************************************) | 
| 15 | (**** SHADOW SYNTAX AND SEMANTICS ****) | |
| 16 | (*********************************************************************************) | |
| 23274 | 17 | |
| 58310 | 18 | datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num | 
| 23274 | 19 | | Mul int num | 
| 20 | ||
| 60533 | 21 | primrec num_size :: "num \<Rightarrow> nat" -- \<open>A size for num to make inductive proofs simpler\<close> | 
| 50313 | 22 | where | 
| 23274 | 23 | "num_size (C c) = 1" | 
| 27456 | 24 | | "num_size (Bound n) = 1" | 
| 25 | | "num_size (Neg a) = 1 + num_size a" | |
| 26 | | "num_size (Add a b) = 1 + num_size a + num_size b" | |
| 27 | | "num_size (Sub a b) = 3 + num_size a + num_size b" | |
| 28 | | "num_size (CN n c a) = 4 + num_size a" | |
| 29 | | "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 | 30 | |
| 55844 | 31 | primrec Inum :: "int list \<Rightarrow> num \<Rightarrow> int" | 
| 32 | where | |
| 23274 | 33 | "Inum bs (C c) = c" | 
| 27456 | 34 | | "Inum bs (Bound n) = bs!n" | 
| 35 | | "Inum bs (CN n c a) = c * (bs!n) + (Inum bs a)" | |
| 36 | | "Inum bs (Neg a) = -(Inum bs a)" | |
| 37 | | "Inum bs (Add a b) = Inum bs a + Inum bs b" | |
| 38 | | "Inum bs (Sub a b) = Inum bs a - Inum bs b" | |
| 39 | | "Inum bs (Mul c a) = c* Inum bs a" | |
| 23274 | 40 | |
| 58310 | 41 | datatype fm = | 
| 23274 | 42 | T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num| | 
| 50313 | 43 | NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm | 
| 23274 | 44 | | Closed nat | NClosed nat | 
| 45 | ||
| 46 | ||
| 60533 | 47 | fun fmsize :: "fm \<Rightarrow> nat" -- \<open>A size for fm\<close> | 
| 50313 | 48 | where | 
| 23274 | 49 | "fmsize (NOT p) = 1 + fmsize p" | 
| 41837 | 50 | | "fmsize (And p q) = 1 + fmsize p + fmsize q" | 
| 51 | | "fmsize (Or p q) = 1 + fmsize p + fmsize q" | |
| 52 | | "fmsize (Imp p q) = 3 + fmsize p + fmsize q" | |
| 53 | | "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)" | |
| 54 | | "fmsize (E p) = 1 + fmsize p" | |
| 55 | | "fmsize (A p) = 4+ fmsize p" | |
| 56 | | "fmsize (Dvd i t) = 2" | |
| 57 | | "fmsize (NDvd i t) = 2" | |
| 58 | | "fmsize p = 1" | |
| 50313 | 59 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31952diff
changeset | 60 | lemma fmsize_pos: "fmsize p > 0" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31952diff
changeset | 61 | 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 | 62 | |
| 60533 | 63 | primrec Ifm :: "bool list \<Rightarrow> int list \<Rightarrow> fm \<Rightarrow> bool" -- \<open>Semantics of formulae (fm)\<close> | 
| 50313 | 64 | where | 
| 55981 | 65 | "Ifm bbs bs T \<longleftrightarrow> True" | 
| 66 | | "Ifm bbs bs F \<longleftrightarrow> False" | |
| 67 | | "Ifm bbs bs (Lt a) \<longleftrightarrow> Inum bs a < 0" | |
| 68 | | "Ifm bbs bs (Gt a) \<longleftrightarrow> Inum bs a > 0" | |
| 69 | | "Ifm bbs bs (Le a) \<longleftrightarrow> Inum bs a \<le> 0" | |
| 70 | | "Ifm bbs bs (Ge a) \<longleftrightarrow> Inum bs a \<ge> 0" | |
| 71 | | "Ifm bbs bs (Eq a) \<longleftrightarrow> Inum bs a = 0" | |
| 72 | | "Ifm bbs bs (NEq a) \<longleftrightarrow> Inum bs a \<noteq> 0" | |
| 73 | | "Ifm bbs bs (Dvd i b) \<longleftrightarrow> i dvd Inum bs b" | |
| 74 | | "Ifm bbs bs (NDvd i b) \<longleftrightarrow> \<not> i dvd Inum bs b" | |
| 75 | | "Ifm bbs bs (NOT p) \<longleftrightarrow> \<not> Ifm bbs bs p" | |
| 76 | | "Ifm bbs bs (And p q) \<longleftrightarrow> Ifm bbs bs p \<and> Ifm bbs bs q" | |
| 77 | | "Ifm bbs bs (Or p q) \<longleftrightarrow> Ifm bbs bs p \<or> Ifm bbs bs q" | |
| 78 | | "Ifm bbs bs (Imp p q) \<longleftrightarrow> (Ifm bbs bs p \<longrightarrow> Ifm bbs bs q)" | |
| 79 | | "Ifm bbs bs (Iff p q) \<longleftrightarrow> Ifm bbs bs p = Ifm bbs bs q" | |
| 80 | | "Ifm bbs bs (E p) \<longleftrightarrow> (\<exists>x. Ifm bbs (x # bs) p)" | |
| 81 | | "Ifm bbs bs (A p) \<longleftrightarrow> (\<forall>x. Ifm bbs (x # bs) p)" | |
| 82 | | "Ifm bbs bs (Closed n) \<longleftrightarrow> bbs!n" | |
| 83 | | "Ifm bbs bs (NClosed n) \<longleftrightarrow> \<not> bbs!n" | |
| 23274 | 84 | |
| 85 | consts prep :: "fm \<Rightarrow> fm" | |
| 86 | recdef prep "measure fmsize" | |
| 87 | "prep (E T) = T" | |
| 88 | "prep (E F) = F" | |
| 89 | "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))" | |
| 90 | "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))" | |
| 50313 | 91 | "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" | 
| 23274 | 92 | "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))" | 
| 93 | "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))" | |
| 94 | "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" | |
| 95 | "prep (E p) = E (prep p)" | |
| 96 | "prep (A (And p q)) = And (prep (A p)) (prep (A q))" | |
| 97 | "prep (A p) = prep (NOT (E (NOT p)))" | |
| 98 | "prep (NOT (NOT p)) = prep p" | |
| 99 | "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))" | |
| 100 | "prep (NOT (A p)) = prep (E (NOT p))" | |
| 101 | "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))" | |
| 102 | "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))" | |
| 103 | "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))" | |
| 104 | "prep (NOT p) = NOT (prep p)" | |
| 105 | "prep (Or p q) = Or (prep p) (prep q)" | |
| 106 | "prep (And p q) = And (prep p) (prep q)" | |
| 107 | "prep (Imp p q) = prep (Or (NOT p) q)" | |
| 108 | "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))" | |
| 109 | "prep p = p" | |
| 50313 | 110 | (hints simp add: fmsize_pos) | 
| 111 | ||
| 23274 | 112 | lemma prep: "Ifm bbs bs (prep p) = Ifm bbs bs p" | 
| 50313 | 113 | by (induct p arbitrary: bs rule: prep.induct) auto | 
| 23274 | 114 | |
| 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 | 115 | |
| 60533 | 116 | fun qfree :: "fm \<Rightarrow> bool" -- \<open>Quantifier freeness\<close> | 
| 50313 | 117 | where | 
| 55999 | 118 | "qfree (E p) \<longleftrightarrow> False" | 
| 119 | | "qfree (A p) \<longleftrightarrow> False" | |
| 120 | | "qfree (NOT p) \<longleftrightarrow> qfree p" | |
| 121 | | "qfree (And p q) \<longleftrightarrow> qfree p \<and> qfree q" | |
| 122 | | "qfree (Or p q) \<longleftrightarrow> qfree p \<and> qfree q" | |
| 123 | | "qfree (Imp p q) \<longleftrightarrow> qfree p \<and> qfree q" | |
| 124 | | "qfree (Iff p q) \<longleftrightarrow> qfree p \<and> qfree q" | |
| 125 | | "qfree p \<longleftrightarrow> True" | |
| 23274 | 126 | |
| 50313 | 127 | |
| 60533 | 128 | text \<open>Boundedness and substitution\<close> | 
| 50313 | 129 | |
| 60533 | 130 | primrec numbound0 :: "num \<Rightarrow> bool" -- \<open>a num is INDEPENDENT of Bound 0\<close> | 
| 50313 | 131 | where | 
| 55981 | 132 | "numbound0 (C c) \<longleftrightarrow> True" | 
| 133 | | "numbound0 (Bound n) \<longleftrightarrow> n > 0" | |
| 134 | | "numbound0 (CN n i a) \<longleftrightarrow> n > 0 \<and> numbound0 a" | |
| 135 | | "numbound0 (Neg a) \<longleftrightarrow> numbound0 a" | |
| 136 | | "numbound0 (Add a b) \<longleftrightarrow> numbound0 a \<and> numbound0 b" | |
| 137 | | "numbound0 (Sub a b) \<longleftrightarrow> numbound0 a \<and> numbound0 b" | |
| 138 | | "numbound0 (Mul i a) \<longleftrightarrow> numbound0 a" | |
| 23274 | 139 | |
| 140 | lemma numbound0_I: | |
| 141 | assumes nb: "numbound0 a" | |
| 55981 | 142 | shows "Inum (b # bs) a = Inum (b' # bs) a" | 
| 50313 | 143 | using nb by (induct a rule: num.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 | 144 | |
| 60533 | 145 | primrec bound0 :: "fm \<Rightarrow> bool" -- \<open>A Formula is independent of Bound 0\<close> | 
| 50313 | 146 | where | 
| 55981 | 147 | "bound0 T \<longleftrightarrow> True" | 
| 148 | | "bound0 F \<longleftrightarrow> True" | |
| 149 | | "bound0 (Lt a) \<longleftrightarrow> numbound0 a" | |
| 150 | | "bound0 (Le a) \<longleftrightarrow> numbound0 a" | |
| 151 | | "bound0 (Gt a) \<longleftrightarrow> numbound0 a" | |
| 152 | | "bound0 (Ge a) \<longleftrightarrow> numbound0 a" | |
| 153 | | "bound0 (Eq a) \<longleftrightarrow> numbound0 a" | |
| 154 | | "bound0 (NEq a) \<longleftrightarrow> numbound0 a" | |
| 155 | | "bound0 (Dvd i a) \<longleftrightarrow> numbound0 a" | |
| 156 | | "bound0 (NDvd i a) \<longleftrightarrow> numbound0 a" | |
| 157 | | "bound0 (NOT p) \<longleftrightarrow> bound0 p" | |
| 158 | | "bound0 (And p q) \<longleftrightarrow> bound0 p \<and> bound0 q" | |
| 159 | | "bound0 (Or p q) \<longleftrightarrow> bound0 p \<and> bound0 q" | |
| 160 | | "bound0 (Imp p q) \<longleftrightarrow> bound0 p \<and> bound0 q" | |
| 161 | | "bound0 (Iff p q) \<longleftrightarrow> bound0 p \<and> bound0 q" | |
| 162 | | "bound0 (E p) \<longleftrightarrow> False" | |
| 163 | | "bound0 (A p) \<longleftrightarrow> False" | |
| 164 | | "bound0 (Closed P) \<longleftrightarrow> True" | |
| 165 | | "bound0 (NClosed P) \<longleftrightarrow> True" | |
| 50313 | 166 | |
| 23274 | 167 | lemma bound0_I: | 
| 168 | assumes bp: "bound0 p" | |
| 55981 | 169 | shows "Ifm bbs (b # bs) p = Ifm bbs (b' # bs) p" | 
| 50313 | 170 | using bp numbound0_I[where b="b" and bs="bs" and b'="b'"] | 
| 171 | by (induct p rule: fm.induct) (auto simp add: gr0_conv_Suc) | |
| 23274 | 172 | |
| 50313 | 173 | fun numsubst0 :: "num \<Rightarrow> num \<Rightarrow> num" | 
| 174 | where | |
| 23274 | 175 | "numsubst0 t (C c) = (C c)" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 176 | | "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 | 177 | | "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 | 178 | | "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 | 179 | | "numsubst0 t (Neg a) = Neg (numsubst0 t a)" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 180 | | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)" | 
| 50313 | 181 | | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 182 | | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)" | 
| 23274 | 183 | |
| 50313 | 184 | lemma numsubst0_I: "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t" | 
| 185 | 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 | 186 | |
| 23274 | 187 | lemma numsubst0_I': | 
| 25134 
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
 nipkow parents: 
25112diff
changeset | 188 | "numbound0 a \<Longrightarrow> Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t" | 
| 50313 | 189 | by (induct t rule: numsubst0.induct) (auto simp: nth_Cons' numbound0_I[where b="b" and b'="b'"]) | 
| 23274 | 190 | |
| 60533 | 191 | primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" -- \<open>substitue a num into a formula for Bound 0\<close> | 
| 50313 | 192 | where | 
| 23274 | 193 | "subst0 t T = T" | 
| 39246 | 194 | | "subst0 t F = F" | 
| 195 | | "subst0 t (Lt a) = Lt (numsubst0 t a)" | |
| 196 | | "subst0 t (Le a) = Le (numsubst0 t a)" | |
| 197 | | "subst0 t (Gt a) = Gt (numsubst0 t a)" | |
| 198 | | "subst0 t (Ge a) = Ge (numsubst0 t a)" | |
| 199 | | "subst0 t (Eq a) = Eq (numsubst0 t a)" | |
| 200 | | "subst0 t (NEq a) = NEq (numsubst0 t a)" | |
| 201 | | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)" | |
| 202 | | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)" | |
| 203 | | "subst0 t (NOT p) = NOT (subst0 t p)" | |
| 204 | | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)" | |
| 205 | | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)" | |
| 206 | | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)" | |
| 207 | | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)" | |
| 208 | | "subst0 t (Closed P) = (Closed P)" | |
| 209 | | "subst0 t (NClosed P) = (NClosed P)" | |
| 23274 | 210 | |
| 50313 | 211 | lemma subst0_I: | 
| 55999 | 212 | assumes "qfree p" | 
| 55885 | 213 | shows "Ifm bbs (b # bs) (subst0 a p) = Ifm bbs (Inum (b # bs) a # bs) p" | 
| 55999 | 214 | using assms numsubst0_I[where b="b" and bs="bs" and a="a"] | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 215 | by (induct p) (simp_all add: gr0_conv_Suc) | 
| 23274 | 216 | |
| 50313 | 217 | fun decrnum:: "num \<Rightarrow> num" | 
| 218 | where | |
| 23274 | 219 | "decrnum (Bound n) = Bound (n - 1)" | 
| 41837 | 220 | | "decrnum (Neg a) = Neg (decrnum a)" | 
| 221 | | "decrnum (Add a b) = Add (decrnum a) (decrnum b)" | |
| 222 | | "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" | |
| 223 | | "decrnum (Mul c a) = Mul c (decrnum a)" | |
| 224 | | "decrnum (CN n i a) = (CN (n - 1) i (decrnum a))" | |
| 225 | | "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 | 226 | |
| 50313 | 227 | fun decr :: "fm \<Rightarrow> fm" | 
| 228 | where | |
| 23274 | 229 | "decr (Lt a) = Lt (decrnum a)" | 
| 41837 | 230 | | "decr (Le a) = Le (decrnum a)" | 
| 231 | | "decr (Gt a) = Gt (decrnum a)" | |
| 232 | | "decr (Ge a) = Ge (decrnum a)" | |
| 233 | | "decr (Eq a) = Eq (decrnum a)" | |
| 234 | | "decr (NEq a) = NEq (decrnum a)" | |
| 235 | | "decr (Dvd i a) = Dvd i (decrnum a)" | |
| 236 | | "decr (NDvd i a) = NDvd i (decrnum a)" | |
| 50313 | 237 | | "decr (NOT p) = NOT (decr p)" | 
| 41837 | 238 | | "decr (And p q) = And (decr p) (decr q)" | 
| 239 | | "decr (Or p q) = Or (decr p) (decr q)" | |
| 240 | | "decr (Imp p q) = Imp (decr p) (decr q)" | |
| 241 | | "decr (Iff p q) = Iff (decr p) (decr q)" | |
| 242 | | "decr p = p" | |
| 23274 | 243 | |
| 50313 | 244 | lemma decrnum: | 
| 245 | assumes nb: "numbound0 t" | |
| 55885 | 246 | shows "Inum (x # bs) t = Inum bs (decrnum t)" | 
| 50313 | 247 | using nb by (induct t rule: decrnum.induct) (auto simp add: gr0_conv_Suc) | 
| 23274 | 248 | |
| 50313 | 249 | lemma decr: | 
| 250 | assumes nb: "bound0 p" | |
| 55885 | 251 | shows "Ifm bbs (x # bs) p = Ifm bbs bs (decr p)" | 
| 50313 | 252 | using nb by (induct p rule: decr.induct) (simp_all add: gr0_conv_Suc decrnum) | 
| 23274 | 253 | |
| 254 | lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)" | |
| 50313 | 255 | by (induct p) simp_all | 
| 23274 | 256 | |
| 60533 | 257 | fun isatom :: "fm \<Rightarrow> bool" -- \<open>test for atomicity\<close> | 
| 50313 | 258 | where | 
| 55981 | 259 | "isatom T \<longleftrightarrow> True" | 
| 260 | | "isatom F \<longleftrightarrow> True" | |
| 261 | | "isatom (Lt a) \<longleftrightarrow> True" | |
| 262 | | "isatom (Le a) \<longleftrightarrow> True" | |
| 263 | | "isatom (Gt a) \<longleftrightarrow> True" | |
| 264 | | "isatom (Ge a) \<longleftrightarrow> True" | |
| 265 | | "isatom (Eq a) \<longleftrightarrow> True" | |
| 266 | | "isatom (NEq a) \<longleftrightarrow> True" | |
| 267 | | "isatom (Dvd i b) \<longleftrightarrow> True" | |
| 268 | | "isatom (NDvd i b) \<longleftrightarrow> True" | |
| 269 | | "isatom (Closed P) \<longleftrightarrow> True" | |
| 270 | | "isatom (NClosed P) \<longleftrightarrow> True" | |
| 271 | | "isatom p \<longleftrightarrow> 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 | 272 | |
| 50313 | 273 | lemma numsubst0_numbound0: | 
| 55844 | 274 | assumes "numbound0 t" | 
| 23274 | 275 | shows "numbound0 (numsubst0 t a)" | 
| 55844 | 276 | using assms | 
| 60708 | 277 | proof (induct a) | 
| 278 | case (CN n _ _) | |
| 279 | then show ?case by (cases n) simp_all | |
| 280 | qed simp_all | |
| 23274 | 281 | |
| 50313 | 282 | lemma subst0_bound0: | 
| 55844 | 283 | assumes qf: "qfree p" | 
| 284 | and nb: "numbound0 t" | |
| 23274 | 285 | shows "bound0 (subst0 t p)" | 
| 50313 | 286 | using qf numsubst0_numbound0[OF nb] by (induct p) auto | 
| 23274 | 287 | |
| 288 | lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p" | |
| 50313 | 289 | by (induct p) simp_all | 
| 23274 | 290 | |
| 291 | ||
| 50313 | 292 | definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
 | 
| 293 | where | |
| 294 | "djf f p q = | |
| 55885 | 295 | (if q = T then T | 
| 296 | else if q = F then f p | |
| 297 | else | |
| 298 | let fp = f p | |
| 299 | in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q)" | |
| 50313 | 300 | |
| 301 | definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
 | |
| 302 | where "evaldjf f ps = foldr (djf f) ps F" | |
| 23274 | 303 | |
| 304 | lemma djf_Or: "Ifm bbs bs (djf f p q) = Ifm bbs bs (Or (f p) q)" | |
| 55885 | 305 | by (cases "q=T", simp add: djf_def, cases "q = F", simp add: djf_def) | 
| 50313 | 306 | (cases "f p", simp_all add: Let_def djf_def) | 
| 23274 | 307 | |
| 55885 | 308 | lemma evaldjf_ex: "Ifm bbs bs (evaldjf f ps) \<longleftrightarrow> (\<exists>p \<in> set ps. Ifm bbs bs (f p))" | 
| 50313 | 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 | |
| 50313 | 311 | lemma evaldjf_bound0: | 
| 312 | assumes nb: "\<forall>x\<in> set xs. bound0 (f x)" | |
| 23274 | 313 | shows "bound0 (evaldjf f xs)" | 
| 55422 | 314 | using nb by (induct xs) (auto simp add: evaldjf_def djf_def Let_def, case_tac "f a", auto) | 
| 23274 | 315 | |
| 50313 | 316 | lemma evaldjf_qf: | 
| 317 | assumes nb: "\<forall>x\<in> set xs. qfree (f x)" | |
| 23274 | 318 | shows "qfree (evaldjf f xs)" | 
| 55422 | 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 | |
| 50313 | 321 | fun disjuncts :: "fm \<Rightarrow> fm list" | 
| 322 | where | |
| 323 | "disjuncts (Or p q) = disjuncts p @ disjuncts q" | |
| 41837 | 324 | | "disjuncts F = []" | 
| 325 | | "disjuncts p = [p]" | |
| 23274 | 326 | |
| 55885 | 327 | lemma disjuncts: "(\<exists>q \<in> set (disjuncts p). Ifm bbs bs q) \<longleftrightarrow> Ifm bbs bs p" | 
| 328 | by (induct p rule: disjuncts.induct) auto | |
| 23274 | 329 | |
| 50313 | 330 | lemma disjuncts_nb: | 
| 55999 | 331 | assumes "bound0 p" | 
| 50313 | 332 | shows "\<forall>q \<in> set (disjuncts p). bound0 q" | 
| 333 | proof - | |
| 55999 | 334 | from assms have "list_all bound0 (disjuncts p)" | 
| 50313 | 335 | by (induct p rule: disjuncts.induct) auto | 
| 55999 | 336 | then show ?thesis | 
| 337 | 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 | 338 | 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 | 339 | |
| 50313 | 340 | lemma disjuncts_qf: | 
| 55999 | 341 | assumes "qfree p" | 
| 50313 | 342 | shows "\<forall>q \<in> set (disjuncts p). qfree q" | 
| 343 | proof - | |
| 55999 | 344 | from assms have "list_all qfree (disjuncts p)" | 
| 50313 | 345 | by (induct p rule: disjuncts.induct) auto | 
| 55885 | 346 | then show ?thesis by (simp only: list_all_iff) | 
| 23274 | 347 | 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 | 348 | |
| 50313 | 349 | definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" | 
| 350 | where "DJ f p = 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 | 351 | |
| 50313 | 352 | lemma DJ: | 
| 55999 | 353 | assumes "\<forall>p q. f (Or p q) = Or (f p) (f q)" | 
| 354 | and "f F = F" | |
| 23274 | 355 | shows "Ifm bbs bs (DJ f p) = Ifm bbs bs (f p)" | 
| 50313 | 356 | proof - | 
| 55999 | 357 | have "Ifm bbs bs (DJ f p) \<longleftrightarrow> (\<exists>q \<in> set (disjuncts p). Ifm bbs bs (f q))" | 
| 50313 | 358 | by (simp add: DJ_def evaldjf_ex) | 
| 55999 | 359 | also from assms have "\<dots> = Ifm bbs bs (f p)" | 
| 360 | by (induct p rule: disjuncts.induct) auto | |
| 23274 | 361 | finally show ?thesis . | 
| 362 | 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 | 363 | |
| 50313 | 364 | lemma DJ_qf: | 
| 55999 | 365 | assumes "\<forall>p. qfree p \<longrightarrow> qfree (f p)" | 
| 23274 | 366 | shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) " | 
| 50313 | 367 | proof clarify | 
| 55844 | 368 | fix p | 
| 369 | assume qf: "qfree p" | |
| 370 | have th: "DJ f p = evaldjf f (disjuncts p)" | |
| 371 | by (simp add: DJ_def) | |
| 55925 | 372 | from disjuncts_qf[OF qf] have "\<forall>q \<in> set (disjuncts p). qfree q" . | 
| 55999 | 373 | with assms have th': "\<forall>q \<in> set (disjuncts p). qfree (f q)" | 
| 55844 | 374 | by blast | 
| 375 | from evaldjf_qf[OF th'] th show "qfree (DJ f p)" | |
| 376 | 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 | 377 | 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 | 378 | |
| 50313 | 379 | lemma DJ_qe: | 
| 55885 | 380 | assumes qe: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> Ifm bbs bs (qe p) = Ifm bbs bs (E p)" | 
| 381 | shows "\<forall>bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> Ifm bbs bs ((DJ qe p)) = Ifm bbs bs (E p)" | |
| 50313 | 382 | proof clarify | 
| 55844 | 383 | fix p :: fm | 
| 384 | fix bs | |
| 23274 | 385 | assume qf: "qfree p" | 
| 55844 | 386 | from qe have qth: "\<forall>p. qfree p \<longrightarrow> qfree (qe p)" | 
| 387 | by blast | |
| 55925 | 388 | from DJ_qf[OF qth] qf have qfth: "qfree (DJ qe p)" | 
| 55844 | 389 | by auto | 
| 50313 | 390 | have "Ifm bbs bs (DJ qe p) = (\<exists>q\<in> set (disjuncts p). Ifm bbs bs (qe q))" | 
| 23274 | 391 | by (simp add: DJ_def evaldjf_ex) | 
| 55999 | 392 | also have "\<dots> \<longleftrightarrow> (\<exists>q \<in> set (disjuncts p). Ifm bbs bs (E q))" | 
| 50313 | 393 | using qe disjuncts_qf[OF qf] by auto | 
| 55925 | 394 | also have "\<dots> \<longleftrightarrow> Ifm bbs bs (E p)" | 
| 50313 | 395 | by (induct p rule: disjuncts.induct) auto | 
| 396 | finally show "qfree (DJ qe p) \<and> Ifm bbs bs (DJ qe p) = Ifm bbs bs (E p)" | |
| 397 | using qfth by blast | |
| 23274 | 398 | qed | 
| 50313 | 399 | |
| 400 | ||
| 60533 | 401 | text \<open>Simplification\<close> | 
| 23274 | 402 | |
| 60533 | 403 | text \<open>Algebraic simplifications for nums\<close> | 
| 41837 | 404 | |
| 50313 | 405 | fun bnds :: "num \<Rightarrow> nat list" | 
| 406 | where | |
| 23274 | 407 | "bnds (Bound n) = [n]" | 
| 55921 | 408 | | "bnds (CN n c a) = n # bnds a" | 
| 41837 | 409 | | "bnds (Neg a) = bnds a" | 
| 55921 | 410 | | "bnds (Add a b) = bnds a @ bnds b" | 
| 55999 | 411 | | "bnds (Sub a b) = bnds a @ bnds b" | 
| 41837 | 412 | | "bnds (Mul i a) = bnds a" | 
| 413 | | "bnds a = []" | |
| 414 | ||
| 50313 | 415 | fun lex_ns:: "nat list \<Rightarrow> nat list \<Rightarrow> bool" | 
| 416 | where | |
| 55844 | 417 | "lex_ns [] ms \<longleftrightarrow> True" | 
| 418 | | "lex_ns ns [] \<longleftrightarrow> False" | |
| 419 | | "lex_ns (n # ns) (m # ms) \<longleftrightarrow> n < m \<or> (n = m \<and> lex_ns ns ms)" | |
| 23274 | 420 | |
| 50313 | 421 | definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" | 
| 422 | where "lex_bnd t s = lex_ns (bnds t) (bnds s)" | |
| 423 | ||
| 424 | consts numadd:: "num \<times> num \<Rightarrow> num" | |
| 55921 | 425 | recdef numadd "measure (\<lambda>(t, s). num_size t + num_size s)" | 
| 55981 | 426 | "numadd (CN n1 c1 r1, CN n2 c2 r2) = | 
| 50313 | 427 | (if n1 = n2 then | 
| 55981 | 428 | let c = c1 + c2 | 
| 429 | in if c = 0 then numadd (r1, r2) else CN n1 c (numadd (r1, r2)) | |
| 55921 | 430 | else if n1 \<le> n2 then CN n1 c1 (numadd (r1, Add (Mul c2 (Bound n2)) r2)) | 
| 50313 | 431 | else CN n2 c2 (numadd (Add (Mul c1 (Bound n1)) r1, r2)))" | 
| 432 | "numadd (CN n1 c1 r1, t) = CN n1 c1 (numadd (r1, t))" | |
| 433 | "numadd (t, CN n2 c2 r2) = CN n2 c2 (numadd (t, r2))" | |
| 434 | "numadd (C b1, C b2) = C (b1 + b2)" | |
| 435 | "numadd (a, b) = Add a b" | |
| 23274 | 436 | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 437 | (*function (sequential) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 438 | numadd :: "num \<Rightarrow> num \<Rightarrow> num" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 439 | where | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 440 | "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 | 441 | (if n1 = n2 then (let c = c1 + c2 | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 442 | in (if c = 0 then numadd r1 r2 else | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 443 | Add (Mul c (Bound n1)) (numadd r1 r2))) | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 444 | else if n1 \<le> n2 then | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 445 | 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 | 446 | else | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 447 | 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 | 448 | | "numadd (Add (Mul c1 (Bound n1)) r1) t = | 
| 50313 | 449 | Add (Mul c1 (Bound n1)) (numadd r1 t)" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 450 | | "numadd t (Add (Mul c2 (Bound n2)) r2) = | 
| 50313 | 451 | Add (Mul c2 (Bound n2)) (numadd t r2)" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 452 | | "numadd (C b1) (C b2) = C (b1 + b2)" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 453 | | "numadd a b = Add a b" | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 454 | apply pat_completeness apply auto*) | 
| 50313 | 455 | |
| 23274 | 456 | lemma numadd: "Inum bs (numadd (t,s)) = Inum bs (Add t s)" | 
| 55885 | 457 | apply (induct t s rule: numadd.induct) | 
| 458 | apply (simp_all add: Let_def) | |
| 60708 | 459 | subgoal for n1 c1 r1 n2 c2 r2 | 
| 460 | apply (cases "c1 + c2 = 0") | |
| 461 | apply (cases "n1 \<le> n2") | |
| 462 | apply simp_all | |
| 463 | apply (cases "n1 = n2") | |
| 464 | apply (simp_all add: algebra_simps) | |
| 465 | apply (simp add: distrib_right[symmetric]) | |
| 466 | done | |
| 50313 | 467 | done | 
| 23274 | 468 | |
| 50313 | 469 | lemma numadd_nb: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numadd (t, s))" | 
| 470 | by (induct t s rule: numadd.induct) (auto simp add: Let_def) | |
| 23274 | 471 | |
| 50313 | 472 | fun nummul :: "int \<Rightarrow> num \<Rightarrow> num" | 
| 473 | where | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 474 | "nummul i (C j) = C (i * j)" | 
| 55844 | 475 | | "nummul i (CN n c t) = CN n (c * i) (nummul i t)" | 
| 41837 | 476 | | "nummul i t = Mul i t" | 
| 23274 | 477 | |
| 50313 | 478 | lemma nummul: "Inum bs (nummul i t) = Inum bs (Mul i t)" | 
| 479 | by (induct t arbitrary: i rule: nummul.induct) (auto simp add: algebra_simps numadd) | |
| 23274 | 480 | |
| 50313 | 481 | lemma nummul_nb: "numbound0 t \<Longrightarrow> numbound0 (nummul i t)" | 
| 482 | by (induct t arbitrary: i rule: nummul.induct) (auto simp add: numadd_nb) | |
| 23274 | 483 | |
| 50313 | 484 | definition numneg :: "num \<Rightarrow> num" | 
| 485 | where "numneg t = nummul (- 1) t" | |
| 23274 | 486 | |
| 50313 | 487 | definition numsub :: "num \<Rightarrow> num \<Rightarrow> num" | 
| 488 | where "numsub s t = (if s = t then C 0 else numadd (s, numneg t))" | |
| 23274 | 489 | |
| 490 | lemma numneg: "Inum bs (numneg t) = Inum bs (Neg t)" | |
| 50313 | 491 | using numneg_def nummul by simp | 
| 23274 | 492 | |
| 493 | lemma numneg_nb: "numbound0 t \<Longrightarrow> numbound0 (numneg t)" | |
| 50313 | 494 | using numneg_def nummul_nb by simp | 
| 23274 | 495 | |
| 496 | lemma numsub: "Inum bs (numsub a b) = Inum bs (Sub a b)" | |
| 50313 | 497 | using numneg numadd numsub_def by simp | 
| 23274 | 498 | |
| 50313 | 499 | lemma numsub_nb: "numbound0 t \<Longrightarrow> numbound0 s \<Longrightarrow> numbound0 (numsub t s)" | 
| 500 | using numsub_def numadd_nb numneg_nb by simp | |
| 23274 | 501 | |
| 50313 | 502 | fun simpnum :: "num \<Rightarrow> num" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 503 | where | 
| 23274 | 504 | "simpnum (C j) = C j" | 
| 50313 | 505 | | "simpnum (Bound n) = CN n 1 (C 0)" | 
| 506 | | "simpnum (Neg t) = numneg (simpnum t)" | |
| 507 | | "simpnum (Add t s) = numadd (simpnum t, simpnum s)" | |
| 508 | | "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" | |
| 509 | | "simpnum (Mul i t) = (if i = 0 then C 0 else nummul i (simpnum t))" | |
| 510 | | "simpnum t = t" | |
| 23274 | 511 | |
| 512 | lemma simpnum_ci: "Inum bs (simpnum t) = Inum bs t" | |
| 50313 | 513 | by (induct t rule: simpnum.induct) (auto simp add: numneg numadd numsub nummul) | 
| 23274 | 514 | |
| 50313 | 515 | lemma simpnum_numbound0: "numbound0 t \<Longrightarrow> numbound0 (simpnum t)" | 
| 516 | by (induct t rule: simpnum.induct) (auto simp add: numadd_nb numsub_nb nummul_nb numneg_nb) | |
| 23274 | 517 | |
| 50313 | 518 | fun not :: "fm \<Rightarrow> fm" | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 519 | where | 
| 23274 | 520 | "not (NOT p) = p" | 
| 50313 | 521 | | "not T = F" | 
| 522 | | "not F = T" | |
| 523 | | "not p = NOT p" | |
| 524 | ||
| 23274 | 525 | lemma not: "Ifm bbs bs (not p) = Ifm bbs bs (NOT p)" | 
| 41807 | 526 | by (cases p) auto | 
| 50313 | 527 | |
| 23274 | 528 | lemma not_qf: "qfree p \<Longrightarrow> qfree (not p)" | 
| 41807 | 529 | by (cases p) auto | 
| 50313 | 530 | |
| 23274 | 531 | lemma not_bn: "bound0 p \<Longrightarrow> bound0 (not p)" | 
| 41807 | 532 | by (cases p) auto | 
| 23274 | 533 | |
| 50313 | 534 | definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" | 
| 60708 | 535 | where | 
| 536 | "conj p q = | |
| 537 | (if p = F \<or> q = F then F | |
| 538 | else if p = T then q | |
| 539 | else if q = T then p | |
| 540 | else And p q)" | |
| 50313 | 541 | |
| 23274 | 542 | lemma conj: "Ifm bbs bs (conj p q) = Ifm bbs bs (And p q)" | 
| 55844 | 543 | by (cases "p = F \<or> q = F", simp_all add: conj_def) (cases p, simp_all) | 
| 23274 | 544 | |
| 50313 | 545 | lemma conj_qf: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (conj p q)" | 
| 546 | using conj_def by auto | |
| 23274 | 547 | |
| 50313 | 548 | lemma conj_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (conj p q)" | 
| 549 | using conj_def by auto | |
| 550 | ||
| 551 | definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" | |
| 60708 | 552 | where | 
| 553 | "disj p q = | |
| 554 | (if p = T \<or> q = T then T | |
| 555 | else if p = F then q | |
| 556 | else if q = F then p | |
| 557 | else Or p q)" | |
| 23274 | 558 | |
| 559 | lemma disj: "Ifm bbs bs (disj p q) = Ifm bbs bs (Or p q)" | |
| 55885 | 560 | by (cases "p = T \<or> q = T", simp_all add: disj_def) (cases p, simp_all) | 
| 50313 | 561 | |
| 55844 | 562 | lemma disj_qf: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (disj p q)" | 
| 50313 | 563 | using disj_def by auto | 
| 564 | ||
| 55844 | 565 | lemma disj_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (disj p q)" | 
| 50313 | 566 | using disj_def by auto | 
| 23274 | 567 | |
| 50313 | 568 | definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" | 
| 60708 | 569 | where | 
| 570 | "imp p q = | |
| 571 | (if p = F \<or> q = T then T | |
| 572 | else if p = T then q | |
| 573 | else if q = F then not p | |
| 574 | else Imp p q)" | |
| 50313 | 575 | |
| 23274 | 576 | lemma imp: "Ifm bbs bs (imp p q) = Ifm bbs bs (Imp p q)" | 
| 55844 | 577 | by (cases "p = F \<or> q = T", simp_all add: imp_def, cases p) (simp_all add: not) | 
| 50313 | 578 | |
| 579 | lemma imp_qf: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (imp p q)" | |
| 55844 | 580 | using imp_def by (cases "p = F \<or> q = T", simp_all add: imp_def, cases p) (simp_all add: not_qf) | 
| 50313 | 581 | |
| 582 | lemma imp_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (imp p q)" | |
| 55844 | 583 | using imp_def by (cases "p = F \<or> q = T", simp_all add: imp_def, cases p) simp_all | 
| 23274 | 584 | |
| 50313 | 585 | definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" | 
| 55885 | 586 | where | 
| 587 | "iff p q = | |
| 588 | (if p = q then T | |
| 589 | else if p = not q \<or> not p = q then F | |
| 590 | else if p = F then not q | |
| 591 | else if q = F then not p | |
| 592 | else if p = T then q | |
| 593 | else if q = T then p | |
| 594 | else Iff p q)" | |
| 50313 | 595 | |
| 23274 | 596 | lemma iff: "Ifm bbs bs (iff p q) = Ifm bbs bs (Iff p q)" | 
| 55885 | 597 | by (unfold iff_def, cases "p = q", simp, cases "p = not q", simp add: not) | 
| 598 | (cases "not p = q", auto simp add: not) | |
| 50313 | 599 | |
| 55885 | 600 | lemma iff_qf: "qfree p \<Longrightarrow> qfree q \<Longrightarrow> qfree (iff p q)" | 
| 601 | by (unfold iff_def, cases "p = q", auto simp add: not_qf) | |
| 23274 | 602 | |
| 55885 | 603 | lemma iff_nb: "bound0 p \<Longrightarrow> bound0 q \<Longrightarrow> bound0 (iff p q)" | 
| 604 | using iff_def by (unfold iff_def, cases "p = q", auto simp add: not_bn) | |
| 50313 | 605 | |
| 606 | function (sequential) simpfm :: "fm \<Rightarrow> fm" | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 607 | where | 
| 23274 | 608 | "simpfm (And p q) = conj (simpfm p) (simpfm q)" | 
| 50313 | 609 | | "simpfm (Or p q) = disj (simpfm p) (simpfm q)" | 
| 610 | | "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" | |
| 611 | | "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" | |
| 612 | | "simpfm (NOT p) = not (simpfm p)" | |
| 55921 | 613 | | "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v < 0 then T else F | _ \<Rightarrow> Lt a')" | 
| 614 | | "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')" | |
| 615 | | "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v > 0 then T else F | _ \<Rightarrow> Gt a')" | |
| 616 | | "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')" | |
| 617 | | "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if v = 0 then T else F | _ \<Rightarrow> Eq a')" | |
| 618 | | "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')" | |
| 50313 | 619 | | "simpfm (Dvd i a) = | 
| 55885 | 620 | (if i = 0 then simpfm (Eq a) | 
| 621 | else if abs i = 1 then T | |
| 55921 | 622 | else let a' = simpnum a in case a' of C v \<Rightarrow> if i dvd v then T else F | _ \<Rightarrow> Dvd i a')" | 
| 50313 | 623 | | "simpfm (NDvd i a) = | 
| 55885 | 624 | (if i = 0 then simpfm (NEq a) | 
| 625 | else if abs i = 1 then F | |
| 55921 | 626 | 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')" | 
| 50313 | 627 | | "simpfm p = p" | 
| 628 | by pat_completeness auto | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 629 | termination by (relation "measure fmsize") auto | 
| 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 630 | |
| 23274 | 631 | lemma simpfm: "Ifm bbs bs (simpfm p) = Ifm bbs bs p" | 
| 55844 | 632 | proof (induct p rule: simpfm.induct) | 
| 50313 | 633 | case (6 a) | 
| 634 | let ?sa = "simpnum a" | |
| 55925 | 635 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" | 
| 636 | by simp | |
| 60708 | 637 | consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast | 
| 638 | then show ?case | |
| 639 | proof cases | |
| 640 | case 1 | |
| 641 | with sa show ?thesis by simp | |
| 642 | next | |
| 643 | case 2 | |
| 644 | with sa show ?thesis by (cases ?sa) (simp_all add: Let_def) | |
| 645 | 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 | 646 | next | 
| 50313 | 647 | case (7 a) | 
| 648 | let ?sa = "simpnum a" | |
| 55925 | 649 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" | 
| 650 | by simp | |
| 60708 | 651 | consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast | 
| 652 | then show ?case | |
| 653 | proof cases | |
| 654 | case 1 | |
| 655 | with sa show ?thesis by simp | |
| 656 | next | |
| 657 | case 2 | |
| 658 | with sa show ?thesis by (cases ?sa) (simp_all add: Let_def) | |
| 659 | qed | |
| 23274 | 660 | next | 
| 50313 | 661 | case (8 a) | 
| 662 | let ?sa = "simpnum a" | |
| 55925 | 663 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" | 
| 664 | by simp | |
| 60708 | 665 | consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast | 
| 666 | then show ?case | |
| 667 | proof cases | |
| 668 | case 1 | |
| 669 | with sa show ?thesis by simp | |
| 670 | next | |
| 671 | case 2 | |
| 672 | with sa show ?thesis by (cases ?sa) (simp_all add: Let_def) | |
| 673 | qed | |
| 23274 | 674 | next | 
| 50313 | 675 | case (9 a) | 
| 676 | let ?sa = "simpnum a" | |
| 55925 | 677 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" | 
| 678 | by simp | |
| 60708 | 679 | consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast | 
| 680 | then show ?case | |
| 681 | proof cases | |
| 682 | case 1 | |
| 683 | with sa show ?thesis by simp | |
| 684 | next | |
| 685 | case 2 | |
| 686 | with sa show ?thesis by (cases ?sa) (simp_all add: Let_def) | |
| 687 | qed | |
| 23274 | 688 | next | 
| 50313 | 689 | case (10 a) | 
| 690 | let ?sa = "simpnum a" | |
| 55925 | 691 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" | 
| 692 | by simp | |
| 60708 | 693 | consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast | 
| 694 | then show ?case | |
| 695 | proof cases | |
| 696 | case 1 | |
| 697 | with sa show ?thesis by simp | |
| 698 | next | |
| 699 | case 2 | |
| 700 | with sa show ?thesis by (cases ?sa) (simp_all add: Let_def) | |
| 701 | qed | |
| 23274 | 702 | next | 
| 50313 | 703 | case (11 a) | 
| 704 | let ?sa = "simpnum a" | |
| 55925 | 705 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" | 
| 706 | by simp | |
| 60708 | 707 | consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast | 
| 708 | then show ?case | |
| 709 | proof cases | |
| 710 | case 1 | |
| 711 | with sa show ?thesis by simp | |
| 712 | next | |
| 713 | case 2 | |
| 714 | with sa show ?thesis by (cases ?sa) (simp_all add: Let_def) | |
| 715 | qed | |
| 23274 | 716 | next | 
| 50313 | 717 | case (12 i a) | 
| 718 | let ?sa = "simpnum a" | |
| 55925 | 719 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" | 
| 720 | by simp | |
| 60708 | 721 | consider "i = 0" | "abs i = 1" | "i \<noteq> 0" "abs i \<noteq> 1" by blast | 
| 722 | then show ?case | |
| 723 | proof cases | |
| 724 | case 1 | |
| 725 | then show ?thesis | |
| 726 | using "12.hyps" by (simp add: dvd_def Let_def) | |
| 727 | next | |
| 728 | case 2 | |
| 729 | with one_dvd[of "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"] | |
| 730 | show ?thesis | |
| 55925 | 731 | apply (cases "i = 0") | 
| 732 | apply (simp_all add: Let_def) | |
| 733 | apply (cases "i > 0") | |
| 734 | apply simp_all | |
| 50313 | 735 | done | 
| 60708 | 736 | next | 
| 737 | case i: 3 | |
| 738 | consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast | |
| 739 | then show ?thesis | |
| 740 | proof cases | |
| 741 | case 1 | |
| 742 | with sa[symmetric] i show ?thesis | |
| 55925 | 743 | by (cases "abs i = 1") auto | 
| 60708 | 744 | next | 
| 745 | case 2 | |
| 55925 | 746 | then have "simpfm (Dvd i a) = Dvd i ?sa" | 
| 60708 | 747 | using i by (cases ?sa) (auto simp add: Let_def) | 
| 748 | with sa show ?thesis by simp | |
| 749 | qed | |
| 750 | qed | |
| 50313 | 751 | next | 
| 752 | case (13 i a) | |
| 55925 | 753 | let ?sa = "simpnum a" | 
| 754 | from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" | |
| 755 | by simp | |
| 60708 | 756 | consider "i = 0" | "abs i = 1" | "i \<noteq> 0" "abs i \<noteq> 1" by blast | 
| 757 | then show ?case | |
| 758 | proof cases | |
| 759 | case 1 | |
| 760 | then show ?thesis | |
| 761 | using "13.hyps" by (simp add: dvd_def Let_def) | |
| 762 | next | |
| 763 | case 2 | |
| 764 | with one_dvd[of "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"] | |
| 765 | show ?thesis | |
| 55925 | 766 | apply (cases "i = 0") | 
| 767 | apply (simp_all add: Let_def) | |
| 768 | apply (cases "i > 0") | |
| 769 | apply simp_all | |
| 50313 | 770 | done | 
| 60708 | 771 | next | 
| 772 | case i: 3 | |
| 773 | consider v where "?sa = C v" | "\<not> (\<exists>v. ?sa = C v)" by blast | |
| 774 | then show ?thesis | |
| 775 | proof cases | |
| 776 | case 1 | |
| 777 | with sa[symmetric] i show ?thesis | |
| 778 | by (cases "abs i = 1") auto | |
| 779 | next | |
| 780 | case 2 | |
| 55925 | 781 | then have "simpfm (NDvd i a) = NDvd i ?sa" | 
| 60708 | 782 | using i by (cases ?sa) (auto simp add: Let_def) | 
| 783 | with sa show ?thesis by simp | |
| 784 | qed | |
| 785 | qed | |
| 50313 | 786 | qed (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 | 787 | |
| 23274 | 788 | lemma simpfm_bound0: "bound0 p \<Longrightarrow> bound0 (simpfm p)" | 
| 50313 | 789 | proof (induct p rule: simpfm.induct) | 
| 55925 | 790 | case (6 a) | 
| 791 | then have nb: "numbound0 a" by simp | |
| 55885 | 792 | then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | 
| 793 | then show ?case by (cases "simpnum a") (auto simp add: Let_def) | |
| 23274 | 794 | next | 
| 55925 | 795 | case (7 a) | 
| 796 | then have nb: "numbound0 a" by simp | |
| 55885 | 797 | then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | 
| 798 | then show ?case by (cases "simpnum a") (auto simp add: Let_def) | |
| 23274 | 799 | next | 
| 55925 | 800 | case (8 a) | 
| 801 | then have nb: "numbound0 a" by simp | |
| 55885 | 802 | then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | 
| 803 | then show ?case by (cases "simpnum a") (auto simp add: Let_def) | |
| 23274 | 804 | next | 
| 55925 | 805 | case (9 a) | 
| 806 | then have nb: "numbound0 a" by simp | |
| 55885 | 807 | then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | 
| 808 | then show ?case by (cases "simpnum a") (auto simp add: Let_def) | |
| 23274 | 809 | next | 
| 55925 | 810 | case (10 a) | 
| 811 | then have nb: "numbound0 a" by simp | |
| 55885 | 812 | then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | 
| 813 | then show ?case by (cases "simpnum a") (auto simp add: Let_def) | |
| 23274 | 814 | next | 
| 55925 | 815 | case (11 a) | 
| 816 | then have nb: "numbound0 a" by simp | |
| 55885 | 817 | then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | 
| 818 | then show ?case by (cases "simpnum a") (auto simp add: Let_def) | |
| 23274 | 819 | next | 
| 55925 | 820 | case (12 i a) | 
| 821 | then have nb: "numbound0 a" by simp | |
| 55885 | 822 | then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | 
| 823 | then show ?case by (cases "simpnum a") (auto simp add: Let_def) | |
| 23274 | 824 | next | 
| 55925 | 825 | case (13 i a) | 
| 826 | then have nb: "numbound0 a" by simp | |
| 55885 | 827 | then have "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) | 
| 828 | then show ?case by (cases "simpnum a") (auto simp add: Let_def) | |
| 50313 | 829 | 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 | 830 | |
| 23274 | 831 | lemma simpfm_qf: "qfree p \<Longrightarrow> qfree (simpfm p)" | 
| 60708 | 832 | apply (induct p rule: simpfm.induct) | 
| 833 | apply (auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def) | |
| 834 | apply (case_tac "simpnum a", auto)+ | |
| 835 | done | |
| 23274 | 836 | |
| 60533 | 837 | text \<open>Generic quantifier elimination\<close> | 
| 50313 | 838 | function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm" | 
| 839 | where | |
| 840 | "qelim (E p) = (\<lambda>qe. DJ qe (qelim p qe))" | |
| 841 | | "qelim (A p) = (\<lambda>qe. not (qe ((qelim (NOT p) qe))))" | |
| 842 | | "qelim (NOT p) = (\<lambda>qe. not (qelim p qe))" | |
| 843 | | "qelim (And p q) = (\<lambda>qe. conj (qelim p qe) (qelim q qe))" | |
| 844 | | "qelim (Or p q) = (\<lambda>qe. disj (qelim p qe) (qelim q qe))" | |
| 845 | | "qelim (Imp p q) = (\<lambda>qe. imp (qelim p qe) (qelim q qe))" | |
| 846 | | "qelim (Iff p q) = (\<lambda>qe. iff (qelim p qe) (qelim q qe))" | |
| 847 | | "qelim p = (\<lambda>y. simpfm p)" | |
| 848 | by pat_completeness auto | |
| 41837 | 849 | termination by (relation "measure fmsize") auto | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 850 | |
| 23274 | 851 | lemma qelim_ci: | 
| 55885 | 852 | assumes qe_inv: "\<forall>bs p. qfree p \<longrightarrow> qfree (qe p) \<and> Ifm bbs bs (qe p) = Ifm bbs bs (E p)" | 
| 853 | shows "\<And>bs. qfree (qelim p qe) \<and> Ifm bbs bs (qelim p qe) = Ifm bbs bs p" | |
| 50313 | 854 | using qe_inv DJ_qe[OF qe_inv] | 
| 55964 | 855 | by (induct p rule: qelim.induct) | 
| 856 | (auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf | |
| 857 | simpfm simpfm_qf simp del: simpfm.simps) | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 858 | |
| 60533 | 859 | text \<open>Linearity for fm where Bound 0 ranges over @{text "\<int>"}\<close>
 | 
| 50313 | 860 | |
| 60533 | 861 | fun zsplit0 :: "num \<Rightarrow> int \<times> num" -- \<open>splits the bounded from the unbounded part\<close> | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 862 | where | 
| 55885 | 863 | "zsplit0 (C c) = (0, C c)" | 
| 864 | | "zsplit0 (Bound n) = (if n = 0 then (1, C 0) else (0, Bound n))" | |
| 50313 | 865 | | "zsplit0 (CN n i a) = | 
| 55885 | 866 | (let (i', a') = zsplit0 a | 
| 867 | in if n = 0 then (i + i', a') else (i', CN n i a'))" | |
| 868 | | "zsplit0 (Neg a) = (let (i', a') = zsplit0 a in (-i', Neg a'))" | |
| 869 | | "zsplit0 (Add a b) = | |
| 870 | (let | |
| 871 | (ia, a') = zsplit0 a; | |
| 872 | (ib, b') = zsplit0 b | |
| 873 | in (ia + ib, Add a' b'))" | |
| 874 | | "zsplit0 (Sub a b) = | |
| 875 | (let | |
| 876 | (ia, a') = zsplit0 a; | |
| 877 | (ib, b') = zsplit0 b | |
| 878 | in (ia - ib, Sub a' b'))" | |
| 879 | | "zsplit0 (Mul i a) = (let (i', a') = zsplit0 a in (i*i', Mul i a'))" | |
| 23274 | 880 | |
| 881 | lemma zsplit0_I: | |
| 55964 | 882 | "\<And>n a. zsplit0 t = (n, a) \<Longrightarrow> | 
| 55921 | 883 | (Inum ((x::int) # bs) (CN 0 n a) = Inum (x # bs) t) \<and> numbound0 a" | 
| 50313 | 884 | (is "\<And>n a. ?S t = (n,a) \<Longrightarrow> (?I x (CN 0 n a) = ?I x t) \<and> ?N a") | 
| 885 | proof (induct t rule: zsplit0.induct) | |
| 55844 | 886 | case (1 c n a) | 
| 887 | then show ?case by auto | |
| 23274 | 888 | next | 
| 55844 | 889 | case (2 m n a) | 
| 890 | then show ?case by (cases "m = 0") auto | |
| 23274 | 891 | next | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 892 | case (3 m i a n a') | 
| 23274 | 893 | let ?j = "fst (zsplit0 a)" | 
| 894 | let ?b = "snd (zsplit0 a)" | |
| 55844 | 895 | have abj: "zsplit0 a = (?j, ?b)" by simp | 
| 60708 | 896 | show ?case | 
| 897 | proof (cases "m = 0") | |
| 898 | case False | |
| 899 | with 3(1)[OF abj] 3(2) show ?thesis | |
| 55844 | 900 | by (auto simp add: Let_def split_def) | 
| 60708 | 901 | next | 
| 902 | case m: True | |
| 55964 | 903 | with abj have th: "a' = ?b \<and> n = i + ?j" | 
| 904 | using 3 by (simp add: Let_def split_def) | |
| 60708 | 905 | from abj 3 m have th2: "(?I x (CN 0 ?j ?b) = ?I x a) \<and> ?N ?b" | 
| 55844 | 906 | by blast | 
| 55964 | 907 | from th have "?I x (CN 0 n a') = ?I x (CN 0 (i + ?j) ?b)" | 
| 55844 | 908 | by simp | 
| 909 | also from th2 have "\<dots> = ?I x (CN 0 i (CN 0 ?j ?b))" | |
| 910 | by (simp add: distrib_right) | |
| 911 | finally have "?I x (CN 0 n a') = ?I x (CN 0 i a)" | |
| 912 | using th2 by simp | |
| 60708 | 913 | with th2 th m show ?thesis | 
| 55844 | 914 | by blast | 
| 60708 | 915 | qed | 
| 23274 | 916 | next | 
| 917 | case (4 t n a) | |
| 918 | let ?nt = "fst (zsplit0 t)" | |
| 919 | let ?at = "snd (zsplit0 t)" | |
| 55964 | 920 | have abj: "zsplit0 t = (?nt, ?at)" | 
| 921 | by simp | |
| 922 | then have th: "a = Neg ?at \<and> n = - ?nt" | |
| 55844 | 923 | using 4 by (simp add: Let_def split_def) | 
| 924 | from abj 4 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" | |
| 925 | by blast | |
| 926 | from th2[simplified] th[simplified] show ?case | |
| 927 | by simp | |
| 23274 | 928 | next | 
| 929 | case (5 s t n a) | |
| 930 | let ?ns = "fst (zsplit0 s)" | |
| 931 | let ?as = "snd (zsplit0 s)" | |
| 932 | let ?nt = "fst (zsplit0 t)" | |
| 933 | let ?at = "snd (zsplit0 t)" | |
| 55844 | 934 | have abjs: "zsplit0 s = (?ns, ?as)" | 
| 935 | by simp | |
| 936 | moreover have abjt: "zsplit0 t = (?nt, ?at)" | |
| 937 | by simp | |
| 55964 | 938 | ultimately have th: "a = Add ?as ?at \<and> n = ?ns + ?nt" | 
| 55844 | 939 | using 5 by (simp add: Let_def split_def) | 
| 55964 | 940 | from abjs[symmetric] have bluddy: "\<exists>x y. (x, y) = zsplit0 s" | 
| 55844 | 941 | by blast | 
| 942 | from 5 have "(\<exists>x y. (x, y) = zsplit0 s) \<longrightarrow> | |
| 943 | (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (x # bs) (CN 0 xa xb) = Inum (x # bs) t \<and> numbound0 xb)" | |
| 944 | by auto | |
| 945 | with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" | |
| 946 | by blast | |
| 947 | from abjs 5 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" | |
| 948 | by blast | |
| 50313 | 949 | from th3[simplified] th2[simplified] th[simplified] show ?case | 
| 49962 
a8cc904a6820
Renamed {left,right}_distrib to distrib_{right,left}.
 webertj parents: 
48891diff
changeset | 950 | by (simp add: distrib_right) | 
| 23274 | 951 | next | 
| 952 | case (6 s t n a) | |
| 953 | let ?ns = "fst (zsplit0 s)" | |
| 954 | let ?as = "snd (zsplit0 s)" | |
| 955 | let ?nt = "fst (zsplit0 t)" | |
| 956 | let ?at = "snd (zsplit0 t)" | |
| 55844 | 957 | have abjs: "zsplit0 s = (?ns, ?as)" | 
| 958 | by simp | |
| 959 | moreover have abjt: "zsplit0 t = (?nt, ?at)" | |
| 960 | by simp | |
| 55964 | 961 | ultimately have th: "a = Sub ?as ?at \<and> n = ?ns - ?nt" | 
| 55844 | 962 | using 6 by (simp add: Let_def split_def) | 
| 55964 | 963 | from abjs[symmetric] have bluddy: "\<exists>x y. (x, y) = zsplit0 s" | 
| 55844 | 964 | by blast | 
| 50313 | 965 | from 6 have "(\<exists>x y. (x,y) = zsplit0 s) \<longrightarrow> | 
| 966 | (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (x # bs) (CN 0 xa xb) = Inum (x # bs) t \<and> numbound0 xb)" | |
| 967 | by auto | |
| 55844 | 968 | with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" | 
| 969 | by blast | |
| 970 | from abjs 6 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" | |
| 971 | by blast | |
| 50313 | 972 | from th3[simplified] th2[simplified] th[simplified] show ?case | 
| 23274 | 973 | by (simp add: left_diff_distrib) | 
| 974 | next | |
| 975 | case (7 i t n a) | |
| 976 | let ?nt = "fst (zsplit0 t)" | |
| 977 | let ?at = "snd (zsplit0 t)" | |
| 55844 | 978 | have abj: "zsplit0 t = (?nt,?at)" | 
| 979 | by simp | |
| 55964 | 980 | then have th: "a = Mul i ?at \<and> n = i * ?nt" | 
| 55844 | 981 | using 7 by (simp add: Let_def split_def) | 
| 982 | from abj 7 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" | |
| 983 | by blast | |
| 984 | then have "?I x (Mul i t) = i * ?I x (CN 0 ?nt ?at)" | |
| 985 | by simp | |
| 986 | also have "\<dots> = ?I x (CN 0 (i*?nt) (Mul i ?at))" | |
| 987 | by (simp add: distrib_left) | |
| 988 | finally show ?case using th th2 | |
| 989 | 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 | 990 | 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 | 991 | |
| 60533 | 992 | consts iszlfm :: "fm \<Rightarrow> bool" -- \<open>Linearity test for fm\<close> | 
| 23274 | 993 | recdef iszlfm "measure size" | 
| 55885 | 994 | "iszlfm (And p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q" | 
| 995 | "iszlfm (Or p q) \<longleftrightarrow> iszlfm p \<and> iszlfm q" | |
| 996 | "iszlfm (Eq (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e" | |
| 997 | "iszlfm (NEq (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e" | |
| 998 | "iszlfm (Lt (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e" | |
| 999 | "iszlfm (Le (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e" | |
| 1000 | "iszlfm (Gt (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e" | |
| 1001 | "iszlfm (Ge (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> numbound0 e" | |
| 1002 | "iszlfm (Dvd i (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> i > 0 \<and> numbound0 e" | |
| 1003 | "iszlfm (NDvd i (CN 0 c e)) \<longleftrightarrow> c > 0 \<and> i > 0 \<and> numbound0 e" | |
| 1004 | "iszlfm p \<longleftrightarrow> 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 | 1005 | |
| 23274 | 1006 | lemma zlin_qfree: "iszlfm p \<Longrightarrow> qfree p" | 
| 1007 | 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 | 1008 | |
| 60533 | 1009 | consts zlfm :: "fm \<Rightarrow> fm" -- \<open>Linearity transformation for fm\<close> | 
| 23274 | 1010 | recdef zlfm "measure fmsize" | 
| 1011 | "zlfm (And p q) = And (zlfm p) (zlfm q)" | |
| 1012 | "zlfm (Or p q) = Or (zlfm p) (zlfm q)" | |
| 1013 | "zlfm (Imp p q) = Or (zlfm (NOT p)) (zlfm q)" | |
| 1014 | "zlfm (Iff p q) = Or (And (zlfm p) (zlfm q)) (And (zlfm (NOT p)) (zlfm (NOT q)))" | |
| 55964 | 1015 | "zlfm (Lt a) = | 
| 1016 | (let (c, r) = zsplit0 a in | |
| 1017 | if c = 0 then Lt r else | |
| 1018 | if c > 0 then (Lt (CN 0 c r)) | |
| 1019 | else Gt (CN 0 (- c) (Neg r)))" | |
| 1020 | "zlfm (Le a) = | |
| 1021 | (let (c, r) = zsplit0 a in | |
| 1022 | if c = 0 then Le r | |
| 1023 | else if c > 0 then Le (CN 0 c r) | |
| 1024 | else Ge (CN 0 (- c) (Neg r)))" | |
| 1025 | "zlfm (Gt a) = | |
| 1026 | (let (c, r) = zsplit0 a in | |
| 1027 | if c = 0 then Gt r else | |
| 1028 | if c > 0 then Gt (CN 0 c r) | |
| 1029 | else Lt (CN 0 (- c) (Neg r)))" | |
| 1030 | "zlfm (Ge a) = | |
| 1031 | (let (c, r) = zsplit0 a in | |
| 1032 | if c = 0 then Ge r | |
| 1033 | else if c > 0 then Ge (CN 0 c r) | |
| 1034 | else Le (CN 0 (- c) (Neg r)))" | |
| 1035 | "zlfm (Eq a) = | |
| 1036 | (let (c, r) = zsplit0 a in | |
| 1037 | if c = 0 then Eq r | |
| 1038 | else if c > 0 then Eq (CN 0 c r) | |
| 1039 | else Eq (CN 0 (- c) (Neg r)))" | |
| 1040 | "zlfm (NEq a) = | |
| 1041 | (let (c, r) = zsplit0 a in | |
| 1042 | if c = 0 then NEq r | |
| 1043 | else if c > 0 then NEq (CN 0 c r) | |
| 1044 | else NEq (CN 0 (- c) (Neg r)))" | |
| 1045 | "zlfm (Dvd i a) = | |
| 1046 | (if i = 0 then zlfm (Eq a) | |
| 1047 | else | |
| 1048 | let (c, r) = zsplit0 a in | |
| 1049 | if c = 0 then Dvd (abs i) r | |
| 1050 | else if c > 0 then Dvd (abs i) (CN 0 c r) | |
| 1051 | else Dvd (abs i) (CN 0 (- c) (Neg r)))" | |
| 1052 | "zlfm (NDvd i a) = | |
| 1053 | (if i = 0 then zlfm (NEq a) | |
| 1054 | else | |
| 1055 | let (c, r) = zsplit0 a in | |
| 1056 | if c = 0 then NDvd (abs i) r | |
| 1057 | else if c > 0 then NDvd (abs i) (CN 0 c r) | |
| 1058 | else NDvd (abs i) (CN 0 (- c) (Neg r)))" | |
| 23274 | 1059 | "zlfm (NOT (And p q)) = Or (zlfm (NOT p)) (zlfm (NOT q))" | 
| 1060 | "zlfm (NOT (Or p q)) = And (zlfm (NOT p)) (zlfm (NOT q))" | |
| 1061 | "zlfm (NOT (Imp p q)) = And (zlfm p) (zlfm (NOT q))" | |
| 1062 | "zlfm (NOT (Iff p q)) = Or (And(zlfm p) (zlfm(NOT q))) (And (zlfm(NOT p)) (zlfm q))" | |
| 1063 | "zlfm (NOT (NOT p)) = zlfm p" | |
| 1064 | "zlfm (NOT T) = F" | |
| 1065 | "zlfm (NOT F) = T" | |
| 1066 | "zlfm (NOT (Lt a)) = zlfm (Ge a)" | |
| 1067 | "zlfm (NOT (Le a)) = zlfm (Gt a)" | |
| 1068 | "zlfm (NOT (Gt a)) = zlfm (Le a)" | |
| 1069 | "zlfm (NOT (Ge a)) = zlfm (Lt a)" | |
| 1070 | "zlfm (NOT (Eq a)) = zlfm (NEq a)" | |
| 1071 | "zlfm (NOT (NEq a)) = zlfm (Eq a)" | |
| 1072 | "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)" | |
| 1073 | "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)" | |
| 1074 | "zlfm (NOT (Closed P)) = NClosed P" | |
| 1075 | "zlfm (NOT (NClosed P)) = Closed P" | |
| 1076 | "zlfm p = p" (hints simp add: fmsize_pos) | |
| 1077 | ||
| 1078 | lemma zlfm_I: | |
| 1079 | assumes qfp: "qfree p" | |
| 55981 | 1080 | shows "Ifm bbs (i # bs) (zlfm p) = Ifm bbs (i # bs) p \<and> iszlfm (zlfm p)" | 
| 60708 | 1081 | (is "?I (?l p) = ?I p \<and> ?L (?l p)") | 
| 50313 | 1082 | using qfp | 
| 1083 | proof (induct p rule: zlfm.induct) | |
| 1084 | case (5 a) | |
| 23274 | 1085 | let ?c = "fst (zsplit0 a)" | 
| 1086 | let ?r = "snd (zsplit0 a)" | |
| 55844 | 1087 | have spl: "zsplit0 a = (?c, ?r)" | 
| 1088 | by simp | |
| 50313 | 1089 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | 
| 55964 | 1090 | have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" | 
| 55844 | 1091 | by auto | 
| 55964 | 1092 | let ?N = "\<lambda>t. Inum (i # bs) t" | 
| 1093 | from 5 Ia nb show ?case | |
| 50313 | 1094 | apply (auto simp add: Let_def split_def algebra_simps) | 
| 55844 | 1095 | apply (cases "?r") | 
| 1096 | apply auto | |
| 60708 | 1097 | subgoal for nat a b by (cases nat) auto | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1098 | done | 
| 23274 | 1099 | next | 
| 50313 | 1100 | case (6 a) | 
| 23274 | 1101 | let ?c = "fst (zsplit0 a)" | 
| 1102 | let ?r = "snd (zsplit0 a)" | |
| 55844 | 1103 | have spl: "zsplit0 a = (?c, ?r)" | 
| 1104 | by simp | |
| 50313 | 1105 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | 
| 55964 | 1106 | have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" | 
| 55844 | 1107 | by auto | 
| 55964 | 1108 | let ?N = "\<lambda>t. Inum (i # bs) t" | 
| 50313 | 1109 | from 6 Ia nb show ?case | 
| 1110 | apply (auto simp add: Let_def split_def algebra_simps) | |
| 55844 | 1111 | apply (cases "?r") | 
| 1112 | apply auto | |
| 60708 | 1113 | subgoal for nat a b by (cases nat) auto | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1114 | done | 
| 23274 | 1115 | next | 
| 50313 | 1116 | case (7 a) | 
| 23274 | 1117 | let ?c = "fst (zsplit0 a)" | 
| 1118 | let ?r = "snd (zsplit0 a)" | |
| 55844 | 1119 | have spl: "zsplit0 a = (?c, ?r)" | 
| 1120 | by simp | |
| 50313 | 1121 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | 
| 55844 | 1122 | have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" | 
| 1123 | by auto | |
| 55964 | 1124 | let ?N = "\<lambda>t. Inum (i # bs) t" | 
| 50313 | 1125 | from 7 Ia nb show ?case | 
| 1126 | apply (auto simp add: Let_def split_def algebra_simps) | |
| 55844 | 1127 | apply (cases "?r") | 
| 1128 | apply auto | |
| 60708 | 1129 | subgoal for nat a b by (cases nat) auto | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1130 | done | 
| 23274 | 1131 | next | 
| 50313 | 1132 | case (8 a) | 
| 23274 | 1133 | let ?c = "fst (zsplit0 a)" | 
| 1134 | let ?r = "snd (zsplit0 a)" | |
| 55844 | 1135 | have spl: "zsplit0 a = (?c, ?r)" | 
| 1136 | by simp | |
| 50313 | 1137 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | 
| 55964 | 1138 | have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" | 
| 55844 | 1139 | by auto | 
| 55964 | 1140 | let ?N = "\<lambda>t. Inum (i # bs) t" | 
| 55844 | 1141 | from 8 Ia nb show ?case | 
| 50313 | 1142 | apply (auto simp add: Let_def split_def algebra_simps) | 
| 55844 | 1143 | apply (cases "?r") | 
| 1144 | apply auto | |
| 60708 | 1145 | subgoal for nat a b by (cases nat) auto | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1146 | done | 
| 23274 | 1147 | next | 
| 50313 | 1148 | case (9 a) | 
| 23274 | 1149 | let ?c = "fst (zsplit0 a)" | 
| 1150 | let ?r = "snd (zsplit0 a)" | |
| 55844 | 1151 | have spl: "zsplit0 a = (?c, ?r)" | 
| 1152 | by simp | |
| 50313 | 1153 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | 
| 55844 | 1154 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" | 
| 1155 | by auto | |
| 55964 | 1156 | let ?N = "\<lambda>t. Inum (i # bs) t" | 
| 55844 | 1157 | from 9 Ia nb show ?case | 
| 50313 | 1158 | apply (auto simp add: Let_def split_def algebra_simps) | 
| 55844 | 1159 | apply (cases "?r") | 
| 1160 | apply auto | |
| 60708 | 1161 | subgoal for nat a b by (cases nat) auto | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1162 | done | 
| 23274 | 1163 | next | 
| 50313 | 1164 | case (10 a) | 
| 23274 | 1165 | let ?c = "fst (zsplit0 a)" | 
| 1166 | let ?r = "snd (zsplit0 a)" | |
| 55844 | 1167 | have spl: "zsplit0 a = (?c, ?r)" | 
| 1168 | by simp | |
| 50313 | 1169 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | 
| 55844 | 1170 | have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" | 
| 1171 | by auto | |
| 55964 | 1172 | let ?N = "\<lambda>t. Inum (i # bs) t" | 
| 55844 | 1173 | from 10 Ia nb show ?case | 
| 50313 | 1174 | apply (auto simp add: Let_def split_def algebra_simps) | 
| 55844 | 1175 | apply (cases "?r") | 
| 1176 | apply auto | |
| 60708 | 1177 | subgoal for nat a b by (cases nat) auto | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1178 | 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 | 1179 | next | 
| 50313 | 1180 | case (11 j a) | 
| 23274 | 1181 | let ?c = "fst (zsplit0 a)" | 
| 1182 | let ?r = "snd (zsplit0 a)" | |
| 55844 | 1183 | have spl: "zsplit0 a = (?c,?r)" | 
| 1184 | by simp | |
| 50313 | 1185 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | 
| 55844 | 1186 | have Ia: "Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" | 
| 1187 | by auto | |
| 50313 | 1188 | let ?N = "\<lambda>t. Inum (i#bs) t" | 
| 60708 | 1189 | consider "j = 0" | "j \<noteq> 0" "?c = 0" | "j \<noteq> 0" "?c > 0" | "j \<noteq> 0" "?c < 0" | 
| 55844 | 1190 | by arith | 
| 60708 | 1191 | then show ?case | 
| 1192 | proof cases | |
| 1193 | case 1 | |
| 55844 | 1194 | then have z: "zlfm (Dvd j a) = (zlfm (Eq a))" | 
| 1195 | by (simp add: Let_def) | |
| 60708 | 1196 | with 11 \<open>j = 0\<close> show ?thesis | 
| 55844 | 1197 | by (simp del: zlfm.simps) | 
| 60708 | 1198 | next | 
| 1199 | case 2 | |
| 1200 | with zsplit0_I[OF spl, where x="i" and bs="bs"] show ?thesis | |
| 1201 | apply (auto simp add: Let_def split_def algebra_simps) | |
| 1202 | apply (cases "?r") | |
| 1203 | apply auto | |
| 1204 | subgoal for nat a b by (cases nat) auto | |
| 1205 | done | |
| 1206 | next | |
| 1207 | case 3 | |
| 55844 | 1208 | then have l: "?L (?l (Dvd j a))" | 
| 23274 | 1209 | by (simp add: nb Let_def split_def) | 
| 60708 | 1210 | with Ia 3 show ?thesis | 
| 1211 | by (simp add: Let_def split_def) | |
| 1212 | next | |
| 1213 | case 4 | |
| 55844 | 1214 | then have l: "?L (?l (Dvd j a))" | 
| 23274 | 1215 | by (simp add: nb Let_def split_def) | 
| 60708 | 1216 | with Ia 4 dvd_minus_iff[of "abs j" "?c*i + ?N ?r"] show ?thesis | 
| 55844 | 1217 | by (simp add: Let_def split_def) | 
| 60708 | 1218 | 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 | 1219 | next | 
| 50313 | 1220 | case (12 j a) | 
| 23274 | 1221 | let ?c = "fst (zsplit0 a)" | 
| 1222 | let ?r = "snd (zsplit0 a)" | |
| 55844 | 1223 | have spl: "zsplit0 a = (?c, ?r)" | 
| 1224 | by simp | |
| 50313 | 1225 | from zsplit0_I[OF spl, where x="i" and bs="bs"] | 
| 55844 | 1226 | have Ia:"Inum (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" | 
| 1227 | by auto | |
| 55964 | 1228 | let ?N = "\<lambda>t. Inum (i # bs) t" | 
| 60708 | 1229 | consider "j = 0" | "j \<noteq> 0" "?c = 0" | "j \<noteq> 0" "?c > 0" | "j \<noteq> 0" "?c < 0" | 
| 55844 | 1230 | by arith | 
| 60708 | 1231 | then show ?case | 
| 1232 | proof cases | |
| 1233 | case 1 | |
| 55964 | 1234 | then have z: "zlfm (NDvd j a) = zlfm (NEq a)" | 
| 55844 | 1235 | by (simp add: Let_def) | 
| 60708 | 1236 | with assms 12 \<open>j = 0\<close> show ?thesis | 
| 1237 | by (simp del: zlfm.simps) | |
| 1238 | next | |
| 1239 | case 2 | |
| 1240 | with zsplit0_I[OF spl, where x="i" and bs="bs"] show ?thesis | |
| 1241 | apply (auto simp add: Let_def split_def algebra_simps) | |
| 1242 | apply (cases "?r") | |
| 1243 | apply auto | |
| 1244 | subgoal for nat a b by (cases nat) auto | |
| 1245 | done | |
| 1246 | next | |
| 1247 | case 3 | |
| 55844 | 1248 | then have l: "?L (?l (Dvd j a))" | 
| 23274 | 1249 | by (simp add: nb Let_def split_def) | 
| 60708 | 1250 | with Ia 3 show ?thesis | 
| 55844 | 1251 | by (simp add: Let_def split_def) | 
| 60708 | 1252 | next | 
| 1253 | case 4 | |
| 55844 | 1254 | then have l: "?L (?l (Dvd j a))" | 
| 23274 | 1255 | by (simp add: nb Let_def split_def) | 
| 60708 | 1256 | with Ia 4 dvd_minus_iff[of "abs j" "?c*i + ?N ?r"] show ?thesis | 
| 55844 | 1257 | by (simp add: Let_def split_def) | 
| 60708 | 1258 | qed | 
| 23274 | 1259 | qed auto | 
| 1260 | ||
| 60533 | 1261 | consts minusinf :: "fm \<Rightarrow> fm" -- \<open>Virtual substitution of @{text "-\<infinity>"}\<close>
 | 
| 23274 | 1262 | recdef minusinf "measure size" | 
| 50313 | 1263 | "minusinf (And p q) = And (minusinf p) (minusinf q)" | 
| 1264 | "minusinf (Or p q) = Or (minusinf p) (minusinf q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1265 | "minusinf (Eq (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1266 | "minusinf (NEq (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1267 | "minusinf (Lt (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1268 | "minusinf (Le (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1269 | "minusinf (Gt (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1270 | "minusinf (Ge (CN 0 c e)) = F" | 
| 23274 | 1271 | "minusinf p = p" | 
| 1272 | ||
| 1273 | lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)" | |
| 50313 | 1274 | by (induct p rule: minusinf.induct) auto | 
| 23274 | 1275 | |
| 60533 | 1276 | consts plusinf :: "fm \<Rightarrow> fm"  -- \<open>Virtual substitution of @{text "+\<infinity>"}\<close>
 | 
| 23274 | 1277 | recdef plusinf "measure size" | 
| 50313 | 1278 | "plusinf (And p q) = And (plusinf p) (plusinf q)" | 
| 1279 | "plusinf (Or p q) = Or (plusinf p) (plusinf q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1280 | "plusinf (Eq (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1281 | "plusinf (NEq (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1282 | "plusinf (Lt (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1283 | "plusinf (Le (CN 0 c e)) = F" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1284 | "plusinf (Gt (CN 0 c e)) = T" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1285 | "plusinf (Ge (CN 0 c e)) = T" | 
| 23274 | 1286 | "plusinf p = p" | 
| 1287 | ||
| 60533 | 1288 | consts \<delta> :: "fm \<Rightarrow> int"  -- \<open>Compute @{text "lcm {d| N\<^sup>? Dvd c*x+t \<in> p}"}\<close>
 | 
| 23274 | 1289 | recdef \<delta> "measure size" | 
| 50313 | 1290 | "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)" | 
| 1291 | "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1292 | "\<delta> (Dvd i (CN 0 c e)) = i" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1293 | "\<delta> (NDvd i (CN 0 c e)) = i" | 
| 23274 | 1294 | "\<delta> p = 1" | 
| 1295 | ||
| 60533 | 1296 | consts d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool" -- \<open>check if a given l divides all the ds above\<close> | 
| 50252 | 1297 | recdef d_\<delta> "measure size" | 
| 50313 | 1298 | "d_\<delta> (And p q) = (\<lambda>d. d_\<delta> p d \<and> d_\<delta> q d)" | 
| 1299 | "d_\<delta> (Or p q) = (\<lambda>d. d_\<delta> p d \<and> d_\<delta> q d)" | |
| 1300 | "d_\<delta> (Dvd i (CN 0 c e)) = (\<lambda>d. i dvd d)" | |
| 1301 | "d_\<delta> (NDvd i (CN 0 c e)) = (\<lambda>d. i dvd d)" | |
| 1302 | "d_\<delta> p = (\<lambda>d. True)" | |
| 23274 | 1303 | |
| 50313 | 1304 | lemma delta_mono: | 
| 23274 | 1305 | assumes lin: "iszlfm p" | 
| 50313 | 1306 | and d: "d dvd d'" | 
| 1307 | and ad: "d_\<delta> p d" | |
| 50252 | 1308 | shows "d_\<delta> p d'" | 
| 55999 | 1309 | using lin ad | 
| 50313 | 1310 | proof (induct p rule: iszlfm.induct) | 
| 55844 | 1311 | case (9 i c e) | 
| 1312 | then show ?case using d | |
| 30042 | 1313 | by (simp add: dvd_trans[of "i" "d" "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 | 1314 | next | 
| 55844 | 1315 | case (10 i c e) | 
| 1316 | then show ?case using d | |
| 30042 | 1317 | by (simp add: dvd_trans[of "i" "d" "d'"]) | 
| 23274 | 1318 | 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 | 1319 | |
| 50313 | 1320 | lemma \<delta>: | 
| 55885 | 1321 | assumes lin: "iszlfm p" | 
| 50252 | 1322 | shows "d_\<delta> p (\<delta> p) \<and> \<delta> p >0" | 
| 50313 | 1323 | using lin | 
| 23274 | 1324 | proof (induct p rule: iszlfm.induct) | 
| 50313 | 1325 | case (1 p q) | 
| 23274 | 1326 | let ?d = "\<delta> (And p q)" | 
| 55844 | 1327 | from 1 lcm_pos_int have dp: "?d > 0" | 
| 1328 | by simp | |
| 1329 | have d1: "\<delta> p dvd \<delta> (And p q)" | |
| 1330 | using 1 by simp | |
| 1331 | then have th: "d_\<delta> p ?d" | |
| 1332 | using delta_mono 1(2,3) by (simp only: iszlfm.simps) | |
| 1333 | have "\<delta> q dvd \<delta> (And p q)" | |
| 1334 | using 1 by simp | |
| 1335 | then have th': "d_\<delta> q ?d" | |
| 1336 | using delta_mono 1 by (simp only: iszlfm.simps) | |
| 1337 | from th th' dp show ?case | |
| 1338 | by simp | |
| 23274 | 1339 | next | 
| 50313 | 1340 | case (2 p q) | 
| 23274 | 1341 | let ?d = "\<delta> (And p q)" | 
| 55844 | 1342 | from 2 lcm_pos_int have dp: "?d > 0" | 
| 1343 | by simp | |
| 1344 | have "\<delta> p dvd \<delta> (And p q)" | |
| 1345 | using 2 by simp | |
| 1346 | then have th: "d_\<delta> p ?d" | |
| 1347 | using delta_mono 2 by (simp only: iszlfm.simps) | |
| 1348 | have "\<delta> q dvd \<delta> (And p q)" | |
| 1349 | using 2 by simp | |
| 1350 | then have th': "d_\<delta> q ?d" | |
| 1351 | using delta_mono 2 by (simp only: iszlfm.simps) | |
| 1352 | from th th' dp show ?case | |
| 1353 | 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 | 1354 | 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 | 1355 | |
| 
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 | 1356 | |
| 60533 | 1357 | consts a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" -- \<open>adjust the coefficients of a formula\<close> | 
| 50252 | 1358 | recdef a_\<beta> "measure size" | 
| 50313 | 1359 | "a_\<beta> (And p q) = (\<lambda>k. And (a_\<beta> p k) (a_\<beta> q k))" | 
| 1360 | "a_\<beta> (Or p q) = (\<lambda>k. Or (a_\<beta> p k) (a_\<beta> q k))" | |
| 1361 | "a_\<beta> (Eq (CN 0 c e)) = (\<lambda>k. Eq (CN 0 1 (Mul (k div c) e)))" | |
| 1362 | "a_\<beta> (NEq (CN 0 c e)) = (\<lambda>k. NEq (CN 0 1 (Mul (k div c) e)))" | |
| 1363 | "a_\<beta> (Lt (CN 0 c e)) = (\<lambda>k. Lt (CN 0 1 (Mul (k div c) e)))" | |
| 1364 | "a_\<beta> (Le (CN 0 c e)) = (\<lambda>k. Le (CN 0 1 (Mul (k div c) e)))" | |
| 1365 | "a_\<beta> (Gt (CN 0 c e)) = (\<lambda>k. Gt (CN 0 1 (Mul (k div c) e)))" | |
| 1366 | "a_\<beta> (Ge (CN 0 c e)) = (\<lambda>k. Ge (CN 0 1 (Mul (k div c) e)))" | |
| 1367 | "a_\<beta> (Dvd i (CN 0 c e)) =(\<lambda>k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" | |
| 1368 | "a_\<beta> (NDvd i (CN 0 c e))=(\<lambda>k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" | |
| 1369 | "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 | 1370 | |
| 60533 | 1371 | consts d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" -- \<open>test if all coeffs c of c divide a given l\<close> | 
| 50252 | 1372 | recdef d_\<beta> "measure size" | 
| 50313 | 1373 | "d_\<beta> (And p q) = (\<lambda>k. (d_\<beta> p k) \<and> (d_\<beta> q k))" | 
| 1374 | "d_\<beta> (Or p q) = (\<lambda>k. (d_\<beta> p k) \<and> (d_\<beta> q k))" | |
| 1375 | "d_\<beta> (Eq (CN 0 c e)) = (\<lambda>k. c dvd k)" | |
| 1376 | "d_\<beta> (NEq (CN 0 c e)) = (\<lambda>k. c dvd k)" | |
| 1377 | "d_\<beta> (Lt (CN 0 c e)) = (\<lambda>k. c dvd k)" | |
| 1378 | "d_\<beta> (Le (CN 0 c e)) = (\<lambda>k. c dvd k)" | |
| 1379 | "d_\<beta> (Gt (CN 0 c e)) = (\<lambda>k. c dvd k)" | |
| 1380 | "d_\<beta> (Ge (CN 0 c e)) = (\<lambda>k. c dvd k)" | |
| 1381 | "d_\<beta> (Dvd i (CN 0 c e)) =(\<lambda>k. c dvd k)" | |
| 1382 | "d_\<beta> (NDvd i (CN 0 c e))=(\<lambda>k. c dvd k)" | |
| 1383 | "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 | 1384 | |
| 60533 | 1385 | consts \<zeta> :: "fm \<Rightarrow> int" -- \<open>computes the lcm of all coefficients of x\<close> | 
| 23274 | 1386 | recdef \<zeta> "measure size" | 
| 50313 | 1387 | "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)" | 
| 1388 | "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1389 | "\<zeta> (Eq (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1390 | "\<zeta> (NEq (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1391 | "\<zeta> (Lt (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1392 | "\<zeta> (Le (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1393 | "\<zeta> (Gt (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1394 | "\<zeta> (Ge (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1395 | "\<zeta> (Dvd i (CN 0 c e)) = c" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1396 | "\<zeta> (NDvd i (CN 0 c e))= c" | 
| 23274 | 1397 | "\<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 | 1398 | |
| 50313 | 1399 | consts \<beta> :: "fm \<Rightarrow> num list" | 
| 23274 | 1400 | recdef \<beta> "measure size" | 
| 50313 | 1401 | "\<beta> (And p q) = (\<beta> p @ \<beta> q)" | 
| 1402 | "\<beta> (Or p q) = (\<beta> p @ \<beta> q)" | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58310diff
changeset | 1403 | "\<beta> (Eq (CN 0 c e)) = [Sub (C (- 1)) e]" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1404 | "\<beta> (NEq (CN 0 c e)) = [Neg e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1405 | "\<beta> (Lt (CN 0 c e)) = []" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1406 | "\<beta> (Le (CN 0 c e)) = []" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1407 | "\<beta> (Gt (CN 0 c e)) = [Neg e]" | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58310diff
changeset | 1408 | "\<beta> (Ge (CN 0 c e)) = [Sub (C (- 1)) e]" | 
| 23274 | 1409 | "\<beta> p = []" | 
| 19736 | 1410 | |
| 50313 | 1411 | consts \<alpha> :: "fm \<Rightarrow> num list" | 
| 23274 | 1412 | recdef \<alpha> "measure size" | 
| 55999 | 1413 | "\<alpha> (And p q) = \<alpha> p @ \<alpha> q" | 
| 1414 | "\<alpha> (Or p q) = \<alpha> p @ \<alpha> q" | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58310diff
changeset | 1415 | "\<alpha> (Eq (CN 0 c e)) = [Add (C (- 1)) e]" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1416 | "\<alpha> (NEq (CN 0 c e)) = [e]" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1417 | "\<alpha> (Lt (CN 0 c e)) = [e]" | 
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58310diff
changeset | 1418 | "\<alpha> (Le (CN 0 c e)) = [Add (C (- 1)) e]" | 
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1419 | "\<alpha> (Gt (CN 0 c e)) = []" | 
| 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1420 | "\<alpha> (Ge (CN 0 c e)) = []" | 
| 23274 | 1421 | "\<alpha> p = []" | 
| 50313 | 1422 | |
| 23274 | 1423 | consts mirror :: "fm \<Rightarrow> fm" | 
| 1424 | recdef mirror "measure size" | |
| 50313 | 1425 | "mirror (And p q) = And (mirror p) (mirror q)" | 
| 1426 | "mirror (Or p q) = Or (mirror p) (mirror q)" | |
| 23995 
c34490f1e0ff
Updated proofs; changed shadow syntax to improve (processing) time
 chaieb parents: 
23984diff
changeset | 1427 | "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 | 1428 | "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 | 1429 | "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 | 1430 | "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 | 1431 | "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 | 1432 | "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 | 1433 | "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 | 1434 | "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))" | 
| 23274 | 1435 | "mirror p = p" | 
| 50313 | 1436 | |
| 60533 | 1437 | text \<open>Lemmas for the correctness of @{text "\<sigma>_\<rho>"}\<close>
 | 
| 50313 | 1438 | |
| 55844 | 1439 | lemma dvd1_eq1: | 
| 1440 | fixes x :: int | |
| 1441 | shows "x > 0 \<Longrightarrow> x dvd 1 \<longleftrightarrow> x = 1" | |
| 41807 | 1442 | 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 | 1443 | |
| 23274 | 1444 | lemma minusinf_inf: | 
| 1445 | assumes linp: "iszlfm p" | |
| 50313 | 1446 | and u: "d_\<beta> p 1" | 
| 55964 | 1447 | shows "\<exists>z::int. \<forall>x < z. Ifm bbs (x # bs) (minusinf p) = Ifm bbs (x # bs) p" | 
| 50313 | 1448 | (is "?P p" is "\<exists>(z::int). \<forall>x < z. ?I x (?M p) = ?I x p") | 
| 1449 | using linp u | |
| 23274 | 1450 | proof (induct p rule: minusinf.induct) | 
| 55844 | 1451 | case (1 p q) | 
| 1452 | then show ?case | |
| 60708 | 1453 | apply auto | 
| 1454 | subgoal for z z' by (rule exI [where x = "min z z'"]) simp | |
| 1455 | done | |
| 23274 | 1456 | next | 
| 55844 | 1457 | case (2 p q) | 
| 1458 | then show ?case | |
| 60708 | 1459 | apply auto | 
| 1460 | subgoal for z z' by (rule exI [where x = "min z z'"]) simp | |
| 1461 | 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 | 1462 | next | 
| 55844 | 1463 | case (3 c e) | 
| 1464 | then have c1: "c = 1" and nb: "numbound0 e" | |
| 1465 | by simp_all | |
| 26934 | 1466 | fix a | 
| 55999 | 1467 | from 3 have "\<forall>x<(- Inum (a # bs) e). c * x + Inum (x # bs) e \<noteq> 0" | 
| 55844 | 1468 | proof clarsimp | 
| 1469 | fix x | |
| 55999 | 1470 | assume "x < (- Inum (a # bs) e)" and "x + Inum (x # bs) e = 0" | 
| 23274 | 1471 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | 
| 55844 | 1472 | show False by simp | 
| 23274 | 1473 | qed | 
| 55844 | 1474 | then show ?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 | 1475 | next | 
| 55844 | 1476 | case (4 c e) | 
| 1477 | then have c1: "c = 1" and nb: "numbound0 e" | |
| 1478 | by simp_all | |
| 26934 | 1479 | fix a | 
| 55964 | 1480 | from 4 have "\<forall>x < (- Inum (a # bs) e). c * x + Inum (x # bs) e \<noteq> 0" | 
| 55921 | 1481 | proof clarsimp | 
| 1482 | fix x | |
| 55964 | 1483 | assume "x < (- Inum (a # bs) e)" and "x + Inum (x # bs) e = 0" | 
| 23274 | 1484 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | 
| 1485 | show "False" by simp | |
| 1486 | qed | |
| 55885 | 1487 | then show ?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 | 1488 | next | 
| 55921 | 1489 | case (5 c e) | 
| 1490 | then have c1: "c = 1" and nb: "numbound0 e" | |
| 1491 | by simp_all | |
| 26934 | 1492 | fix a | 
| 55999 | 1493 | from 5 have "\<forall>x<(- Inum (a # bs) e). c * x + Inum (x # bs) e < 0" | 
| 55921 | 1494 | proof clarsimp | 
| 1495 | fix x | |
| 55964 | 1496 | assume "x < (- Inum (a # bs) e)" | 
| 23274 | 1497 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | 
| 55921 | 1498 | show "x + Inum (x # bs) e < 0" | 
| 1499 | by simp | |
| 23274 | 1500 | qed | 
| 55885 | 1501 | then show ?case by auto | 
| 23274 | 1502 | next | 
| 55921 | 1503 | case (6 c e) | 
| 1504 | then have c1: "c = 1" and nb: "numbound0 e" | |
| 1505 | by simp_all | |
| 26934 | 1506 | fix a | 
| 55964 | 1507 | from 6 have "\<forall>x<(- Inum (a # bs) e). c * x + Inum (x # bs) e \<le> 0" | 
| 55921 | 1508 | proof clarsimp | 
| 1509 | fix x | |
| 55964 | 1510 | assume "x < (- Inum (a # bs) e)" | 
| 23274 | 1511 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | 
| 55964 | 1512 | show "x + Inum (x # bs) e \<le> 0" by simp | 
| 23274 | 1513 | qed | 
| 55885 | 1514 | then show ?case by auto | 
| 23274 | 1515 | next | 
| 55921 | 1516 | case (7 c e) | 
| 1517 | then have c1: "c = 1" and nb: "numbound0 e" | |
| 1518 | by simp_all | |
| 26934 | 1519 | fix a | 
| 55964 | 1520 | from 7 have "\<forall>x<(- Inum (a # bs) e). \<not> (c * x + Inum (x # bs) e > 0)" | 
| 55921 | 1521 | proof clarsimp | 
| 1522 | fix x | |
| 55964 | 1523 | assume "x < - Inum (a # bs) e" and "x + Inum (x # bs) e > 0" | 
| 23274 | 1524 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | 
| 55921 | 1525 | show False by simp | 
| 23274 | 1526 | qed | 
| 55885 | 1527 | then show ?case by auto | 
| 23274 | 1528 | next | 
| 55921 | 1529 | case (8 c e) | 
| 1530 | then have c1: "c = 1" and nb: "numbound0 e" | |
| 1531 | by simp_all | |
| 26934 | 1532 | fix a | 
| 55999 | 1533 | from 8 have "\<forall>x<(- Inum (a # bs) e). \<not> c * x + Inum (x # bs) e \<ge> 0" | 
| 55921 | 1534 | proof clarsimp | 
| 1535 | fix x | |
| 55999 | 1536 | assume "x < (- Inum (a # bs) e)" and "x + Inum (x # bs) e \<ge> 0" | 
| 23274 | 1537 | with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] | 
| 55921 | 1538 | show False by simp | 
| 23274 | 1539 | qed | 
| 55885 | 1540 | then show ?case by auto | 
| 23274 | 1541 | 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 | 1542 | |
| 23274 | 1543 | lemma minusinf_repeats: | 
| 55921 | 1544 | assumes d: "d_\<delta> p d" | 
| 1545 | and linp: "iszlfm p" | |
| 1546 | shows "Ifm bbs ((x - k * d) # bs) (minusinf p) = Ifm bbs (x # bs) (minusinf p)" | |
| 50313 | 1547 | using linp d | 
| 1548 | proof (induct p rule: iszlfm.induct) | |
| 1549 | case (9 i c e) | |
| 55921 | 1550 | then have nbe: "numbound0 e" and id: "i dvd d" | 
| 1551 | by simp_all | |
| 1552 | then have "\<exists>k. d = i * k" | |
| 1553 | by (simp add: dvd_def) | |
| 1554 | then obtain "di" where di_def: "d = i * di" | |
| 1555 | by blast | |
| 50313 | 1556 | show ?case | 
| 1557 | proof (simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib, | |
| 1558 | rule iffI) | |
| 55921 | 1559 | assume "i dvd c * x - c * (k * d) + Inum (x # bs) e" | 
| 55999 | 1560 | (is "?ri dvd ?rc * ?rx - ?rc * (?rk * ?rd) + ?I x e" is "?ri dvd ?rt") | 
| 55921 | 1561 | then have "\<exists>l::int. ?rt = i * l" | 
| 1562 | by (simp add: dvd_def) | |
| 55964 | 1563 | then have "\<exists>l::int. c * x + ?I x e = i * l + c * (k * i * di)" | 
| 50313 | 1564 | by (simp add: algebra_simps di_def) | 
| 55964 | 1565 | then have "\<exists>l::int. c * x + ?I x e = i* (l + c * k * di)" | 
| 50313 | 1566 | by (simp add: algebra_simps) | 
| 55921 | 1567 | then have "\<exists>l::int. c * x + ?I x e = i * l" | 
| 1568 | by blast | |
| 1569 | then show "i dvd c * x + Inum (x # bs) e" | |
| 1570 | by (simp add: dvd_def) | |
| 50313 | 1571 | next | 
| 55964 | 1572 | assume "i dvd c * x + Inum (x # bs) e" (is "?ri dvd ?rc * ?rx + ?e") | 
| 55921 | 1573 | then have "\<exists>l::int. c * x + ?e = i * l" | 
| 1574 | by (simp add: dvd_def) | |
| 1575 | then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l - c * (k * d)" | |
| 1576 | by simp | |
| 1577 | then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l - c * (k * i * di)" | |
| 1578 | by (simp add: di_def) | |
| 55964 | 1579 | then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * (l - c * k * di)" | 
| 55921 | 1580 | by (simp add: algebra_simps) | 
| 1581 | then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l" | |
| 1582 | by blast | |
| 1583 | then show "i dvd c * x - c * (k * d) + Inum (x # bs) e" | |
| 1584 | by (simp add: dvd_def) | |
| 50313 | 1585 | qed | 
| 23274 | 1586 | next | 
| 50313 | 1587 | case (10 i c e) | 
| 55921 | 1588 | then have nbe: "numbound0 e" and id: "i dvd d" | 
| 1589 | by simp_all | |
| 1590 | then have "\<exists>k. d = i * k" | |
| 1591 | by (simp add: dvd_def) | |
| 1592 | then obtain di where di_def: "d = i * di" | |
| 1593 | by blast | |
| 50313 | 1594 | show ?case | 
| 55999 | 1595 | proof (simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib, | 
| 1596 | rule iffI) | |
| 55921 | 1597 | assume "i dvd c * x - c * (k * d) + Inum (x # bs) e" | 
| 55999 | 1598 | (is "?ri dvd ?rc * ?rx - ?rc * (?rk * ?rd) + ?I x e" is "?ri dvd ?rt") | 
| 55921 | 1599 | then have "\<exists>l::int. ?rt = i * l" | 
| 1600 | by (simp add: dvd_def) | |
| 1601 | then have "\<exists>l::int. c * x + ?I x e = i * l + c * (k * i * di)" | |
| 50313 | 1602 | by (simp add: algebra_simps di_def) | 
| 55921 | 1603 | then have "\<exists>l::int. c * x+ ?I x e = i * (l + c * k * di)" | 
| 50313 | 1604 | by (simp add: algebra_simps) | 
| 55921 | 1605 | then have "\<exists>l::int. c * x + ?I x e = i * l" | 
| 1606 | by blast | |
| 1607 | then show "i dvd c * x + Inum (x # bs) e" | |
| 1608 | by (simp add: dvd_def) | |
| 50313 | 1609 | next | 
| 55921 | 1610 | assume "i dvd c * x + Inum (x # bs) e" (is "?ri dvd ?rc * ?rx + ?e") | 
| 1611 | then have "\<exists>l::int. c * x + ?e = i * l" | |
| 1612 | by (simp add: dvd_def) | |
| 1613 | then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l - c * (k * d)" | |
| 1614 | by simp | |
| 1615 | then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l - c * (k * i * di)" | |
| 1616 | by (simp add: di_def) | |
| 55999 | 1617 | then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * (l - c * k * di)" | 
| 55921 | 1618 | by (simp add: algebra_simps) | 
| 1619 | then have "\<exists>l::int. c * x - c * (k * d) + ?e = i * l" | |
| 50313 | 1620 | by blast | 
| 55921 | 1621 | then show "i dvd c * x - c * (k * d) + Inum (x # bs) e" | 
| 1622 | by (simp add: dvd_def) | |
| 50313 | 1623 | qed | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1624 | 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 | 1625 | |
| 50252 | 1626 | lemma mirror_\<alpha>_\<beta>: | 
| 23274 | 1627 | assumes lp: "iszlfm p" | 
| 55964 | 1628 | shows "Inum (i # bs) ` set (\<alpha> p) = Inum (i # bs) ` set (\<beta> (mirror p))" | 
| 50313 | 1629 | using lp 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 | 1630 | |
| 50313 | 1631 | lemma mirror: | 
| 23274 | 1632 | assumes lp: "iszlfm p" | 
| 55921 | 1633 | shows "Ifm bbs (x # bs) (mirror p) = Ifm bbs ((- x) # bs) p" | 
| 50313 | 1634 | using lp | 
| 1635 | proof (induct p rule: iszlfm.induct) | |
| 1636 | case (9 j c e) | |
| 55964 | 1637 | then have nb: "numbound0 e" | 
| 1638 | by simp | |
| 1639 | have "Ifm bbs (x # bs) (mirror (Dvd j (CN 0 c e))) \<longleftrightarrow> j dvd c * x - Inum (x # bs) e" | |
| 50313 | 1640 | (is "_ = (j dvd c*x - ?e)") by simp | 
| 55964 | 1641 | also have "\<dots> \<longleftrightarrow> j dvd (- (c * x - ?e))" | 
| 30042 | 1642 | by (simp only: dvd_minus_iff) | 
| 55964 | 1643 | also have "\<dots> \<longleftrightarrow> j dvd (c * (- x)) + ?e" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
55999diff
changeset | 1644 | by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] ac_simps minus_add_distrib) | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53168diff
changeset | 1645 | (simp add: algebra_simps) | 
| 55964 | 1646 | also have "\<dots> = Ifm bbs ((- x) # bs) (Dvd j (CN 0 c e))" | 
| 50313 | 1647 | using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp | 
| 23274 | 1648 | finally show ?case . | 
| 1649 | next | |
| 55964 | 1650 | case (10 j c e) | 
| 1651 | then have nb: "numbound0 e" | |
| 1652 | by simp | |
| 1653 | have "Ifm bbs (x # bs) (mirror (Dvd j (CN 0 c e))) \<longleftrightarrow> j dvd c * x - Inum (x # bs) e" | |
| 1654 | (is "_ = (j dvd c * x - ?e)") by simp | |
| 1655 | also have "\<dots> \<longleftrightarrow> j dvd (- (c * x - ?e))" | |
| 30042 | 1656 | by (simp only: dvd_minus_iff) | 
| 55964 | 1657 | also have "\<dots> \<longleftrightarrow> j dvd (c * (- x)) + ?e" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
55999diff
changeset | 1658 | by (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] ac_simps minus_add_distrib) | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53168diff
changeset | 1659 | (simp add: algebra_simps) | 
| 55964 | 1660 | also have "\<dots> \<longleftrightarrow> Ifm bbs ((- x) # bs) (Dvd j (CN 0 c e))" | 
| 50313 | 1661 | using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp | 
| 23274 | 1662 | finally show ?case by simp | 
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1663 | 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 | 1664 | |
| 50313 | 1665 | lemma mirror_l: "iszlfm p \<and> d_\<beta> p 1 \<Longrightarrow> iszlfm (mirror p) \<and> d_\<beta> (mirror p) 1" | 
| 41807 | 1666 | 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 | 1667 | |
| 23274 | 1668 | lemma mirror_\<delta>: "iszlfm p \<Longrightarrow> \<delta> (mirror p) = \<delta> p" | 
| 41807 | 1669 | by (induct p rule: mirror.induct) auto | 
| 23274 | 1670 | |
| 50313 | 1671 | lemma \<beta>_numbound0: | 
| 1672 | assumes lp: "iszlfm p" | |
| 55964 | 1673 | shows "\<forall>b \<in> set (\<beta> p). numbound0 b" | 
| 41807 | 1674 | 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 | 1675 | |
| 50313 | 1676 | lemma d_\<beta>_mono: | 
| 23274 | 1677 | assumes linp: "iszlfm p" | 
| 50313 | 1678 | and dr: "d_\<beta> p l" | 
| 1679 | and d: "l dvd l'" | |
| 50252 | 1680 | shows "d_\<beta> p l'" | 
| 50313 | 1681 | using dr linp dvd_trans[of _ "l" "l'", simplified d] | 
| 41807 | 1682 | by (induct p rule: iszlfm.induct) simp_all | 
| 23274 | 1683 | |
| 50313 | 1684 | lemma \<alpha>_l: | 
| 55999 | 1685 | assumes "iszlfm p" | 
| 50313 | 1686 | shows "\<forall>b \<in> set (\<alpha> p). numbound0 b" | 
| 55999 | 1687 | using assms 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 | 1688 | |
| 50313 | 1689 | lemma \<zeta>: | 
| 55999 | 1690 | assumes "iszlfm p" | 
| 50252 | 1691 | shows "\<zeta> p > 0 \<and> d_\<beta> p (\<zeta> p)" | 
| 55999 | 1692 | using assms | 
| 50313 | 1693 | proof (induct p rule: iszlfm.induct) | 
| 23274 | 1694 | case (1 p q) | 
| 55964 | 1695 | from 1 have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" | 
| 1696 | by simp | |
| 1697 | from 1 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" | |
| 1698 | by simp | |
| 50313 | 1699 | from 1 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"] | 
| 55964 | 1700 | d_\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"] | 
| 1701 | dl1 dl2 | |
| 1702 | show ?case | |
| 1703 | by (auto simp add: lcm_pos_int) | |
| 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 | 1704 | next | 
| 23274 | 1705 | case (2 p q) | 
| 55964 | 1706 | from 2 have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" | 
| 1707 | by simp | |
| 1708 | from 2 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" | |
| 1709 | by simp | |
| 50313 | 1710 | from 2 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"] | 
| 55964 | 1711 | d_\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"] | 
| 1712 | dl1 dl2 | |
| 1713 | show ?case | |
| 1714 | by (auto simp add: lcm_pos_int) | |
| 31952 
40501bb2d57c
renamed lemmas: nat_xyz/int_xyz -> xyz_nat/xyz_int
 nipkow parents: 
31730diff
changeset | 1715 | qed (auto simp add: lcm_pos_int) | 
| 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 | 1716 | |
| 50313 | 1717 | lemma a_\<beta>: | 
| 55921 | 1718 | assumes linp: "iszlfm p" | 
| 1719 | and d: "d_\<beta> p l" | |
| 1720 | and lp: "l > 0" | |
| 55964 | 1721 | 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" | 
| 50313 | 1722 | using linp d | 
| 23274 | 1723 | proof (induct p rule: iszlfm.induct) | 
| 50313 | 1724 | case (5 c e) | 
| 55964 | 1725 | then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l" | 
| 55921 | 1726 | by simp_all | 
| 1727 | from lp cp have clel: "c \<le> l" | |
| 1728 | by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1729 | from cp have cnz: "c \<noteq> 0" | |
| 1730 | by simp | |
| 1731 | have "c div c \<le> l div c" | |
| 50313 | 1732 | by (simp add: zdiv_mono1[OF clel cp]) | 
| 55999 | 1733 | then have ldcp: "0 < l div c" | 
| 50313 | 1734 | by (simp add: div_self[OF cnz]) | 
| 55921 | 1735 | have "c * (l div c) = c * (l div c) + l mod c" | 
| 1736 | using d' dvd_eq_mod_eq_0[of "c" "l"] by simp | |
| 1737 | then have cl: "c * (l div c) =l" | |
| 1738 | using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp | |
| 55964 | 1739 | then have "(l * x + (l div c) * Inum (x # bs) e < 0) \<longleftrightarrow> | 
| 50313 | 1740 | ((c * (l div c)) * x + (l div c) * Inum (x # bs) e < 0)" | 
| 1741 | by simp | |
| 55999 | 1742 | also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) < (l div c) * 0" | 
| 50313 | 1743 | by (simp add: algebra_simps) | 
| 55964 | 1744 | also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e < 0" | 
| 1745 | using mult_less_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp | |
| 1746 | by simp | |
| 50313 | 1747 | finally show ?case | 
| 55964 | 1748 | using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] be | 
| 1749 | 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 | 1750 | next | 
| 50313 | 1751 | case (6 c e) | 
| 55921 | 1752 | then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l" | 
| 1753 | by simp_all | |
| 1754 | from lp cp have clel: "c \<le> l" | |
| 1755 | by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1756 | from cp have cnz: "c \<noteq> 0" | |
| 1757 | by simp | |
| 55964 | 1758 | have "c div c \<le> l div c" | 
| 50313 | 1759 | by (simp add: zdiv_mono1[OF clel cp]) | 
| 1760 | then have ldcp:"0 < l div c" | |
| 1761 | by (simp add: div_self[OF cnz]) | |
| 55964 | 1762 | have "c * (l div c) = c * (l div c) + l mod c" | 
| 55921 | 1763 | using d' dvd_eq_mod_eq_0[of "c" "l"] by simp | 
| 1764 | then have cl: "c * (l div c) = l" | |
| 1765 | using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp | |
| 55964 | 1766 | then have "l * x + (l div c) * Inum (x # bs) e \<le> 0 \<longleftrightarrow> | 
| 1767 | (c * (l div c)) * x + (l div c) * Inum (x # bs) e \<le> 0" | |
| 50313 | 1768 | by simp | 
| 55964 | 1769 | also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) \<le> (l div c) * 0" | 
| 50313 | 1770 | by (simp add: algebra_simps) | 
| 55964 | 1771 | also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e \<le> 0" | 
| 23274 | 1772 | using mult_le_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp by simp | 
| 50313 | 1773 | finally show ?case | 
| 1774 | 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 | 1775 | next | 
| 50313 | 1776 | case (7 c e) | 
| 55921 | 1777 | then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l" | 
| 1778 | by simp_all | |
| 1779 | from lp cp have clel: "c \<le> l" | |
| 1780 | by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1781 | from cp have cnz: "c \<noteq> 0" | |
| 1782 | by simp | |
| 1783 | have "c div c \<le> l div c" | |
| 1784 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 55964 | 1785 | then have ldcp: "0 < l div c" | 
| 55921 | 1786 | by (simp add: div_self[OF cnz]) | 
| 55964 | 1787 | have "c * (l div c) = c * (l div c) + l mod c" | 
| 55921 | 1788 | using d' dvd_eq_mod_eq_0[of "c" "l"] by simp | 
| 55964 | 1789 | then have cl: "c * (l div c) = l" | 
| 55921 | 1790 | using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp | 
| 55964 | 1791 | then have "l * x + (l div c) * Inum (x # bs) e > 0 \<longleftrightarrow> | 
| 1792 | (c * (l div c)) * x + (l div c) * Inum (x # bs) e > 0" | |
| 55921 | 1793 | by simp | 
| 55964 | 1794 | also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) > (l div c) * 0" | 
| 55921 | 1795 | by (simp add: algebra_simps) | 
| 55964 | 1796 | also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e > 0" | 
| 55921 | 1797 | using zero_less_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp | 
| 1798 | by simp | |
| 1799 | finally show ?case | |
| 1800 | using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be | |
| 1801 | by simp | |
| 1802 | next | |
| 1803 | case (8 c e) | |
| 1804 | then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l" | |
| 1805 | by simp_all | |
| 1806 | from lp cp have clel: "c \<le> l" | |
| 1807 | by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1808 | from cp have cnz: "c \<noteq> 0" | |
| 1809 | by simp | |
| 1810 | have "c div c \<le> l div c" | |
| 1811 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1812 | then have ldcp: "0 < l div c" | |
| 1813 | by (simp add: div_self[OF cnz]) | |
| 55964 | 1814 | have "c * (l div c) = c * (l div c) + l mod c" | 
| 55921 | 1815 | using d' dvd_eq_mod_eq_0[of "c" "l"] by simp | 
| 1816 | then have cl: "c * (l div c) =l" | |
| 1817 | using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 1818 | by simp | |
| 55964 | 1819 | then have "l * x + (l div c) * Inum (x # bs) e \<ge> 0 \<longleftrightarrow> | 
| 1820 | (c * (l div c)) * x + (l div c) * Inum (x # bs) e \<ge> 0" | |
| 55921 | 1821 | by simp | 
| 55964 | 1822 | also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) \<ge> (l div c) * 0" | 
| 55921 | 1823 | by (simp add: algebra_simps) | 
| 55964 | 1824 | also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e \<ge> 0" | 
| 55921 | 1825 | using ldcp zero_le_mult_iff [where a="l div c" and b="c*x + Inum (x # bs) e"] | 
| 1826 | by simp | |
| 1827 | finally show ?case | |
| 1828 | using be numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] | |
| 1829 | by simp | |
| 1830 | next | |
| 1831 | case (3 c e) | |
| 1832 | then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l" | |
| 1833 | by simp_all | |
| 1834 | from lp cp have clel: "c \<le> l" | |
| 1835 | by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1836 | from cp have cnz: "c \<noteq> 0" | |
| 1837 | by simp | |
| 1838 | have "c div c \<le> l div c" | |
| 50313 | 1839 | by (simp add: zdiv_mono1[OF clel cp]) | 
| 1840 | then have ldcp:"0 < l div c" | |
| 1841 | by (simp add: div_self[OF cnz]) | |
| 55964 | 1842 | have "c * (l div c) = c * (l div c) + l mod c" | 
| 50313 | 1843 | using d' dvd_eq_mod_eq_0[of "c" "l"] by simp | 
| 55885 | 1844 | then have cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] | 
| 50313 | 1845 | by simp | 
| 55964 | 1846 | then have "l * x + (l div c) * Inum (x # bs) e = 0 \<longleftrightarrow> | 
| 1847 | (c * (l div c)) * x + (l div c) * Inum (x # bs) e = 0" | |
| 23274 | 1848 | by simp | 
| 55964 | 1849 | also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) = ((l div c)) * 0" | 
| 50313 | 1850 | by (simp add: algebra_simps) | 
| 55964 | 1851 | also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e = 0" | 
| 55921 | 1852 | using mult_eq_0_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp | 
| 1853 | by simp | |
| 50313 | 1854 | finally show ?case | 
| 55921 | 1855 | using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be | 
| 1856 | 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 | 1857 | next | 
| 50313 | 1858 | case (4 c e) | 
| 55921 | 1859 | then have cp: "c > 0" and be: "numbound0 e" and d': "c dvd l" | 
| 1860 | by simp_all | |
| 1861 | from lp cp have clel: "c \<le> l" | |
| 1862 | by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1863 | from cp have cnz: "c \<noteq> 0" | |
| 1864 | by simp | |
| 1865 | have "c div c \<le> l div c" | |
| 50313 | 1866 | by (simp add: zdiv_mono1[OF clel cp]) | 
| 1867 | then have ldcp:"0 < l div c" | |
| 1868 | by (simp add: div_self[OF cnz]) | |
| 55964 | 1869 | have "c * (l div c) = c * (l div c) + l mod c" | 
| 55921 | 1870 | using d' dvd_eq_mod_eq_0[of "c" "l"] by simp | 
| 1871 | then have cl: "c * (l div c) = l" | |
| 1872 | using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp | |
| 55964 | 1873 | then have "l * x + (l div c) * Inum (x # bs) e \<noteq> 0 \<longleftrightarrow> | 
| 55921 | 1874 | (c * (l div c)) * x + (l div c) * Inum (x # bs) e \<noteq> 0" | 
| 50313 | 1875 | by simp | 
| 55921 | 1876 | also have "\<dots> \<longleftrightarrow> (l div c) * (c * x + Inum (x # bs) e) \<noteq> (l div c) * 0" | 
| 50313 | 1877 | by (simp add: algebra_simps) | 
| 55921 | 1878 | also have "\<dots> \<longleftrightarrow> c * x + Inum (x # bs) e \<noteq> 0" | 
| 1879 | using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp | |
| 1880 | by simp | |
| 50313 | 1881 | finally show ?case | 
| 55921 | 1882 | using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be | 
| 1883 | 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 | 1884 | next | 
| 50313 | 1885 | case (9 j c e) | 
| 55921 | 1886 | then have cp: "c > 0" and be: "numbound0 e" and jp: "j > 0" and d': "c dvd l" | 
| 1887 | by simp_all | |
| 1888 | from lp cp have clel: "c \<le> l" | |
| 1889 | by (simp add: zdvd_imp_le [OF d' lp]) | |
| 50313 | 1890 | from cp have cnz: "c \<noteq> 0" by simp | 
| 1891 | have "c div c\<le> l div c" | |
| 1892 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1893 | then have ldcp:"0 < l div c" | |
| 1894 | by (simp add: div_self[OF cnz]) | |
| 55964 | 1895 | have "c * (l div c) = c * (l div c) + l mod c" | 
| 55885 | 1896 | using d' dvd_eq_mod_eq_0[of "c" "l"] by simp | 
| 55921 | 1897 | then have cl: "c * (l div c) = l" | 
| 55885 | 1898 | using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp | 
| 55921 | 1899 | then have "(\<exists>k::int. l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) \<longleftrightarrow> | 
| 55964 | 1900 | (\<exists>k::int. (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)" | 
| 55921 | 1901 | by simp | 
| 55964 | 1902 | also have "\<dots> \<longleftrightarrow> (\<exists>k::int. (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c) * 0)" | 
| 55921 | 1903 | by (simp add: algebra_simps) | 
| 60708 | 1904 | also have "\<dots> \<longleftrightarrow> (\<exists>k::int. c * x + Inum (x # bs) e - j * k = 0)" | 
| 1905 | using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k" for k] ldcp | |
| 55921 | 1906 | by simp | 
| 1907 | also have "\<dots> \<longleftrightarrow> (\<exists>k::int. c * x + Inum (x # bs) e = j * k)" | |
| 1908 | by simp | |
| 1909 | finally show ?case | |
| 1910 | using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] | |
| 1911 | be mult_strict_mono[OF ldcp jp ldcp ] | |
| 1912 | by (simp add: dvd_def) | |
| 1913 | next | |
| 1914 | case (10 j c e) | |
| 1915 | then have cp: "c > 0" and be: "numbound0 e" and jp: "j > 0" and d': "c dvd l" | |
| 1916 | by simp_all | |
| 1917 | from lp cp have clel: "c \<le> l" | |
| 1918 | by (simp add: zdvd_imp_le [OF d' lp]) | |
| 1919 | from cp have cnz: "c \<noteq> 0" | |
| 50313 | 1920 | by simp | 
| 55921 | 1921 | have "c div c \<le> l div c" | 
| 1922 | by (simp add: zdiv_mono1[OF clel cp]) | |
| 1923 | then have ldcp: "0 < l div c" | |
| 1924 | by (simp add: div_self[OF cnz]) | |
| 1925 | have "c * (l div c) = c* (l div c) + l mod c" | |
| 1926 | using d' dvd_eq_mod_eq_0[of "c" "l"] by simp | |
| 1927 | then have cl:"c * (l div c) =l" | |
| 1928 | using zmod_zdiv_equality[where a="l" and b="c", symmetric] | |
| 1929 | by simp | |
| 1930 | then have "(\<exists>k::int. l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) \<longleftrightarrow> | |
| 1931 | (\<exists>k::int. (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)" | |
| 1932 | by simp | |
| 1933 | also have "\<dots> \<longleftrightarrow> (\<exists>k::int. (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c) * 0)" | |
| 1934 | by (simp add: algebra_simps) | |
| 60708 | 1935 | also have "\<dots> \<longleftrightarrow> (\<exists>k::int. c * x + Inum (x # bs) e - j * k = 0)" | 
| 1936 | using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k" for k] ldcp | |
| 55921 | 1937 | by simp | 
| 55964 | 1938 | also have "\<dots> \<longleftrightarrow> (\<exists>k::int. c * x + Inum (x # bs) e = j * k)" | 
| 55921 | 1939 | by simp | 
| 1940 | finally show ?case | |
| 1941 | using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be | |
| 1942 | mult_strict_mono[OF ldcp jp ldcp ] | |
| 1943 | by (simp add: dvd_def) | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 1944 | 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 | 1945 | |
| 55921 | 1946 | lemma a_\<beta>_ex: | 
| 1947 | assumes linp: "iszlfm p" | |
| 1948 | and d: "d_\<beta> p l" | |
| 1949 | and lp: "l > 0" | |
| 1950 | shows "(\<exists>x. l dvd x \<and> Ifm bbs (x #bs) (a_\<beta> p l)) \<longleftrightarrow> (\<exists>x::int. Ifm bbs (x#bs) p)" | |
| 1951 | (is "(\<exists>x. l dvd x \<and> ?P x) \<longleftrightarrow> (\<exists>x. ?P' x)") | |
| 23274 | 1952 | proof- | 
| 55999 | 1953 | have "(\<exists>x. l dvd x \<and> ?P x) \<longleftrightarrow> (\<exists>x::int. ?P (l * x))" | 
| 23274 | 1954 | using unity_coeff_ex[where l="l" and P="?P", simplified] by simp | 
| 55921 | 1955 | also have "\<dots> = (\<exists>x::int. ?P' x)" | 
| 1956 | using a_\<beta>[OF linp d lp] by simp | |
| 50313 | 1957 | 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 | 1958 | 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 | 1959 | |
| 23274 | 1960 | lemma \<beta>: | 
| 55999 | 1961 | assumes "iszlfm p" | 
| 1962 | and "d_\<beta> p 1" | |
| 1963 | and "d_\<delta> p d" | |
| 55885 | 1964 | and dp: "d > 0" | 
| 55999 | 1965 |     and "\<not> (\<exists>j::int \<in> {1 .. d}. \<exists>b \<in> Inum (a # bs) ` set (\<beta> p). x = b + j)"
 | 
| 55964 | 1966 | and p: "Ifm bbs (x # bs) p" (is "?P x") | 
| 23274 | 1967 | shows "?P (x - d)" | 
| 55999 | 1968 | using assms | 
| 55885 | 1969 | proof (induct p rule: iszlfm.induct) | 
| 1970 | case (5 c e) | |
| 1971 | then have c1: "c = 1" and bn: "numbound0 e" | |
| 1972 | by simp_all | |
| 55964 | 1973 | with dp p c1 numbound0_I[OF bn,where b = "(x - d)" and b' = "x" and bs = "bs"] 5 | 
| 41807 | 1974 | show ?case by simp | 
| 23274 | 1975 | next | 
| 55885 | 1976 | case (6 c e) | 
| 1977 | then have c1: "c = 1" and bn: "numbound0 e" | |
| 1978 | by simp_all | |
| 41807 | 1979 | with dp p c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] 6 | 
| 1980 | show ?case by simp | |
| 23274 | 1981 | next | 
| 55885 | 1982 | case (7 c e) | 
| 55964 | 1983 | then have p: "Ifm bbs (x # bs) (Gt (CN 0 c e))" and c1: "c=1" and bn: "numbound0 e" | 
| 55885 | 1984 | by simp_all | 
| 41807 | 1985 | let ?e = "Inum (x # bs) e" | 
| 60708 | 1986 | show ?case | 
| 1987 | proof (cases "(x - d) + ?e > 0") | |
| 1988 | case True | |
| 1989 | then show ?thesis | |
| 55885 | 1990 | using c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] by simp | 
| 60708 | 1991 | next | 
| 1992 | case False | |
| 55964 | 1993 | let ?v = "Neg e" | 
| 1994 | have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" | |
| 1995 | by simp | |
| 57816 
d8bbb97689d3
no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
 blanchet parents: 
57514diff
changeset | 1996 | from 7(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]] | 
| 55964 | 1997 |     have nob: "\<not> (\<exists>j\<in> {1 ..d}. x = - ?e + j)"
 | 
| 55885 | 1998 | by auto | 
| 60708 | 1999 | from False p have "x + ?e > 0 \<and> x + ?e \<le> d" | 
| 55885 | 2000 | by (simp add: c1) | 
| 2001 | then have "x + ?e \<ge> 1 \<and> x + ?e \<le> d" | |
| 2002 | by simp | |
| 55964 | 2003 |     then have "\<exists>j::int \<in> {1 .. d}. j = x + ?e"
 | 
| 55885 | 2004 | by simp | 
| 55964 | 2005 |     then have "\<exists>j::int \<in> {1 .. d}. x = (- ?e + j)"
 | 
| 41807 | 2006 | by (simp add: algebra_simps) | 
| 60708 | 2007 | with nob show ?thesis | 
| 55885 | 2008 | by auto | 
| 60708 | 2009 | qed | 
| 23274 | 2010 | next | 
| 55885 | 2011 | case (8 c e) | 
| 2012 | then have p: "Ifm bbs (x # bs) (Ge (CN 0 c e))" and c1: "c = 1" and bn: "numbound0 e" | |
| 50313 | 2013 | by simp_all | 
| 55885 | 2014 | let ?e = "Inum (x # bs) e" | 
| 60708 | 2015 | show ?case | 
| 2016 | proof (cases "(x - d) + ?e \<ge> 0") | |
| 2017 | case True | |
| 2018 | then show ?thesis | |
| 55885 | 2019 | using c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] | 
| 2020 | by simp | |
| 60708 | 2021 | next | 
| 2022 | case False | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58310diff
changeset | 2023 | let ?v = "Sub (C (- 1)) e" | 
| 55885 | 2024 | have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" | 
| 2025 | by simp | |
| 57816 
d8bbb97689d3
no need for 'set_simps' now that 'datatype_new' generates the desired 'set' property
 blanchet parents: 
57514diff
changeset | 2026 | from 8(5)[simplified simp_thms Inum.simps \<beta>.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]] | 
| 55885 | 2027 |     have nob: "\<not> (\<exists>j\<in> {1 ..d}. x =  - ?e - 1 + j)"
 | 
| 2028 | by auto | |
| 60708 | 2029 | from False p have "x + ?e \<ge> 0 \<and> x + ?e < d" | 
| 55885 | 2030 | by (simp add: c1) | 
| 2031 | then have "x + ?e +1 \<ge> 1 \<and> x + ?e + 1 \<le> d" | |
| 2032 | by simp | |
| 55964 | 2033 |     then have "\<exists>j::int \<in> {1 .. d}. j = x + ?e + 1"
 | 
| 55885 | 2034 | by simp | 
| 55964 | 2035 |     then have "\<exists>j::int \<in> {1 .. d}. x= - ?e - 1 + j"
 | 
| 55885 | 2036 | by (simp add: algebra_simps) | 
| 60708 | 2037 | with nob show ?thesis | 
| 55885 | 2038 | by simp | 
| 60708 | 2039 | qed | 
| 23274 | 2040 | next | 
| 55885 | 2041 | case (3 c e) | 
| 2042 | then | |
| 2043 | have p: "Ifm bbs (x #bs) (Eq (CN 0 c e))" (is "?p x") | |
| 55964 | 2044 | and c1: "c = 1" | 
| 55885 | 2045 | and bn: "numbound0 e" | 
| 2046 | by simp_all | |
| 2047 | let ?e = "Inum (x # bs) e" | |
| 58410 
6d46ad54a2ab
explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
 haftmann parents: 
58310diff
changeset | 2048 | let ?v="(Sub (C (- 1)) e)" | 
| 55885 | 2049 | have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" | 
| 2050 | by simp | |
| 55964 | 2051 | from p have "x= - ?e" | 
| 2052 | by (simp add: c1) with 3(5) | |
| 2053 | show ?case | |
| 60708 | 2054 | using dp apply simp | 
| 2055 | apply (erule ballE[where x="1"]) | |
| 2056 | apply (simp_all add:algebra_simps numbound0_I[OF bn,where b="x"and b'="a"and bs="bs"]) | |
| 2057 | done | |
| 23274 | 2058 | next | 
| 55885 | 2059 | case (4 c e) | 
| 2060 | then | |
| 55964 | 2061 | have p: "Ifm bbs (x # bs) (NEq (CN 0 c e))" (is "?p x") | 
| 55885 | 2062 | and c1: "c = 1" | 
| 2063 | and bn: "numbound0 e" | |
| 2064 | by simp_all | |
| 2065 | let ?e = "Inum (x # bs) e" | |
| 2066 | let ?v="Neg e" | |
| 55964 | 2067 | have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))" | 
| 2068 | by simp | |
| 60708 | 2069 | show ?case | 
| 2070 | proof (cases "x - d + Inum ((x - d) # bs) e = 0") | |
| 2071 | case False | |
| 2072 | then show ?thesis by (simp add: c1) | |
| 2073 | next | |
| 2074 | case True | |
| 55964 | 2075 | then have "x = - Inum ((x - d) # bs) e + d" | 
| 55885 | 2076 | by simp | 
| 2077 | then have "x = - Inum (a # bs) e + d" | |
| 2078 | by (simp add: numbound0_I[OF bn,where b="x - d"and b'="a"and bs="bs"]) | |
| 60708 | 2079 | with 4(5) show ?thesis | 
| 55885 | 2080 | using dp by simp | 
| 60708 | 2081 | qed | 
| 50313 | 2082 | next | 
| 55885 | 2083 | case (9 j c e) | 
| 2084 | then | |
| 2085 | have p: "Ifm bbs (x # bs) (Dvd j (CN 0 c e))" (is "?p x") | |
| 2086 | and c1: "c = 1" | |
| 2087 | and bn: "numbound0 e" | |
| 2088 | by simp_all | |
| 2089 | let ?e = "Inum (x # bs) e" | |
| 2090 | from 9 have id: "j dvd d" | |
| 2091 | by simp | |
| 55964 | 2092 | from c1 have "?p x \<longleftrightarrow> j dvd (x + ?e)" | 
| 55885 | 2093 | by simp | 
| 55964 | 2094 | also have "\<dots> \<longleftrightarrow> j dvd x - d + ?e" | 
| 55885 | 2095 | using zdvd_period[OF id, where x="x" and c="-1" and t="?e"] | 
| 2096 | by simp | |
| 2097 | finally show ?case | |
| 2098 | using numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] c1 p | |
| 2099 | by simp | |
| 23274 | 2100 | next | 
| 55885 | 2101 | case (10 j c e) | 
| 2102 | then | |
| 55964 | 2103 | have p: "Ifm bbs (x # bs) (NDvd j (CN 0 c e))" (is "?p x") | 
| 55885 | 2104 | and c1: "c = 1" | 
| 2105 | and bn: "numbound0 e" | |
| 2106 | by simp_all | |
| 2107 | let ?e = "Inum (x # bs) e" | |
| 2108 | from 10 have id: "j dvd d" | |
| 2109 | by simp | |
| 55964 | 2110 | from c1 have "?p x \<longleftrightarrow> \<not> j dvd (x + ?e)" | 
| 55885 | 2111 | by simp | 
| 55964 | 2112 | also have "\<dots> \<longleftrightarrow> \<not> j dvd x - d + ?e" | 
| 55885 | 2113 | using zdvd_period[OF id, where x="x" and c="-1" and t="?e"] | 
| 2114 | by simp | |
| 2115 | finally show ?case | |
| 2116 | using numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] c1 p | |
| 2117 | by simp | |
| 23689 
0410269099dc
replaced code generator framework for reflected cooper
 haftmann parents: 
23515diff
changeset | 2118 | 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 | 2119 | |
| 50313 | 2120 | lemma \<beta>': | 
| 23274 | 2121 | assumes lp: "iszlfm p" | 
| 60708 | 2122 | and u: "d_\<beta> p 1" | 
| 2123 | and d: "d_\<delta> p d" | |
| 2124 | and dp: "d > 0" | |
| 55964 | 2125 |   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>
 | 
| 2126 | Ifm bbs (x # bs) p \<longrightarrow> Ifm bbs ((x - d) # bs) p" (is "\<forall>x. ?b \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)") | |
| 55885 | 2127 | proof clarify | 
| 50313 | 2128 | fix x | 
| 60708 | 2129 | assume nb: "?b" and px: "?P x" | 
| 55964 | 2130 |   then have nb2: "\<not> (\<exists>j::int \<in> {1 .. d}. \<exists>b \<in> Inum (a # bs) ` set (\<beta> p). x = b + j)"
 | 
| 23274 | 2131 | by auto | 
| 60708 | 2132 | show "?P (x - d)" by (rule \<beta>[OF lp u d dp nb2 px]) | 
| 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 | 2133 | qed | 
| 55885 | 2134 | |
| 2135 | lemma cpmi_eq: | |
| 55999 | 2136 | fixes P P1 :: "int \<Rightarrow> bool" | 
| 2137 | assumes "0 < D" | |
| 2138 | and "\<exists>z. \<forall>x. x < z \<longrightarrow> P x = P1 x" | |
| 2139 |     and "\<forall>x. \<not> (\<exists>j \<in> {1..D}. \<exists>b \<in> B. P (b + j)) \<longrightarrow> P x \<longrightarrow> P (x - D)"
 | |
| 2140 | and "\<forall>x k. P1 x = P1 (x - k * D)" | |
| 2141 |   shows "(\<exists>x. P x) \<longleftrightarrow> (\<exists>j \<in> {1..D}. P1 j) \<or> (\<exists>j \<in> {1..D}. \<exists>b \<in> B. P (b + j))"
 | |
| 2142 | apply (insert assms) | |
| 2143 | apply (rule iffI) | |
| 55885 | 2144 | prefer 2 | 
| 55981 | 2145 | apply (drule minusinfinity) | 
| 55885 | 2146 | apply assumption+ | 
| 55981 | 2147 | apply fastforce | 
| 55885 | 2148 | apply clarsimp | 
| 55981 | 2149 | apply (subgoal_tac "\<And>k. 0 \<le> k \<Longrightarrow> \<forall>x. P x \<longrightarrow> P (x - k * D)") | 
| 2150 | apply (frule_tac x = x and z=z in decr_lemma) | |
| 2151 | apply (subgoal_tac "P1 (x - (\<bar>x - z\<bar> + 1) * D)") | |
| 55885 | 2152 | prefer 2 | 
| 55981 | 2153 | apply (subgoal_tac "0 \<le> \<bar>x - z\<bar> + 1") | 
| 55885 | 2154 | prefer 2 apply arith | 
| 2155 | apply fastforce | |
| 55981 | 2156 | apply (drule (1) periodic_finite_ex) | 
| 55885 | 2157 | apply blast | 
| 55981 | 2158 | apply (blast dest: decr_mult_lemma) | 
| 55885 | 2159 | 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 | 2160 | |
| 23274 | 2161 | theorem cp_thm: | 
| 2162 | assumes lp: "iszlfm p" | |
| 55885 | 2163 | and u: "d_\<beta> p 1" | 
| 2164 | and d: "d_\<delta> p d" | |
| 2165 | and dp: "d > 0" | |
| 55999 | 2166 | shows "(\<exists>x. Ifm bbs (x # bs) p) \<longleftrightarrow> | 
| 55964 | 2167 |     (\<exists>j \<in> {1.. d}. Ifm bbs (j # bs) (minusinf p) \<or>
 | 
| 2168 | (\<exists>b \<in> set (\<beta> p). Ifm bbs ((Inum (i # bs) b + j) # bs) p))" | |
| 55999 | 2169 | (is "(\<exists>x. ?P x) \<longleftrightarrow> (\<exists>j \<in> ?D. ?M j \<or> (\<exists>b \<in> ?B. ?P (?I b + j)))") | 
| 55885 | 2170 | proof - | 
| 50313 | 2171 | from minusinf_inf[OF lp u] | 
| 55999 | 2172 | have th: "\<exists>z. \<forall>x<z. ?P x = ?M x" | 
| 55885 | 2173 | by blast | 
| 55964 | 2174 |   let ?B' = "{?I b | b. b \<in> ?B}"
 | 
| 2175 | have BB': "(\<exists>j\<in>?D. \<exists>b \<in> ?B. ?P (?I b + j)) \<longleftrightarrow> (\<exists>j \<in> ?D. \<exists>b \<in> ?B'. ?P (b + j))" | |
| 55885 | 2176 | by auto | 
| 55964 | 2177 | then have th2: "\<forall>x. \<not> (\<exists>j \<in> ?D. \<exists>b \<in> ?B'. ?P (b + j)) \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)" | 
| 23274 | 2178 | using \<beta>'[OF lp u d dp, where a="i" and bbs = "bbs"] by blast | 
| 2179 | from minusinf_repeats[OF d lp] | |
| 55885 | 2180 | have th3: "\<forall>x k. ?M x = ?M (x-k*d)" | 
| 2181 | by simp | |
| 2182 | from cpmi_eq[OF dp th th2 th3] BB' show ?thesis | |
| 2183 | 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 | 2184 | 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 | 2185 | |
| 55885 | 2186 | (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *) | 
| 50313 | 2187 | lemma mirror_ex: | 
| 55999 | 2188 | assumes "iszlfm p" | 
| 2189 | shows "(\<exists>x. Ifm bbs (x#bs) (mirror p)) \<longleftrightarrow> (\<exists>x. Ifm bbs (x#bs) p)" | |
| 50313 | 2190 | (is "(\<exists>x. ?I x ?mp) = (\<exists>x. ?I x p)") | 
| 55964 | 2191 | proof auto | 
| 2192 | fix x | |
| 2193 | assume "?I x ?mp" | |
| 2194 | then have "?I (- x) p" | |
| 55999 | 2195 | using mirror[OF assms] by blast | 
| 55964 | 2196 | then show "\<exists>x. ?I x p" | 
| 2197 | by blast | |
| 23274 | 2198 | next | 
| 55964 | 2199 | fix x | 
| 2200 | assume "?I x p" | |
| 2201 | then have "?I (- x) ?mp" | |
| 55999 | 2202 | using mirror[OF assms, where x="- x", symmetric] by auto | 
| 55964 | 2203 | then show "\<exists>x. ?I x ?mp" | 
| 2204 | by blast | |
| 23274 | 2205 | qed | 
| 24349 | 2206 | |
| 50313 | 2207 | lemma cp_thm': | 
| 55999 | 2208 | assumes "iszlfm p" | 
| 2209 | and "d_\<beta> p 1" | |
| 2210 | and "d_\<delta> p d" | |
| 2211 | and "d > 0" | |
| 55964 | 2212 | shows "(\<exists>x. Ifm bbs (x # bs) p) \<longleftrightarrow> | 
| 2213 |     ((\<exists>j\<in> {1 .. d}. Ifm bbs (j#bs) (minusinf p)) \<or>
 | |
| 2214 |       (\<exists>j\<in> {1.. d}. \<exists>b\<in> (Inum (i#bs)) ` set (\<beta> p). Ifm bbs ((b + j) # bs) p))"
 | |
| 55999 | 2215 | using cp_thm[OF assms,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 | 2216 | |
| 50313 | 2217 | definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int" | 
| 2218 | where | |
| 55964 | 2219 | "unit p = | 
| 2220 | (let | |
| 2221 | p' = zlfm p; | |
| 2222 | l = \<zeta> p'; | |
| 2223 | q = And (Dvd l (CN 0 1 (C 0))) (a_\<beta> p' l); | |
| 2224 | d = \<delta> q; | |
| 2225 | B = remdups (map simpnum (\<beta> q)); | |
| 2226 | a = remdups (map simpnum (\<alpha> q)) | |
| 2227 | 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 | 2228 | |
| 50313 | 2229 | lemma unit: | 
| 2230 | assumes qf: "qfree p" | |
| 60708 | 2231 | fixes q B d | 
| 2232 | assumes qBd: "unit p = (q, B, d)" | |
| 2233 | shows "((\<exists>x. Ifm bbs (x # bs) p) \<longleftrightarrow> (\<exists>x. Ifm bbs (x # bs) q)) \<and> | |
| 55964 | 2234 | (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> | 
| 2235 | iszlfm q \<and> (\<forall>b\<in> set B. numbound0 b)" | |
| 50313 | 2236 | proof - | 
| 2237 | let ?I = "\<lambda>x p. Ifm bbs (x#bs) p" | |
| 23274 | 2238 | let ?p' = "zlfm p" | 
| 2239 | let ?l = "\<zeta> ?p'" | |
| 50252 | 2240 | let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a_\<beta> ?p' ?l)" | 
| 23274 | 2241 | let ?d = "\<delta> ?q" | 
| 2242 | let ?B = "set (\<beta> ?q)" | |
| 2243 | let ?B'= "remdups (map simpnum (\<beta> ?q))" | |
| 2244 | let ?A = "set (\<alpha> ?q)" | |
| 2245 | let ?A'= "remdups (map simpnum (\<alpha> ?q))" | |
| 50313 | 2246 | from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] | 
| 2247 | have pp': "\<forall>i. ?I i ?p' = ?I i p" by auto | |
| 23274 | 2248 | from conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]] | 
| 50313 | 2249 | have lp': "iszlfm ?p'" . | 
| 50252 | 2250 | from lp' \<zeta>[where p="?p'"] have lp: "?l >0" and dl: "d_\<beta> ?p' ?l" by auto | 
| 2251 | from a_\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp' dl lp] pp' | |
| 50313 | 2252 | have pq_ex:"(\<exists>(x::int). ?I x p) = (\<exists>x. ?I x ?q)" by simp | 
| 50252 | 2253 | from lp' lp a_\<beta>[OF lp' dl lp] have lq:"iszlfm ?q" and uq: "d_\<beta> ?q 1" by auto | 
| 2254 | from \<delta>[OF lq] have dp:"?d >0" and dd: "d_\<delta> ?q ?d" by blast+ | |
| 50313 | 2255 | let ?N = "\<lambda>t. Inum (i#bs) t" | 
| 55981 | 2256 | have "?N ` set ?B' = ((?N \<circ> simpnum) ` ?B)" | 
| 2257 | by auto | |
| 2258 | also have "\<dots> = ?N ` ?B" | |
| 2259 | using simpnum_ci[where bs="i#bs"] by auto | |
| 23274 | 2260 | finally have BB': "?N ` set ?B' = ?N ` ?B" . | 
| 55981 | 2261 | have "?N ` set ?A' = ((?N \<circ> simpnum) ` ?A)" | 
| 2262 | by auto | |
| 2263 | also have "\<dots> = ?N ` ?A" | |
| 2264 | using simpnum_ci[where bs="i#bs"] by auto | |
| 23274 | 2265 | finally have AA': "?N ` set ?A' = ?N ` ?A" . | 
| 50313 | 2266 | from \<beta>_numbound0[OF lq] have B_nb:"\<forall>b\<in> set ?B'. numbound0 b" | 
| 23274 | 2267 | by (simp add: simpnum_numbound0) | 
| 50313 | 2268 | from \<alpha>_l[OF lq] have A_nb: "\<forall>b\<in> set ?A'. numbound0 b" | 
| 23274 | 2269 | by (simp add: simpnum_numbound0) | 
| 60708 | 2270 | show ?thesis | 
| 2271 | proof (cases "length ?B' \<le> length ?A'") | |
| 2272 | case True | |
| 55981 | 2273 | then have q: "q = ?q" and "B = ?B'" and d: "d = ?d" | 
| 23274 | 2274 | using qBd by (auto simp add: Let_def unit_def) | 
| 55981 | 2275 | with BB' B_nb | 
| 2276 | have b: "?N ` (set B) = ?N ` set (\<beta> q)" and bn: "\<forall>b\<in> set B. numbound0 b" | |
| 2277 | by simp_all | |
| 60708 | 2278 | with pq_ex dp uq dd lq q d show ?thesis | 
| 55981 | 2279 | by simp | 
| 60708 | 2280 | next | 
| 2281 | case False | |
| 55885 | 2282 | then have q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" | 
| 23274 | 2283 | using qBd by (auto simp add: Let_def unit_def) | 
| 50313 | 2284 | with AA' mirror_\<alpha>_\<beta>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<beta> q)" | 
| 2285 | and bn: "\<forall>b\<in> set B. numbound0 b" by simp_all | |
| 2286 | from mirror_ex[OF lq] pq_ex q | |
| 55981 | 2287 | have pqm_eq:"(\<exists>(x::int). ?I x p) = (\<exists>(x::int). ?I x q)" | 
| 2288 | by simp | |
| 23274 | 2289 | from lq uq q mirror_l[where p="?q"] | 
| 55981 | 2290 | have lq': "iszlfm q" and uq: "d_\<beta> q 1" | 
| 2291 | by auto | |
| 2292 | from \<delta>[OF lq'] mirror_\<delta>[OF lq] q d have dq: "d_\<delta> q d" | |
| 2293 | by auto | |
| 60708 | 2294 | from pqm_eq b bn uq lq' dp dq q dp d show ?thesis | 
| 55981 | 2295 | by simp | 
| 60708 | 2296 | qed | 
| 23274 | 2297 | qed | 
| 50313 | 2298 | |
| 2299 | ||
| 60533 | 2300 | text \<open>Cooper's Algorithm\<close> | 
| 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 | 2301 | |
| 55981 | 2302 | definition cooper :: "fm \<Rightarrow> fm" | 
| 2303 | where | |
| 50313 | 2304 | "cooper p = | 
| 2305 | (let | |
| 2306 | (q, B, d) = unit p; | |
| 2307 | js = [1..d]; | |
| 2308 | mq = simpfm (minusinf q); | |
| 2309 | md = evaldjf (\<lambda>j. simpfm (subst0 (C j) mq)) js | |
| 2310 | in | |
| 2311 | if md = T then T | |
| 2312 | else | |
| 2313 | (let | |
| 2314 | qd = evaldjf (\<lambda>(b, j). simpfm (subst0 (Add b (C j)) q)) [(b, j). b \<leftarrow> B, j \<leftarrow> js] | |
| 2315 | in decr (disj md qd)))" | |
| 2316 | ||
| 2317 | lemma cooper: | |
| 2318 | assumes qf: "qfree p" | |
| 60708 | 2319 | shows "(\<exists>x. Ifm bbs (x#bs) p) = Ifm bbs bs (cooper p) \<and> qfree (cooper p)" | 
| 2320 | (is "?lhs = ?rhs \<and> _") | |
| 50313 | 2321 | proof - | 
| 2322 | let ?I = "\<lambda>x p. Ifm bbs (x#bs) p" | |
| 23274 | 2323 | let ?q = "fst (unit p)" | 
| 2324 | let ?B = "fst (snd(unit p))" | |
| 2325 | let ?d = "snd (snd (unit p))" | |
| 41836 | 2326 | let ?js = "[1..?d]" | 
| 23274 | 2327 | let ?mq = "minusinf ?q" | 
| 2328 | let ?smq = "simpfm ?mq" | |
| 50313 | 2329 | let ?md = "evaldjf (\<lambda>j. simpfm (subst0 (C j) ?smq)) ?js" | 
| 26934 | 2330 | fix i | 
| 50313 | 2331 | let ?N = "\<lambda>t. Inum (i#bs) t" | 
| 24336 | 2332 | let ?Bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]" | 
| 50313 | 2333 | let ?qd = "evaldjf (\<lambda>(b,j). simpfm (subst0 (Add b (C j)) ?q)) ?Bjs" | 
| 23274 | 2334 | have qbf:"unit p = (?q,?B,?d)" by simp | 
| 55981 | 2335 | from unit[OF qf qbf] | 
| 2336 | have pq_ex: "(\<exists>(x::int). ?I x p) \<longleftrightarrow> (\<exists>(x::int). ?I x ?q)" | |
| 2337 | and B: "?N ` set ?B = ?N ` set (\<beta> ?q)" | |
| 2338 | and uq: "d_\<beta> ?q 1" | |
| 2339 | and dd: "d_\<delta> ?q ?d" | |
| 2340 | and dp: "?d > 0" | |
| 2341 | and lq: "iszlfm ?q" | |
| 2342 | and Bn: "\<forall>b\<in> set ?B. numbound0 b" | |
| 2343 | by auto | |
| 23274 | 2344 | from zlin_qfree[OF lq] have qfq: "qfree ?q" . | 
| 55921 | 2345 | from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq" . | 
| 55981 | 2346 | have jsnb: "\<forall>j \<in> set ?js. numbound0 (C j)" | 
| 2347 | by simp | |
| 55885 | 2348 | then have "\<forall>j\<in> set ?js. bound0 (subst0 (C j) ?smq)" | 
| 23274 | 2349 | by (auto simp only: subst0_bound0[OF qfmq]) | 
| 55885 | 2350 | then have th: "\<forall>j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))" | 
| 23274 | 2351 | by (auto simp add: simpfm_bound0) | 
| 55981 | 2352 | from evaldjf_bound0[OF th] have mdb: "bound0 ?md" | 
| 2353 | by simp | |
| 50313 | 2354 | 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 | 2355 | by simp | 
| 55885 | 2356 | then have "\<forall>(b,j) \<in> set ?Bjs. bound0 (subst0 (Add b (C j)) ?q)" | 
| 23274 | 2357 | using subst0_bound0[OF qfq] by blast | 
| 55885 | 2358 | then have "\<forall>(b,j) \<in> set ?Bjs. bound0 (simpfm (subst0 (Add b (C j)) ?q))" | 
| 55981 | 2359 | using simpfm_bound0 by blast | 
| 55885 | 2360 | then have th': "\<forall>x \<in> set ?Bjs. bound0 ((\<lambda>(b,j). simpfm (subst0 (Add b (C j)) ?q)) x)" | 
| 50313 | 2361 | by auto | 
| 55981 | 2362 | from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" | 
| 2363 | by simp | |
| 2364 | from mdb qdb have mdqdb: "bound0 (disj ?md ?qd)" | |
| 2365 | unfolding disj_def by (cases "?md = T \<or> ?qd = T") simp_all | |
| 23274 | 2366 | from trans [OF pq_ex cp_thm'[OF lq uq dd dp,where i="i"]] B | 
| 55981 | 2367 |   have "?lhs \<longleftrightarrow> (\<exists>j \<in> {1.. ?d}. ?I j ?mq \<or> (\<exists>b \<in> ?N ` set ?B. Ifm bbs ((b + j) # bs) ?q))"
 | 
| 55921 | 2368 | by auto | 
| 55981 | 2369 |   also have "\<dots> \<longleftrightarrow> (\<exists>j \<in> {1.. ?d}. ?I j ?mq \<or> (\<exists>b \<in> set ?B. Ifm bbs ((?N b + j) # bs) ?q))"
 | 
| 55921 | 2370 | by simp | 
| 55981 | 2371 |   also have "\<dots> \<longleftrightarrow> (\<exists>j \<in> {1.. ?d}. ?I j ?mq ) \<or>
 | 
| 2372 |       (\<exists>j\<in> {1.. ?d}. \<exists>b \<in> set ?B. Ifm bbs ((?N (Add b (C j))) # bs) ?q)"
 | |
| 50313 | 2373 | by (simp only: Inum.simps) blast | 
| 55981 | 2374 |   also have "\<dots> \<longleftrightarrow> (\<exists>j \<in> {1.. ?d}. ?I j ?smq) \<or>
 | 
| 2375 |       (\<exists>j \<in> {1.. ?d}. \<exists>b \<in> set ?B. Ifm bbs ((?N (Add b (C j))) # bs) ?q)"
 | |
| 50313 | 2376 | by (simp add: simpfm) | 
| 55981 | 2377 | also have "\<dots> \<longleftrightarrow> (\<exists>j\<in> set ?js. (\<lambda>j. ?I i (simpfm (subst0 (C j) ?smq))) j) \<or> | 
| 2378 | (\<exists>j\<in> set ?js. \<exists>b\<in> set ?B. Ifm bbs ((?N (Add b (C j)))#bs) ?q)" | |
| 41836 | 2379 | by (simp only: simpfm subst0_I[OF qfmq] set_upto) auto | 
| 55981 | 2380 | also have "\<dots> \<longleftrightarrow> ?I i (evaldjf (\<lambda>j. simpfm (subst0 (C j) ?smq)) ?js) \<or> | 
| 2381 | (\<exists>j\<in> set ?js. \<exists>b\<in> set ?B. ?I i (subst0 (Add b (C j)) ?q))" | |
| 50313 | 2382 | by (simp only: evaldjf_ex subst0_I[OF qfq]) | 
| 55981 | 2383 | also have "\<dots> \<longleftrightarrow> ?I i ?md \<or> | 
| 2384 | (\<exists>(b,j) \<in> set ?Bjs. (\<lambda>(b,j). ?I i (simpfm (subst0 (Add b (C j)) ?q))) (b,j))" | |
| 50313 | 2385 | by (simp only: simpfm set_concat set_map concat_map_singleton UN_simps) blast | 
| 55981 | 2386 | also have "\<dots> \<longleftrightarrow> ?I i ?md \<or> ?I i (evaldjf (\<lambda>(b,j). simpfm (subst0 (Add b (C j)) ?q)) ?Bjs)" | 
| 50313 | 2387 | by (simp only: evaldjf_ex[where bs="i#bs" and f="\<lambda>(b,j). simpfm (subst0 (Add b (C j)) ?q)" and ps="?Bjs"]) | 
| 2388 | (auto simp add: split_def) | |
| 55981 | 2389 | finally have mdqd: "?lhs \<longleftrightarrow> ?I i ?md \<or> ?I i ?qd" | 
| 55921 | 2390 | by simp | 
| 55981 | 2391 | also have "\<dots> \<longleftrightarrow> ?I i (disj ?md ?qd)" | 
| 55921 | 2392 | by (simp add: disj) | 
| 55981 | 2393 | also have "\<dots> \<longleftrightarrow> Ifm bbs bs (decr (disj ?md ?qd))" | 
| 55921 | 2394 | by (simp only: decr [OF mdqdb]) | 
| 55981 | 2395 | finally have mdqd2: "?lhs \<longleftrightarrow> Ifm bbs bs (decr (disj ?md ?qd))" . | 
| 60708 | 2396 | show ?thesis | 
| 2397 | proof (cases "?md = T") | |
| 2398 | case True | |
| 55921 | 2399 | then have cT: "cooper p = T" | 
| 23274 | 2400 | by (simp only: cooper_def unit_def split_def Let_def if_True) simp | 
| 60708 | 2401 | from True have lhs: "?lhs" | 
| 55921 | 2402 | using mdqd by simp | 
| 60708 | 2403 | from True have "?rhs" | 
| 55921 | 2404 | by (simp add: cooper_def unit_def split_def) | 
| 60708 | 2405 | with lhs cT show ?thesis | 
| 55981 | 2406 | by simp | 
| 60708 | 2407 | next | 
| 2408 | case False | |
| 55921 | 2409 | then have "cooper p = decr (disj ?md ?qd)" | 
| 50313 | 2410 | by (simp only: cooper_def unit_def split_def Let_def if_False) | 
| 60708 | 2411 | with mdqd2 decr_qf[OF mdqdb] show ?thesis | 
| 55921 | 2412 | by simp | 
| 60708 | 2413 | 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 | 2414 | 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 | 2415 | |
| 55921 | 2416 | definition pa :: "fm \<Rightarrow> fm" | 
| 2417 | where "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 | 2418 | |
| 55921 | 2419 | theorem mirqe: "Ifm bbs bs (pa p) = Ifm bbs bs p \<and> qfree (pa p)" | 
| 23274 | 2420 | 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 | 2421 | |
| 50313 | 2422 | definition cooper_test :: "unit \<Rightarrow> fm" | 
| 2423 | where | |
| 2424 | "cooper_test u = | |
| 2425 | pa (E (A (Imp (Ge (Sub (Bound 0) (Bound 1))) | |
| 2426 | (E (E (Eq (Sub (Add (Mul 3 (Bound 1)) (Mul 5 (Bound 0))) (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 | 2427 | |
| 60533 | 2428 | ML_val \<open>@{code cooper_test} ()\<close>
 | 
| 27456 | 2429 | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2430 | (*code_reflect Cooper_Procedure | 
| 55685 | 2431 | functions pa T Bound nat_of_integer integer_of_nat int_of_integer integer_of_int | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2432 | file "~~/src/HOL/Tools/Qelim/cooper_procedure.ML"*) | 
| 27456 | 2433 | |
| 60533 | 2434 | oracle linzqe_oracle = \<open> | 
| 27456 | 2435 | let | 
| 2436 | ||
| 55814 | 2437 | fun num_of_term vs (t as Free (xn, xT)) = | 
| 2438 | (case AList.lookup (op =) vs t of | |
| 2439 | NONE => error "Variable not found in the list!" | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2440 |       | SOME n => @{code Bound} (@{code nat_of_integer} n))
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2441 |   | num_of_term vs @{term "0::int"} = @{code C} (@{code int_of_integer} 0)
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2442 |   | num_of_term vs @{term "1::int"} = @{code C} (@{code int_of_integer} 1)
 | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 2443 |   | num_of_term vs @{term "- 1::int"} = @{code C} (@{code int_of_integer} (~ 1))
 | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2444 |   | num_of_term vs (@{term "numeral :: _ \<Rightarrow> int"} $ t) =
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2445 |       @{code C} (@{code int_of_integer} (HOLogic.dest_num t))
 | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
54230diff
changeset | 2446 |   | num_of_term vs (@{term "- numeral :: _ \<Rightarrow> int"} $ t) =
 | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2447 |       @{code C} (@{code int_of_integer} (~(HOLogic.dest_num t)))
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2448 |   | num_of_term vs (Bound i) = @{code Bound} (@{code nat_of_integer} i)
 | 
| 27456 | 2449 |   | num_of_term vs (@{term "uminus :: int \<Rightarrow> int"} $ t') = @{code Neg} (num_of_term vs t')
 | 
| 2450 |   | num_of_term vs (@{term "op + :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
 | |
| 2451 |       @{code Add} (num_of_term vs t1, num_of_term vs t2)
 | |
| 2452 |   | num_of_term vs (@{term "op - :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
 | |
| 2453 |       @{code Sub} (num_of_term vs t1, num_of_term vs t2)
 | |
| 2454 |   | num_of_term vs (@{term "op * :: int \<Rightarrow> int \<Rightarrow> int"} $ t1 $ t2) =
 | |
| 55814 | 2455 | (case try HOLogic.dest_number t1 of | 
| 2456 |         SOME (_, i) => @{code Mul} (@{code int_of_integer} i, num_of_term vs t2)
 | |
| 2457 | | NONE => | |
| 2458 | (case try HOLogic.dest_number t2 of | |
| 2459 |             SOME (_, i) => @{code Mul} (@{code int_of_integer} i, num_of_term vs t1)
 | |
| 2460 | | NONE => error "num_of_term: unsupported multiplication")) | |
| 28264 | 2461 |   | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t);
 | 
| 27456 | 2462 | |
| 2463 | fun fm_of_term ps vs @{term True} = @{code T}
 | |
| 2464 |   | fm_of_term ps vs @{term False} = @{code F}
 | |
| 2465 |   | fm_of_term ps vs (@{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
 | |
| 2466 |       @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
 | |
| 2467 |   | fm_of_term ps vs (@{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
 | |
| 2468 |       @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
 | |
| 2469 |   | fm_of_term ps vs (@{term "op = :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
 | |
| 50313 | 2470 |       @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
 | 
| 27456 | 2471 |   | fm_of_term ps vs (@{term "op dvd :: int \<Rightarrow> int \<Rightarrow> bool"} $ t1 $ t2) =
 | 
| 55814 | 2472 | (case try HOLogic.dest_number t1 of | 
| 2473 |         SOME (_, i) => @{code Dvd} (@{code int_of_integer} i, num_of_term vs t2)
 | |
| 2474 | | NONE => error "num_of_term: unsupported dvd") | |
| 27456 | 2475 |   | fm_of_term ps vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) =
 | 
| 2476 |       @{code Iff} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | |
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 2477 |   | fm_of_term ps vs (@{term HOL.conj} $ t1 $ t2) =
 | 
| 27456 | 2478 |       @{code And} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 2479 |   | fm_of_term ps vs (@{term HOL.disj} $ t1 $ t2) =
 | 
| 27456 | 2480 |       @{code Or} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 2481 |   | fm_of_term ps vs (@{term HOL.implies} $ t1 $ t2) =
 | 
| 27456 | 2482 |       @{code Imp} (fm_of_term ps vs t1, fm_of_term ps vs t2)
 | 
| 2483 |   | fm_of_term ps vs (@{term "Not"} $ t') =
 | |
| 2484 |       @{code NOT} (fm_of_term ps vs t')
 | |
| 38558 | 2485 |   | fm_of_term ps vs (Const (@{const_name Ex}, _) $ Abs (xn, xT, p)) =
 | 
| 27456 | 2486 | let | 
| 42284 | 2487 | val (xn', p') = Syntax_Trans.variant_abs (xn, xT, p); (* FIXME !? *) | 
| 27456 | 2488 | val vs' = (Free (xn', xT), 0) :: map (fn (v, n) => (v, n + 1)) vs; | 
| 2489 |       in @{code E} (fm_of_term ps vs' p) end
 | |
| 38558 | 2490 |   | fm_of_term ps vs (Const (@{const_name All}, _) $ Abs (xn, xT, p)) =
 | 
| 27456 | 2491 | let | 
| 42284 | 2492 | val (xn', p') = Syntax_Trans.variant_abs (xn, xT, p); (* FIXME !? *) | 
| 27456 | 2493 | val vs' = (Free (xn', xT), 0) :: map (fn (v, n) => (v, n + 1)) vs; | 
| 2494 |       in @{code A} (fm_of_term ps vs' p) end
 | |
| 28264 | 2495 |   | fm_of_term ps vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t);
 | 
| 23515 | 2496 | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2497 | fun term_of_num vs (@{code C} i) = HOLogic.mk_number HOLogic.intT (@{code integer_of_int} i)
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2498 |   | term_of_num vs (@{code Bound} n) =
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2499 | let | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2500 |         val q = @{code integer_of_nat} n
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2501 | in fst (the (find_first (fn (_, m) => q = m) vs)) end | 
| 27456 | 2502 |   | term_of_num vs (@{code Neg} t') = @{term "uminus :: int \<Rightarrow> int"} $ term_of_num vs t'
 | 
| 2503 |   | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: int \<Rightarrow> int \<Rightarrow> int"} $
 | |
| 2504 | term_of_num vs t1 $ term_of_num vs t2 | |
| 2505 |   | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: int \<Rightarrow> int \<Rightarrow> int"} $
 | |
| 2506 | term_of_num vs t1 $ term_of_num vs t2 | |
| 2507 |   | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: int \<Rightarrow> int \<Rightarrow> int"} $
 | |
| 2508 |       term_of_num vs (@{code C} i) $ term_of_num vs t2
 | |
| 55814 | 2509 |   | term_of_num vs (@{code CN} (n, i, t)) =
 | 
| 2510 |       term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t));
 | |
| 27456 | 2511 | |
| 50313 | 2512 | fun term_of_fm ps vs @{code T} = @{term True}
 | 
| 45740 | 2513 |   | term_of_fm ps vs @{code F} = @{term False}
 | 
| 27456 | 2514 |   | term_of_fm ps vs (@{code Lt} t) =
 | 
| 2515 |       @{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::int"}
 | |
| 2516 |   | term_of_fm ps vs (@{code Le} t) =
 | |
| 2517 |       @{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::int"}
 | |
| 2518 |   | term_of_fm ps vs (@{code Gt} t) =
 | |
| 2519 |       @{term "op < :: int \<Rightarrow> int \<Rightarrow> bool"} $ @{term "0::int"} $ term_of_num vs t
 | |
| 2520 |   | term_of_fm ps vs (@{code Ge} t) =
 | |
| 2521 |       @{term "op \<le> :: int \<Rightarrow> int \<Rightarrow> bool"} $ @{term "0::int"} $ term_of_num vs t
 | |
| 2522 |   | term_of_fm ps vs (@{code Eq} t) =
 | |
| 2523 |       @{term "op = :: int \<Rightarrow> int \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::int"}
 | |
| 2524 |   | term_of_fm ps vs (@{code NEq} t) =
 | |
| 2525 |       term_of_fm ps vs (@{code NOT} (@{code Eq} t))
 | |
| 2526 |   | term_of_fm ps vs (@{code Dvd} (i, t)) =
 | |
| 2527 |       @{term "op dvd :: int \<Rightarrow> int \<Rightarrow> bool"} $ term_of_num vs (@{code C} i) $ term_of_num vs t
 | |
| 2528 |   | term_of_fm ps vs (@{code NDvd} (i, t)) =
 | |
| 2529 |       term_of_fm ps vs (@{code NOT} (@{code Dvd} (i, t)))
 | |
| 2530 |   | term_of_fm ps vs (@{code NOT} t') =
 | |
| 2531 | HOLogic.Not $ term_of_fm ps vs t' | |
| 2532 |   | term_of_fm ps vs (@{code And} (t1, t2)) =
 | |
| 2533 | HOLogic.conj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | |
| 2534 |   | term_of_fm ps vs (@{code Or} (t1, t2)) =
 | |
| 2535 | HOLogic.disj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | |
| 2536 |   | term_of_fm ps vs (@{code Imp} (t1, t2)) =
 | |
| 2537 | HOLogic.imp $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | |
| 2538 |   | term_of_fm ps vs (@{code Iff} (t1, t2)) =
 | |
| 2539 |       @{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ term_of_fm ps vs t1 $ term_of_fm ps vs t2
 | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2540 |   | term_of_fm ps vs (@{code Closed} n) =
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2541 | let | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2542 |         val q = @{code integer_of_nat} n
 | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
50313diff
changeset | 2543 | in (fst o the) (find_first (fn (_, m) => m = q) ps) end | 
| 29788 | 2544 |   | term_of_fm ps vs (@{code NClosed} n) = term_of_fm ps vs (@{code NOT} (@{code Closed} n));
 | 
| 27456 | 2545 | |
| 2546 | fun term_bools acc t = | |
| 2547 | let | |
| 55814 | 2548 | val is_op = | 
| 2549 |       member (op =) [@{term HOL.conj}, @{term HOL.disj}, @{term HOL.implies},
 | |
| 2550 |       @{term "op = :: bool => _"},
 | |
| 27456 | 2551 |       @{term "op = :: int => _"}, @{term "op < :: int => _"},
 | 
| 2552 |       @{term "op <= :: int => _"}, @{term "Not"}, @{term "All :: (int => _) => _"},
 | |
| 2553 |       @{term "Ex :: (int => _) => _"}, @{term "True"}, @{term "False"}]
 | |
| 50313 | 2554 | fun is_ty t = not (fastype_of t = HOLogic.boolT) | 
| 55814 | 2555 | in | 
| 2556 | (case t of | |
| 2557 | (l as f $ a) $ b => | |
| 2558 | if is_ty t orelse is_op t then term_bools (term_bools acc l) b | |
| 27456 | 2559 | else insert (op aconv) t acc | 
| 55814 | 2560 | | f $ a => | 
| 2561 | if is_ty t orelse is_op t then term_bools (term_bools acc f) a | |
| 27456 | 2562 | else insert (op aconv) t acc | 
| 42284 | 2563 | | Abs p => term_bools acc (snd (Syntax_Trans.variant_abs p)) (* FIXME !? *) | 
| 55814 | 2564 | | _ => if is_ty t orelse is_op t then acc else insert (op aconv) t acc) | 
| 27456 | 2565 | end; | 
| 2566 | ||
| 55814 | 2567 | in | 
| 60325 | 2568 | fn (ctxt, t) => | 
| 55814 | 2569 | let | 
| 2570 | val fs = Misc_Legacy.term_frees t; | |
| 2571 | val bs = term_bools [] t; | |
| 2572 | val vs = map_index swap fs; | |
| 2573 | val ps = map_index swap bs; | |
| 60325 | 2574 |       val t' = term_of_fm ps vs (@{code pa} (fm_of_term ps vs t));
 | 
| 2575 | in Thm.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, t'))) end | |
| 27456 | 2576 | end; | 
| 60533 | 2577 | \<close> | 
| 27456 | 2578 | |
| 48891 | 2579 | ML_file "cooper_tac.ML" | 
| 47432 | 2580 | |
| 60533 | 2581 | method_setup cooper = \<open> | 
| 53168 | 2582 | Scan.lift (Args.mode "no_quantify") >> | 
| 47432 | 2583 | (fn q => fn ctxt => SIMPLE_METHOD' (Cooper_Tac.linz_tac ctxt (not q))) | 
| 60533 | 2584 | \<close> "decision procedure for linear integer arithmetic" | 
| 47432 | 2585 | |
| 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 | 2586 | |
| 60533 | 2587 | text \<open>Tests\<close> | 
| 27456 | 2588 | |
| 55814 | 2589 | lemma "\<exists>(j::int). \<forall>x\<ge>j. \<exists>a b. x = 3*a+5*b" | 
| 27456 | 2590 | 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 | 2591 | |
| 55814 | 2592 | lemma "\<forall>(x::int) \<ge> 8. \<exists>i j. 5*i + 3*j = x" | 
| 27456 | 2593 | by cooper | 
| 2594 | ||
| 55814 | 2595 | theorem "(\<forall>(y::int). 3 dvd y) \<Longrightarrow> \<forall>(x::int). b < x \<longrightarrow> a \<le> x" | 
| 23274 | 2596 | 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 | 2597 | |
| 55814 | 2598 | theorem "\<And>(y::int) (z::int) (n::int). 3 dvd z \<Longrightarrow> 2 dvd (y::int) \<Longrightarrow> | 
| 2599 | (\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)" | |
| 23274 | 2600 | by cooper | 
| 2601 | ||
| 55814 | 2602 | theorem "\<And>(y::int) (z::int) n. Suc n < 6 \<Longrightarrow> 3 dvd z \<Longrightarrow> | 
| 2603 | 2 dvd (y::int) \<Longrightarrow> (\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)" | |
| 23274 | 2604 | by cooper | 
| 2605 | ||
| 55814 | 2606 | theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 \<longrightarrow> y = 5 + x" | 
| 23274 | 2607 | 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 | 2608 | |
| 55814 | 2609 | lemma "\<forall>(x::int) \<ge> 8. \<exists>i j. 5*i + 3*j = x" | 
| 50313 | 2610 | by cooper | 
| 27456 | 2611 | |
| 55814 | 2612 | lemma "\<forall>(y::int) (z::int) (n::int). | 
| 2613 | 3 dvd z \<longrightarrow> 2 dvd (y::int) \<longrightarrow> (\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)" | |
| 27456 | 2614 | by cooper | 
| 2615 | ||
| 55814 | 2616 | lemma "\<forall>(x::int) y. x < y \<longrightarrow> 2 * x + 1 < 2 * y" | 
| 27456 | 2617 | by cooper | 
| 2618 | ||
| 55814 | 2619 | lemma "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y" | 
| 27456 | 2620 | by cooper | 
| 2621 | ||
| 55814 | 2622 | lemma "\<exists>(x::int) y. 0 < x \<and> 0 \<le> y \<and> 3 * x - 5 * y = 1" | 
| 27456 | 2623 | by cooper | 
| 2624 | ||
| 55814 | 2625 | lemma "\<not> (\<exists>(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)" | 
| 27456 | 2626 | by cooper | 
| 2627 | ||
| 55814 | 2628 | lemma "\<forall>(x::int). 2 dvd x \<longrightarrow> (\<exists>(y::int). x = 2*y)" | 
| 27456 | 2629 | by cooper | 
| 2630 | ||
| 55814 | 2631 | lemma "\<forall>(x::int). 2 dvd x \<longleftrightarrow> (\<exists>(y::int). x = 2*y)" | 
| 27456 | 2632 | by cooper | 
| 2633 | ||
| 55814 | 2634 | lemma "\<forall>(x::int). 2 dvd x \<longleftrightarrow> (\<forall>(y::int). x \<noteq> 2*y + 1)" | 
| 27456 | 2635 | by cooper | 
| 2636 | ||
| 55814 | 2637 | lemma "\<not> (\<forall>(x::int). | 
| 55921 | 2638 | (2 dvd x \<longleftrightarrow> (\<forall>(y::int). x \<noteq> 2*y+1) \<or> | 
| 55814 | 2639 | (\<exists>(q::int) (u::int) i. 3*i + 2*q - u < 17) \<longrightarrow> 0 < x \<or> (\<not> 3 dvd x \<and> x + 8 = 0)))" | 
| 23274 | 2640 | by cooper | 
| 27456 | 2641 | |
| 55814 | 2642 | lemma "\<not> (\<forall>(i::int). 4 \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i))" | 
| 27456 | 2643 | 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 | 2644 | |
| 55814 | 2645 | lemma "\<exists>j. \<forall>(x::int) \<ge> j. \<exists>i j. 5*i + 3*j = x" | 
| 2646 | by cooper | |
| 2647 | ||
| 2648 | theorem "(\<forall>(y::int). 3 dvd y) \<Longrightarrow> \<forall>(x::int). b < x \<longrightarrow> a \<le> x" | |
| 23274 | 2649 | 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 | 2650 | |
| 55814 | 2651 | theorem "\<And>(y::int) (z::int) (n::int). 3 dvd z \<Longrightarrow> 2 dvd (y::int) \<Longrightarrow> | 
| 2652 | (\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)" | |
| 23274 | 2653 | 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 | 2654 | |
| 55814 | 2655 | theorem "\<And>(y::int) (z::int) n. Suc n < 6 \<Longrightarrow> 3 dvd z \<Longrightarrow> | 
| 2656 | 2 dvd (y::int) \<Longrightarrow> (\<exists>(x::int). 2*x = y) \<and> (\<exists>(k::int). 3*k = z)" | |
| 23274 | 2657 | 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 | 2658 | |
| 55814 | 2659 | theorem "\<forall>(x::nat). \<exists>(y::nat). (0::nat) \<le> 5 \<longrightarrow> y = 5 + x" | 
| 23274 | 2660 | 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 | 2661 | |
| 55814 | 2662 | theorem "\<forall>(x::nat). \<exists>(y::nat). y = 5 + x \<or> x div 6 + 1 = 2" | 
| 23274 | 2663 | 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 | 2664 | |
| 23274 | 2665 | theorem "\<exists>(x::int). 0 < x" | 
| 2666 | 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 | 2667 | |
| 55814 | 2668 | theorem "\<forall>(x::int) y. x < y \<longrightarrow> 2 * x + 1 < 2 * y" | 
| 23274 | 2669 | by cooper | 
| 50313 | 2670 | |
| 23274 | 2671 | theorem "\<forall>(x::int) y. 2 * x + 1 \<noteq> 2 * y" | 
| 2672 | by cooper | |
| 50313 | 2673 | |
| 23274 | 2674 | theorem "\<exists>(x::int) y. 0 < x & 0 \<le> y & 3 * x - 5 * y = 1" | 
| 2675 | 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 | 2676 | |
| 55814 | 2677 | theorem "\<not> (\<exists>(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)" | 
| 23274 | 2678 | 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 | 2679 | |
| 55814 | 2680 | theorem "\<not> (\<exists>(x::int). False)" | 
| 23274 | 2681 | 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 | 2682 | |
| 55814 | 2683 | theorem "\<forall>(x::int). 2 dvd x \<longrightarrow> (\<exists>(y::int). x = 2*y)" | 
| 50313 | 2684 | by cooper | 
| 23274 | 2685 | |
| 55814 | 2686 | theorem "\<forall>(x::int). 2 dvd x \<longrightarrow> (\<exists>(y::int). x = 2*y)" | 
| 50313 | 2687 | 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 | 2688 | |
| 55814 | 2689 | theorem "\<forall>(x::int). 2 dvd x \<longleftrightarrow> (\<exists>(y::int). x = 2*y)" | 
| 50313 | 2690 | 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 | 2691 | |
| 55814 | 2692 | theorem "\<forall>(x::int). 2 dvd x \<longleftrightarrow> (\<forall>(y::int). x \<noteq> 2*y + 1)" | 
| 50313 | 2693 | 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 | 2694 | |
| 55814 | 2695 | theorem | 
| 2696 | "\<not> (\<forall>(x::int). | |
| 2697 | (2 dvd x \<longleftrightarrow> | |
| 2698 | (\<forall>(y::int). x \<noteq> 2*y+1) \<or> | |
| 2699 | (\<exists>(q::int) (u::int) i. 3*i + 2*q - u < 17) | |
| 2700 | \<longrightarrow> 0 < x \<or> (\<not> 3 dvd x \<and> x + 8 = 0)))" | |
| 23274 | 2701 | by cooper | 
| 50313 | 2702 | |
| 55814 | 2703 | theorem "\<not> (\<forall>(i::int). 4 \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i))" | 
| 23274 | 2704 | 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 | 2705 | |
| 55814 | 2706 | theorem "\<forall>(i::int). 8 \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i)" | 
| 23274 | 2707 | 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 | 2708 | |
| 55814 | 2709 | theorem "\<exists>(j::int). \<forall>i. j \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i)" | 
| 23274 | 2710 | 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 | 2711 | |
| 55814 | 2712 | theorem "\<not> (\<forall>j (i::int). j \<le> i \<longrightarrow> (\<exists>x y. 0 \<le> x \<and> 0 \<le> y \<and> 3 * x + 5 * y = i))" | 
| 23274 | 2713 | 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 | 2714 | |
| 55814 | 2715 | theorem "(\<exists>m::nat. n = 2 * m) \<longrightarrow> (n + 1) div 2 = n div 2" | 
| 23274 | 2716 | by cooper | 
| 17388 | 2717 | |
| 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 | 2718 | end |