1 (* Title: HOL/Decision_Procs/MIR.thy
6 imports Complex_Main Dense_Linear_Order DP_Library
7 "~~/src/HOL/Library/Efficient_Nat" "~~/src/HOL/Library/Old_Recdef"
10 section {* Quantifier elimination for @{text "\<real> (0, 1, +, floor, <)"} *}
12 declare real_of_int_floor_cancel [simp del]
14 lemma myle: fixes a b :: "'a::{ordered_ab_group_add}"
15 shows "(a \<le> b) = (0 \<le> b - a)"
16 by (metis add_0_left add_le_cancel_right diff_add_cancel)
18 lemma myless: fixes a b :: "'a::{ordered_ab_group_add}"
19 shows "(a < b) = (0 < b - a)"
20 by (metis le_iff_diff_le_0 less_le_not_le myle)
22 (* Maybe should be added to the library \<dots> *)
23 lemma floor_int_eq: "(real n\<le> x \<and> x < real (n+1)) = (floor x = n)"
25 assume lb: "real n \<le> x"
26 and ub: "x < real n + 1"
27 have "real (floor x) \<le> x" by simp
28 hence "real (floor x) < real (n + 1) " using ub by arith
29 hence "floor x < n+1" by simp
30 moreover from lb have "n \<le> floor x" using floor_mono[where x="real n" and y="x"]
31 by simp ultimately show "floor x = n" by simp
34 (* Periodicity of dvd *)
36 assumes advdd: "(a::int) dvd d"
37 shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))"
41 from inf_period(3)[OF advdd, rule_format, where x=x and k="-k"]
42 have " ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp}
43 hence "\<forall>x.\<forall>k. ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp
44 then show ?thesis by simp
47 (* The Divisibility relation between reals *)
49 rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50)
51 rdvd_def: "x rdvd y \<longleftrightarrow> (\<exists>k\<Colon>int. y = x * real k)"
54 shows "real (i::int) rdvd x = (i dvd (floor x) \<and> real (floor x) = x)" (is "?l = ?r")
57 hence th: "\<exists> k. x=real (i*k)" by (simp add: rdvd_def)
58 hence th': "real (floor x) = x" by (auto simp del: real_of_int_mult)
59 with th have "\<exists> k. real (floor x) = real (i*k)" by simp
60 hence "\<exists> k. floor x = i*k" by (simp only: real_of_int_inject)
61 thus ?r using th' by (simp add: dvd_def)
63 assume "?r" hence "(i\<Colon>int) dvd \<lfloor>x\<Colon>real\<rfloor>" ..
64 hence "\<exists> k. real (floor x) = real (i*k)"
65 by (simp only: real_of_int_inject) (simp add: dvd_def)
66 thus ?l using `?r` by (simp add: rdvd_def)
69 lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)"
70 by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only :real_of_int_mult[symmetric])
74 "(abs (real d) rdvd t) = (real (d ::int) rdvd t)"
76 assume d: "real d rdvd t"
77 from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real (floor t) = t" by auto
79 from iffD2[OF abs_dvd_iff] d2 have "(abs d) dvd (floor t)" by blast
80 with ti int_rdvd_real[symmetric] have "real (abs d) rdvd t" by blast
81 thus "abs (real d) rdvd t" by simp
83 assume "abs (real d) rdvd t" hence "real (abs d) rdvd t" by simp
84 with int_rdvd_real[where i="abs d" and x="t"] have d2: "abs d dvd floor t" and ti: "real (floor t) =t" by auto
85 from iffD1[OF abs_dvd_iff] d2 have "d dvd floor t" by blast
86 with ti int_rdvd_real[symmetric] show "real d rdvd t" by blast
89 lemma rdvd_minus: "(real (d::int) rdvd t) = (real d rdvd -t)"
90 apply (auto simp add: rdvd_def)
91 apply (rule_tac x="-k" in exI, simp)
92 apply (rule_tac x="-k" in exI, simp)
95 lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)"
96 by (auto simp add: rdvd_def)
99 assumes knz: "k\<noteq>0"
100 shows "(real (n::int) * real (k::int) rdvd x * real k) = (real n rdvd x)"
101 using knz by (simp add:rdvd_def)
103 (*********************************************************************************)
104 (**** SHADOW SYNTAX AND SEMANTICS ****)
105 (*********************************************************************************)
107 datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num
108 | Mul int num | Floor num| CF int num num
110 (* A size for num to make inductive proofs simpler*)
111 primrec num_size :: "num \<Rightarrow> nat" where
113 | "num_size (Bound n) = 1"
114 | "num_size (Neg a) = 1 + num_size a"
115 | "num_size (Add a b) = 1 + num_size a + num_size b"
116 | "num_size (Sub a b) = 3 + num_size a + num_size b"
117 | "num_size (CN n c a) = 4 + num_size a "
118 | "num_size (CF c a b) = 4 + num_size a + num_size b"
119 | "num_size (Mul c a) = 1 + num_size a"
120 | "num_size (Floor a) = 1 + num_size a"
122 (* Semantics of numeral terms (num) *)
123 primrec Inum :: "real list \<Rightarrow> num \<Rightarrow> real" where
124 "Inum bs (C c) = (real c)"
125 | "Inum bs (Bound n) = bs!n"
126 | "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)"
127 | "Inum bs (Neg a) = -(Inum bs a)"
128 | "Inum bs (Add a b) = Inum bs a + Inum bs b"
129 | "Inum bs (Sub a b) = Inum bs a - Inum bs b"
130 | "Inum bs (Mul c a) = (real c) * Inum bs a"
131 | "Inum bs (Floor a) = real (floor (Inum bs a))"
132 | "Inum bs (CF c a b) = real c * real (floor (Inum bs a)) + Inum bs b"
133 definition "isint t bs \<equiv> real (floor (Inum bs t)) = Inum bs t"
135 lemma isint_iff: "isint n bs = (real (floor (Inum bs n)) = Inum bs n)"
136 by (simp add: isint_def)
138 lemma isint_Floor: "isint (Floor n) bs"
139 by (simp add: isint_iff)
141 lemma isint_Mul: "isint e bs \<Longrightarrow> isint (Mul c e) bs"
145 assume be: "isint e bs" hence efe:"real ?fe = ?e" by (simp add: isint_iff)
146 have "real ((floor (Inum bs (Mul c e)))) = real (floor (real (c * ?fe)))" using efe by simp
147 also have "\<dots> = real (c* ?fe)" by (simp only: floor_real_of_int)
148 also have "\<dots> = real c * ?e" using efe by simp
149 finally show ?thesis using isint_iff by simp
152 lemma isint_neg: "isint e bs \<Longrightarrow> isint (Neg e) bs"
154 let ?I = "\<lambda> t. Inum bs t"
155 assume ie: "isint e bs"
156 hence th: "real (floor (?I e)) = ?I e" by (simp add: isint_def)
157 have "real (floor (?I (Neg e))) = real (floor (- (real (floor (?I e)))))" by (simp add: th)
158 also have "\<dots> = - real (floor (?I e))" by(simp add: floor_minus_real_of_int)
159 finally show "isint (Neg e) bs" by (simp add: isint_def th)
163 assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs"
165 let ?I = "\<lambda> t. Inum bs t"
166 from ie have th: "real (floor (?I e)) = ?I e" by (simp add: isint_def)
167 have "real (floor (?I (Sub (C c) e))) = real (floor ((real (c -floor (?I e)))))" by (simp add: th)
168 also have "\<dots> = real (c- floor (?I e))" by(simp add: floor_minus_real_of_int)
169 finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th)
172 lemma isint_add: assumes
173 ai:"isint a bs" and bi: "isint b bs" shows "isint (Add a b) bs"
177 from ai bi isint_iff have "real (floor (?a + ?b)) = real (floor (real (floor ?a) + real (floor ?b)))" by simp
178 also have "\<dots> = real (floor ?a) + real (floor ?b)" by simp
179 also have "\<dots> = ?a + ?b" using ai bi isint_iff by simp
180 finally show "isint (Add a b) bs" by (simp add: isint_iff)
183 lemma isint_c: "isint (C j) bs"
184 by (simp add: isint_iff)
189 T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num|
190 NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
194 fun fmsize :: "fm \<Rightarrow> nat" where
195 "fmsize (NOT p) = 1 + fmsize p"
196 | "fmsize (And p q) = 1 + fmsize p + fmsize q"
197 | "fmsize (Or p q) = 1 + fmsize p + fmsize q"
198 | "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
199 | "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
200 | "fmsize (E p) = 1 + fmsize p"
201 | "fmsize (A p) = 4+ fmsize p"
202 | "fmsize (Dvd i t) = 2"
203 | "fmsize (NDvd i t) = 2"
205 (* several lemmas about fmsize *)
206 lemma fmsize_pos: "fmsize p > 0"
207 by (induct p rule: fmsize.induct) simp_all
209 (* Semantics of formulae (fm) *)
210 primrec Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool" where
213 | "Ifm bs (Lt a) = (Inum bs a < 0)"
214 | "Ifm bs (Gt a) = (Inum bs a > 0)"
215 | "Ifm bs (Le a) = (Inum bs a \<le> 0)"
216 | "Ifm bs (Ge a) = (Inum bs a \<ge> 0)"
217 | "Ifm bs (Eq a) = (Inum bs a = 0)"
218 | "Ifm bs (NEq a) = (Inum bs a \<noteq> 0)"
219 | "Ifm bs (Dvd i b) = (real i rdvd Inum bs b)"
220 | "Ifm bs (NDvd i b) = (\<not>(real i rdvd Inum bs b))"
221 | "Ifm bs (NOT p) = (\<not> (Ifm bs p))"
222 | "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)"
223 | "Ifm bs (Or p q) = (Ifm bs p \<or> Ifm bs q)"
224 | "Ifm bs (Imp p q) = ((Ifm bs p) \<longrightarrow> (Ifm bs q))"
225 | "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)"
226 | "Ifm bs (E p) = (\<exists> x. Ifm (x#bs) p)"
227 | "Ifm bs (A p) = (\<forall> x. Ifm (x#bs) p)"
229 consts prep :: "fm \<Rightarrow> fm"
230 recdef prep "measure fmsize"
233 "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))"
234 "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))"
235 "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
236 "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))"
237 "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
238 "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
239 "prep (E p) = E (prep p)"
240 "prep (A (And p q)) = And (prep (A p)) (prep (A q))"
241 "prep (A p) = prep (NOT (E (NOT p)))"
242 "prep (NOT (NOT p)) = prep p"
243 "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))"
244 "prep (NOT (A p)) = prep (E (NOT p))"
245 "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))"
246 "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))"
247 "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))"
248 "prep (NOT p) = NOT (prep p)"
249 "prep (Or p q) = Or (prep p) (prep q)"
250 "prep (And p q) = And (prep p) (prep q)"
251 "prep (Imp p q) = prep (Or (NOT p) q)"
252 "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))"
254 (hints simp add: fmsize_pos)
255 lemma prep: "\<And> bs. Ifm bs (prep p) = Ifm bs p"
256 by (induct p rule: prep.induct, auto)
259 (* Quantifier freeness *)
260 fun qfree:: "fm \<Rightarrow> bool" where
261 "qfree (E p) = False"
262 | "qfree (A p) = False"
263 | "qfree (NOT p) = qfree p"
264 | "qfree (And p q) = (qfree p \<and> qfree q)"
265 | "qfree (Or p q) = (qfree p \<and> qfree q)"
266 | "qfree (Imp p q) = (qfree p \<and> qfree q)"
267 | "qfree (Iff p q) = (qfree p \<and> qfree q)"
270 (* Boundedness and substitution *)
271 primrec numbound0 :: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *) where
272 "numbound0 (C c) = True"
273 | "numbound0 (Bound n) = (n>0)"
274 | "numbound0 (CN n i a) = (n > 0 \<and> numbound0 a)"
275 | "numbound0 (Neg a) = numbound0 a"
276 | "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)"
277 | "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)"
278 | "numbound0 (Mul i a) = numbound0 a"
279 | "numbound0 (Floor a) = numbound0 a"
280 | "numbound0 (CF c a b) = (numbound0 a \<and> numbound0 b)"
283 assumes nb: "numbound0 a"
284 shows "Inum (b#bs) a = Inum (b'#bs) a"
285 using nb by (induct a) auto
288 assumes nb: "numbound0 t" and ti: "isint t (x#bs)"
289 shows "\<forall> y. isint t (y#bs)"
293 from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def]
294 show "isint t (y#bs)"
295 by (simp add: isint_def)
298 primrec bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *) where
301 | "bound0 (Lt a) = numbound0 a"
302 | "bound0 (Le a) = numbound0 a"
303 | "bound0 (Gt a) = numbound0 a"
304 | "bound0 (Ge a) = numbound0 a"
305 | "bound0 (Eq a) = numbound0 a"
306 | "bound0 (NEq a) = numbound0 a"
307 | "bound0 (Dvd i a) = numbound0 a"
308 | "bound0 (NDvd i a) = numbound0 a"
309 | "bound0 (NOT p) = bound0 p"
310 | "bound0 (And p q) = (bound0 p \<and> bound0 q)"
311 | "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
312 | "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
313 | "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
314 | "bound0 (E p) = False"
315 | "bound0 (A p) = False"
318 assumes bp: "bound0 p"
319 shows "Ifm (b#bs) p = Ifm (b'#bs) p"
320 using bp numbound0_I [where b="b" and bs="bs" and b'="b'"]
323 primrec numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" (* substitute a num into a num for Bound 0 *) where
324 "numsubst0 t (C c) = (C c)"
325 | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)"
326 | "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))"
327 | "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)"
328 | "numsubst0 t (Neg a) = Neg (numsubst0 t a)"
329 | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)"
330 | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)"
331 | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)"
332 | "numsubst0 t (Floor a) = Floor (numsubst0 t a)"
335 shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t"
336 by (induct t) simp_all
338 primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *) where
341 | "subst0 t (Lt a) = Lt (numsubst0 t a)"
342 | "subst0 t (Le a) = Le (numsubst0 t a)"
343 | "subst0 t (Gt a) = Gt (numsubst0 t a)"
344 | "subst0 t (Ge a) = Ge (numsubst0 t a)"
345 | "subst0 t (Eq a) = Eq (numsubst0 t a)"
346 | "subst0 t (NEq a) = NEq (numsubst0 t a)"
347 | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)"
348 | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)"
349 | "subst0 t (NOT p) = NOT (subst0 t p)"
350 | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
351 | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
352 | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
353 | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
355 lemma subst0_I: assumes qfp: "qfree p"
356 shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p"
357 using qfp numsubst0_I[where b="b" and bs="bs" and a="a"]
358 by (induct p) simp_all
360 fun decrnum:: "num \<Rightarrow> num" where
361 "decrnum (Bound n) = Bound (n - 1)"
362 | "decrnum (Neg a) = Neg (decrnum a)"
363 | "decrnum (Add a b) = Add (decrnum a) (decrnum b)"
364 | "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)"
365 | "decrnum (Mul c a) = Mul c (decrnum a)"
366 | "decrnum (Floor a) = Floor (decrnum a)"
367 | "decrnum (CN n c a) = CN (n - 1) c (decrnum a)"
368 | "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)"
371 fun decr :: "fm \<Rightarrow> fm" where
372 "decr (Lt a) = Lt (decrnum a)"
373 | "decr (Le a) = Le (decrnum a)"
374 | "decr (Gt a) = Gt (decrnum a)"
375 | "decr (Ge a) = Ge (decrnum a)"
376 | "decr (Eq a) = Eq (decrnum a)"
377 | "decr (NEq a) = NEq (decrnum a)"
378 | "decr (Dvd i a) = Dvd i (decrnum a)"
379 | "decr (NDvd i a) = NDvd i (decrnum a)"
380 | "decr (NOT p) = NOT (decr p)"
381 | "decr (And p q) = And (decr p) (decr q)"
382 | "decr (Or p q) = Or (decr p) (decr q)"
383 | "decr (Imp p q) = Imp (decr p) (decr q)"
384 | "decr (Iff p q) = Iff (decr p) (decr q)"
387 lemma decrnum: assumes nb: "numbound0 t"
388 shows "Inum (x#bs) t = Inum bs (decrnum t)"
389 using nb by (induct t rule: decrnum.induct, simp_all)
391 lemma decr: assumes nb: "bound0 p"
392 shows "Ifm (x#bs) p = Ifm bs (decr p)"
394 by (induct p rule: decr.induct, simp_all add: decrnum)
396 lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)"
397 by (induct p, simp_all)
399 fun isatom :: "fm \<Rightarrow> bool" (* test for atomicity *) where
402 | "isatom (Lt a) = True"
403 | "isatom (Le a) = True"
404 | "isatom (Gt a) = True"
405 | "isatom (Ge a) = True"
406 | "isatom (Eq a) = True"
407 | "isatom (NEq a) = True"
408 | "isatom (Dvd i b) = True"
409 | "isatom (NDvd i b) = True"
412 lemma numsubst0_numbound0: assumes nb: "numbound0 t"
413 shows "numbound0 (numsubst0 t a)"
414 using nb by (induct a, auto)
416 lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t"
417 shows "bound0 (subst0 t p)"
418 using qf numsubst0_numbound0[OF nb] by (induct p, auto)
420 lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
421 by (induct p, simp_all)
424 definition djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
425 "djf f p q = (if q=T then T else if q=F then f p else
426 (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or fp q))"
428 definition evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
429 "evaldjf f ps = foldr (djf f) ps F"
431 lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)"
432 by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def)
433 (cases "f p", simp_all add: Let_def djf_def)
435 lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm bs (f p))"
436 by(induct ps, simp_all add: evaldjf_def djf_Or)
438 lemma evaldjf_bound0:
439 assumes nb: "\<forall> x\<in> set xs. bound0 (f x)"
440 shows "bound0 (evaldjf f xs)"
441 using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto)
444 assumes nb: "\<forall> x\<in> set xs. qfree (f x)"
445 shows "qfree (evaldjf f xs)"
446 using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto)
448 fun disjuncts :: "fm \<Rightarrow> fm list" where
449 "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)"
451 | "disjuncts p = [p]"
453 fun conjuncts :: "fm \<Rightarrow> fm list" where
454 "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)"
456 | "conjuncts p = [p]"
458 lemma conjuncts: "(\<forall> q\<in> set (conjuncts p). Ifm bs q) = Ifm bs p"
459 by(induct p rule: conjuncts.induct, auto)
461 lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). qfree q"
464 hence "list_all qfree (disjuncts p)"
465 by (induct p rule: disjuncts.induct, auto)
466 thus ?thesis by (simp only: list_all_iff)
468 lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). qfree q"
471 hence "list_all qfree (conjuncts p)"
472 by (induct p rule: conjuncts.induct, auto)
473 thus ?thesis by (simp only: list_all_iff)
476 definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
477 "DJ f p \<equiv> evaldjf f (disjuncts p)"
479 lemma DJ: assumes fdj: "\<forall> p q. f (Or p q) = Or (f p) (f q)"
481 shows "Ifm bs (DJ f p) = Ifm bs (f p)"
483 have "Ifm bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm bs (f q))"
484 by (simp add: DJ_def evaldjf_ex)
485 also have "\<dots> = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto)
486 finally show ?thesis .
490 fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)"
491 shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) "
493 fix p assume qf: "qfree p"
494 have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def)
495 from disjuncts_qf[OF qf] have "\<forall> q\<in> set (disjuncts p). qfree q" .
496 with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast
498 from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp
501 lemma DJ_qe: assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
502 shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm bs ((DJ qe p)) = Ifm bs (E p))"
506 from qe have qth: "\<forall> p. qfree p \<longrightarrow> qfree (qe p)" by blast
507 from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto
508 have "Ifm bs (DJ qe p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (qe q))"
509 by (simp add: DJ_def evaldjf_ex)
510 also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto
511 also have "\<dots> = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto)
512 finally show "qfree (DJ qe p) \<and> Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast
516 (* Algebraic simplifications for nums *)
517 fun bnds:: "num \<Rightarrow> nat list" where
518 "bnds (Bound n) = [n]"
519 | "bnds (CN n c a) = n#(bnds a)"
520 | "bnds (Neg a) = bnds a"
521 | "bnds (Add a b) = (bnds a)@(bnds b)"
522 | "bnds (Sub a b) = (bnds a)@(bnds b)"
523 | "bnds (Mul i a) = bnds a"
524 | "bnds (Floor a) = bnds a"
525 | "bnds (CF c a b) = (bnds a)@(bnds b)"
527 fun lex_ns:: "nat list \<Rightarrow> nat list \<Rightarrow> bool" where
528 "lex_ns [] ms = True"
529 | "lex_ns ns [] = False"
530 | "lex_ns (n#ns) (m#ms) = (n<m \<or> ((n = m) \<and> lex_ns ns ms)) "
531 definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" where
532 "lex_bnd t s \<equiv> lex_ns (bnds t) (bnds s)"
534 fun maxcoeff:: "num \<Rightarrow> int" where
535 "maxcoeff (C i) = abs i"
536 | "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)"
537 | "maxcoeff (CF c t s) = max (abs c) (maxcoeff s)"
540 lemma maxcoeff_pos: "maxcoeff t \<ge> 0"
541 apply (induct t rule: maxcoeff.induct, auto)
544 fun numgcdh:: "num \<Rightarrow> int \<Rightarrow> int" where
545 "numgcdh (C i) = (\<lambda>g. gcd i g)"
546 | "numgcdh (CN n c t) = (\<lambda>g. gcd c (numgcdh t g))"
547 | "numgcdh (CF c s t) = (\<lambda>g. gcd c (numgcdh t g))"
548 | "numgcdh t = (\<lambda>g. 1)"
551 numgcd :: "num \<Rightarrow> int"
553 numgcd_def: "numgcd t = numgcdh t (maxcoeff t)"
555 fun reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num" where
556 "reducecoeffh (C i) = (\<lambda> g. C (i div g))"
557 | "reducecoeffh (CN n c t) = (\<lambda> g. CN n (c div g) (reducecoeffh t g))"
558 | "reducecoeffh (CF c s t) = (\<lambda> g. CF (c div g) s (reducecoeffh t g))"
559 | "reducecoeffh t = (\<lambda>g. t)"
562 reducecoeff :: "num \<Rightarrow> num"
564 reducecoeff_def: "reducecoeff t =
566 if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)"
568 fun dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool" where
569 "dvdnumcoeff (C i) = (\<lambda> g. g dvd i)"
570 | "dvdnumcoeff (CN n c t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
571 | "dvdnumcoeff (CF c s t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))"
572 | "dvdnumcoeff t = (\<lambda>g. False)"
574 lemma dvdnumcoeff_trans:
575 assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'"
576 shows "dvdnumcoeff t g"
578 by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg dvd_trans[OF gdg])
580 declare dvd_trans [trans add]
583 assumes g0: "numgcd t = 0"
584 shows "Inum bs t = 0"
586 have "\<And>x. numgcdh t x= 0 \<Longrightarrow> Inum bs t = 0"
587 by (induct t rule: numgcdh.induct, auto)
588 thus ?thesis using g0[simplified numgcd_def] by blast
591 lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0"
593 by (induct t rule: numgcdh.induct, auto)
595 lemma numgcd_pos: "numgcd t \<ge>0"
596 by (simp add: numgcd_def numgcdh_pos maxcoeff_pos)
599 assumes gt: "dvdnumcoeff t g" and gp: "g > 0"
600 shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t"
602 proof(induct t rule: reducecoeffh.induct)
603 case (1 i) hence gd: "g dvd i" by simp
604 from assms 1 show ?case by (simp add: real_of_int_div[OF gd])
606 case (2 n c t) hence gd: "g dvd c" by simp
607 from assms 2 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps)
609 case (3 c s t) hence gd: "g dvd c" by simp
610 from assms 3 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps)
611 qed (auto simp add: numgcd_def gp)
613 fun ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool" where
614 "ismaxcoeff (C i) = (\<lambda> x. abs i \<le> x)"
615 | "ismaxcoeff (CN n c t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))"
616 | "ismaxcoeff (CF c s t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))"
617 | "ismaxcoeff t = (\<lambda>x. True)"
619 lemma ismaxcoeff_mono: "ismaxcoeff t c \<Longrightarrow> c \<le> c' \<Longrightarrow> ismaxcoeff t c'"
620 by (induct t rule: ismaxcoeff.induct, auto)
622 lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)"
623 proof (induct t rule: maxcoeff.induct)
625 hence H:"ismaxcoeff t (maxcoeff t)" .
626 have thh: "maxcoeff t \<le> max (abs c) (maxcoeff t)" by simp
627 from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1)
630 hence H1:"ismaxcoeff s (maxcoeff s)" by auto
631 have thh1: "maxcoeff s \<le> max \<bar>c\<bar> (maxcoeff s)" by (simp add: max_def)
632 from ismaxcoeff_mono[OF H1 thh1] show ?case by (simp add: le_maxI1)
635 lemma zgcd_gt1: "gcd i j > (1::int) \<Longrightarrow> ((abs i > 1 \<and> abs j > 1) \<or> (abs i = 0 \<and> abs j > 1) \<or> (abs i > 1 \<and> abs j = 0))"
636 apply (unfold gcd_int_def)
637 apply (cases "i = 0", simp_all)
638 apply (cases "j = 0", simp_all)
639 apply (cases "abs i = 1", simp_all)
640 apply (cases "abs j = 1", simp_all)
643 lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow> m =0"
644 by (induct t rule: numgcdh.induct) auto
646 lemma dvdnumcoeff_aux:
647 assumes "ismaxcoeff t m" and mp:"m \<ge> 0" and "numgcdh t m > 1"
648 shows "dvdnumcoeff t (numgcdh t m)"
650 proof(induct t rule: numgcdh.induct)
652 let ?g = "numgcdh t m"
653 from 2 have th:"gcd c ?g > 1" by simp
654 from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
655 have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
656 moreover {assume "abs c > 1" and gp: "?g > 1" with 2
657 have th: "dvdnumcoeff t ?g" by simp
658 have th': "gcd c ?g dvd ?g" by simp
659 from dvdnumcoeff_trans[OF th' th] have ?case by simp }
660 moreover {assume "abs c = 0 \<and> ?g > 1"
661 with 2 have th: "dvdnumcoeff t ?g" by simp
662 have th': "gcd c ?g dvd ?g" by simp
663 from dvdnumcoeff_trans[OF th' th] have ?case by simp
664 hence ?case by simp }
665 moreover {assume "abs c > 1" and g0:"?g = 0"
666 from numgcdh0[OF g0] have "m=0". with 2 g0 have ?case by simp }
667 ultimately show ?case by blast
670 let ?g = "numgcdh t m"
671 from 3 have th:"gcd c ?g > 1" by simp
672 from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"]
673 have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp
674 moreover {assume "abs c > 1" and gp: "?g > 1" with 3
675 have th: "dvdnumcoeff t ?g" by simp
676 have th': "gcd c ?g dvd ?g" by simp
677 from dvdnumcoeff_trans[OF th' th] have ?case by simp }
678 moreover {assume "abs c = 0 \<and> ?g > 1"
679 with 3 have th: "dvdnumcoeff t ?g" by simp
680 have th': "gcd c ?g dvd ?g" by simp
681 from dvdnumcoeff_trans[OF th' th] have ?case by simp
682 hence ?case by simp }
683 moreover {assume "abs c > 1" and g0:"?g = 0"
684 from numgcdh0[OF g0] have "m=0". with 3 g0 have ?case by simp }
685 ultimately show ?case by blast
688 lemma dvdnumcoeff_aux2:
689 assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0"
691 proof (simp add: numgcd_def)
692 let ?mc = "maxcoeff t"
693 let ?g = "numgcdh t ?mc"
694 have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff)
695 have th2: "?mc \<ge> 0" by (rule maxcoeff_pos)
696 assume H: "numgcdh t ?mc > 1"
697 from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" .
700 lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t"
703 have "?g \<ge> 0" by (simp add: numgcd_pos)
704 hence "?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto
705 moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)}
706 moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)}
707 moreover { assume g1:"?g > 1"
708 from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+
709 from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis
710 by (simp add: reducecoeff_def Let_def)}
711 ultimately show ?thesis by blast
714 lemma reducecoeffh_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeffh t g)"
715 by (induct t rule: reducecoeffh.induct, auto)
717 lemma reducecoeff_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeff t)"
718 using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def)
721 numadd:: "num \<times> num \<Rightarrow> num"
723 recdef numadd "measure (\<lambda> (t,s). size t + size s)"
724 "numadd (CN n1 c1 r1,CN n2 c2 r2) =
727 in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2))))
728 else if n1 \<le> n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2))
729 else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))"
730 "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))"
731 "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))"
732 "numadd (CF c1 t1 r1,CF c2 t2 r2) =
734 (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s))
735 else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2))
736 else CF c2 t2 (numadd(CF c1 t1 r1,r2)))"
737 "numadd (CF c1 t1 r1,C c) = CF c1 t1 (numadd (r1, C c))"
738 "numadd (C c,CF c1 t1 r1) = CF c1 t1 (numadd (r1, C c))"
739 "numadd (C b1, C b2) = C (b1+b2)"
740 "numadd (a,b) = Add a b"
742 lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)"
743 apply (induct t s rule: numadd.induct, simp_all add: Let_def)
744 apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all)
745 apply (case_tac "n1 = n2", simp_all add: algebra_simps)
746 apply (simp only: distrib_right[symmetric])
748 apply (case_tac "lex_bnd t1 t2", simp_all)
749 apply (case_tac "c1+c2 = 0")
750 by (case_tac "t1 = t2", simp_all add: algebra_simps distrib_right[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add distrib_right)
752 lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))"
753 by (induct t s rule: numadd.induct, auto simp add: Let_def)
755 fun nummul:: "num \<Rightarrow> int \<Rightarrow> num" where
756 "nummul (C j) = (\<lambda> i. C (i*j))"
757 | "nummul (CN n c t) = (\<lambda> i. CN n (c*i) (nummul t i))"
758 | "nummul (CF c t s) = (\<lambda> i. CF (c*i) t (nummul s i))"
759 | "nummul (Mul c t) = (\<lambda> i. nummul t (i*c))"
760 | "nummul t = (\<lambda> i. Mul i t)"
762 lemma nummul[simp]: "\<And> i. Inum bs (nummul t i) = Inum bs (Mul i t)"
763 by (induct t rule: nummul.induct, auto simp add: algebra_simps)
765 lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
766 by (induct t rule: nummul.induct, auto)
768 definition numneg :: "num \<Rightarrow> num" where
769 "numneg t \<equiv> nummul t (- 1)"
771 definition numsub :: "num \<Rightarrow> num \<Rightarrow> num" where
772 "numsub s t \<equiv> (if s = t then C 0 else numadd (s,numneg t))"
774 lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)"
775 using numneg_def nummul by simp
777 lemma numneg_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numneg t)"
778 using numneg_def by simp
780 lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)"
781 using numsub_def by simp
783 lemma numsub_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numsub t s)"
784 using numsub_def by simp
786 lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs"
788 have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor)
790 have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def)
791 also have "\<dots>" by (simp add: isint_add cti si)
792 finally show ?thesis .
795 fun split_int:: "num \<Rightarrow> num \<times> num" where
796 "split_int (C c) = (C 0, C c)"
797 | "split_int (CN n c b) =
798 (let (bv,bi) = split_int b
800 | "split_int (CF c a b) =
801 (let (bv,bi) = split_int b
803 | "split_int a = (a,C 0)"
805 lemma split_int: "\<And>tv ti. split_int t = (tv,ti) \<Longrightarrow> (Inum bs (Add tv ti) = Inum bs t) \<and> isint ti bs"
806 proof (induct t rule: split_int.induct)
808 let ?bv = "fst (split_int b)"
809 let ?bi = "snd (split_int b)"
810 have "split_int b = (?bv,?bi)" by simp
811 with 2(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
812 from 2(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def)
813 from 2(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def)
816 let ?bv = "fst (split_int b)"
817 let ?bi = "snd (split_int b)"
818 have "split_int b = (?bv,?bi)" by simp
819 with 3(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+
820 from 3(2) have tibi: "ti = CF c a ?bi"
821 by (simp add: Let_def split_def)
822 from 3(2) b[symmetric] bii show ?case
823 by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF)
824 qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def algebra_simps)
826 lemma split_int_nb: "numbound0 t \<Longrightarrow> numbound0 (fst (split_int t)) \<and> numbound0 (snd (split_int t)) "
827 by (induct t rule: split_int.induct) (auto simp add: Let_def split_def)
829 definition numfloor:: "num \<Rightarrow> num"
831 "numfloor t = (let (tv,ti) = split_int t in
832 (case tv of C i \<Rightarrow> numadd (tv,ti)
833 | _ \<Rightarrow> numadd(CF 1 tv (C 0),ti)))"
835 lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)")
837 let ?tv = "fst (split_int t)"
838 let ?ti = "snd (split_int t)"
839 have tvti:"split_int t = (?tv,?ti)" by simp
840 {assume H: "\<forall> v. ?tv \<noteq> C v"
841 hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)"
842 by (cases ?tv, auto simp add: numfloor_def Let_def split_def numadd)
843 from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
844 hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp
845 also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))"
846 by (simp,subst tii[simplified isint_iff, symmetric]) simp
847 also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff])
848 finally have ?thesis using th1 by simp}
849 moreover {fix v assume H:"?tv = C v"
850 from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+
851 hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp
852 also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))"
853 by (simp,subst tii[simplified isint_iff, symmetric]) simp
854 also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff])
855 finally have ?thesis by (simp add: H numfloor_def Let_def split_def numadd) }
856 ultimately show ?thesis by auto
859 lemma numfloor_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numfloor t)"
860 using split_int_nb[where t="t"]
861 by (cases "fst(split_int t)" , auto simp add: numfloor_def Let_def split_def numadd_nb)
863 function simpnum:: "num \<Rightarrow> num" where
864 "simpnum (C j) = C j"
865 | "simpnum (Bound n) = CN n 1 (C 0)"
866 | "simpnum (Neg t) = numneg (simpnum t)"
867 | "simpnum (Add t s) = numadd (simpnum t,simpnum s)"
868 | "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)"
869 | "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)"
870 | "simpnum (Floor t) = numfloor (simpnum t)"
871 | "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))"
872 | "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)"
873 by pat_completeness auto
874 termination by (relation "measure num_size") auto
876 lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t"
877 by (induct t rule: simpnum.induct, auto)
879 lemma simpnum_numbound0[simp]:
880 "numbound0 t \<Longrightarrow> numbound0 (simpnum t)"
881 by (induct t rule: simpnum.induct, auto)
883 fun nozerocoeff:: "num \<Rightarrow> bool" where
884 "nozerocoeff (C c) = True"
885 | "nozerocoeff (CN n c t) = (c\<noteq>0 \<and> nozerocoeff t)"
886 | "nozerocoeff (CF c s t) = (c \<noteq> 0 \<and> nozerocoeff t)"
887 | "nozerocoeff (Mul c t) = (c\<noteq>0 \<and> nozerocoeff t)"
888 | "nozerocoeff t = True"
890 lemma numadd_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numadd (a,b))"
891 by (induct a b rule: numadd.induct,auto simp add: Let_def)
893 lemma nummul_nz : "\<And> i. i\<noteq>0 \<Longrightarrow> nozerocoeff a \<Longrightarrow> nozerocoeff (nummul a i)"
894 by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz)
896 lemma numneg_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff (numneg a)"
897 by (simp add: numneg_def nummul_nz)
899 lemma numsub_nz: "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numsub a b)"
900 by (simp add: numsub_def numneg_nz numadd_nz)
902 lemma split_int_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (fst (split_int t)) \<and> nozerocoeff (snd (split_int t))"
903 by (induct t rule: split_int.induct,auto simp add: Let_def split_def)
905 lemma numfloor_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (numfloor t)"
906 by (simp add: numfloor_def Let_def split_def)
907 (cases "fst (split_int t)", simp_all add: split_int_nz numadd_nz)
909 lemma simpnum_nz: "nozerocoeff (simpnum t)"
910 by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz numfloor_nz)
912 lemma maxcoeff_nz: "nozerocoeff t \<Longrightarrow> maxcoeff t = 0 \<Longrightarrow> t = C 0"
913 proof (induct t rule: maxcoeff.induct)
915 hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+
916 have "max (abs c) (maxcoeff t) \<ge> abs c" by (simp add: le_maxI1)
917 with cnz have "max (abs c) (maxcoeff t) > 0" by arith
918 with 2 show ?case by simp
921 hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+
922 have "max (abs c) (maxcoeff t) \<ge> abs c" by (simp add: le_maxI1)
923 with cnz have "max (abs c) (maxcoeff t) > 0" by arith
924 with 3 show ?case by simp
927 lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0"
929 from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def)
930 from numgcdh0[OF th] have th:"maxcoeff t = 0" .
931 from maxcoeff_nz[OF nz th] show ?thesis .
934 definition simp_num_pair :: "(num \<times> int) \<Rightarrow> num \<times> int" where
935 "simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else
936 (let t' = simpnum t ; g = numgcd t' in
937 if g > 1 then (let g' = gcd n g in
938 if g' = 1 then (t',n)
939 else (reducecoeffh t' g', n div g'))
942 lemma simp_num_pair_ci:
943 shows "((\<lambda> (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\<lambda> (t,n). Inum bs t / real n) (t,n))"
946 let ?t' = "simpnum t"
947 let ?g = "numgcd ?t'"
949 {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
951 { assume nnz: "n \<noteq> 0"
952 {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
954 {assume g1:"?g>1" hence g0: "?g > 0" by simp
955 from g1 nnz have gp0: "?g' \<noteq> 0" by simp
956 hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
957 hence "?g'= 1 \<or> ?g' > 1" by arith
958 moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)}
959 moreover {assume g'1:"?g'>1"
960 from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" ..
961 let ?tt = "reducecoeffh ?t' ?g'"
962 let ?t = "Inum bs ?tt"
963 have gpdg: "?g' dvd ?g" by simp
964 have gpdd: "?g' dvd n" by simp
965 have gpdgp: "?g' dvd ?g'" by simp
966 from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p]
967 have th2:"real ?g' * ?t = Inum bs ?t'" by simp
968 from nnz g1 g'1 have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def)
969 also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp
970 also have "\<dots> = (Inum bs ?t' / real n)"
971 using real_of_int_div[OF gpdd] th2 gp0 by simp
972 finally have "?lhs = Inum bs t / real n" by simp
973 then have ?thesis using nnz g1 g'1 by (simp add: simp_num_pair_def) }
974 ultimately have ?thesis by blast }
975 ultimately have ?thesis by blast }
976 ultimately show ?thesis by blast
979 lemma simp_num_pair_l:
980 assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')"
981 shows "numbound0 t' \<and> n' >0"
983 let ?t' = "simpnum t"
984 let ?g = "numgcd ?t'"
986 { assume nz: "n = 0" hence ?thesis using assms by (simp add: Let_def simp_num_pair_def) }
988 { assume nnz: "n \<noteq> 0"
989 {assume "\<not> ?g > 1" hence ?thesis using assms by (auto simp add: Let_def simp_num_pair_def) }
991 {assume g1:"?g>1" hence g0: "?g > 0" by simp
992 from g1 nnz have gp0: "?g' \<noteq> 0" by simp
993 hence g'p: "?g' > 0" using gcd_ge_0_int[where x="n" and y="numgcd ?t'"] by arith
994 hence "?g'= 1 \<or> ?g' > 1" by arith
995 moreover {assume "?g'=1" hence ?thesis using assms g1 g0
996 by (auto simp add: Let_def simp_num_pair_def) }
997 moreover {assume g'1:"?g'>1"
998 have gpdg: "?g' dvd ?g" by simp
999 have gpdd: "?g' dvd n" by simp
1000 have gpdgp: "?g' dvd ?g'" by simp
1001 from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" .
1002 from zdiv_mono1[OF g'n g'p, simplified div_self[OF gp0]]
1003 have "n div ?g' >0" by simp
1004 hence ?thesis using assms g1 g'1
1005 by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)}
1006 ultimately have ?thesis by blast }
1007 ultimately have ?thesis by blast }
1008 ultimately show ?thesis by blast
1011 fun not:: "fm \<Rightarrow> fm" where
1015 | "not (Lt t) = Ge t"
1016 | "not (Le t) = Gt t"
1017 | "not (Gt t) = Le t"
1018 | "not (Ge t) = Lt t"
1019 | "not (Eq t) = NEq t"
1020 | "not (NEq t) = Eq t"
1021 | "not (Dvd i t) = NDvd i t"
1022 | "not (NDvd i t) = Dvd i t"
1023 | "not (And p q) = Or (not p) (not q)"
1024 | "not (Or p q) = And (not p) (not q)"
1026 lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
1028 lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
1030 lemma not_nb[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
1033 definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
1034 "conj p q \<equiv> (if (p = F \<or> q=F) then F else if p=T then q else if q=T then p else
1035 if p = q then p else And p q)"
1036 lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)"
1037 by (cases "p=F \<or> q=F", simp_all add: conj_def) (cases p, simp_all)
1039 lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
1040 using conj_def by auto
1041 lemma conj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
1042 using conj_def by auto
1044 definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
1045 "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p
1046 else if p=q then p else Or p q)"
1048 lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)"
1049 by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
1050 lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
1051 using disj_def by auto
1052 lemma disj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
1053 using disj_def by auto
1055 definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
1056 "imp p q \<equiv> (if (p = F \<or> q=T \<or> p=q) then T else if p=T then q else if q=F then not p
1058 lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)"
1059 by (cases "p=F \<or> q=T",simp_all add: imp_def)
1060 lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)"
1061 using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def)
1063 definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
1064 "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else
1065 if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else
1067 lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)"
1068 by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp add:not)
1069 (cases "not p= q", auto simp add:not)
1070 lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
1071 by (unfold iff_def,cases "p=q", auto)
1073 fun check_int:: "num \<Rightarrow> bool" where
1074 "check_int (C i) = True"
1075 | "check_int (Floor t) = True"
1076 | "check_int (Mul i t) = check_int t"
1077 | "check_int (Add t s) = (check_int t \<and> check_int s)"
1078 | "check_int (Neg t) = check_int t"
1079 | "check_int (CF c t s) = check_int s"
1080 | "check_int t = False"
1081 lemma check_int: "check_int t \<Longrightarrow> isint t bs"
1082 by (induct t, auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF)
1084 lemma rdvd_left1_int: "real \<lfloor>t\<rfloor> = t \<Longrightarrow> 1 rdvd t"
1085 by (simp add: rdvd_def,rule_tac x="\<lfloor>t\<rfloor>" in exI) simp
1088 assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0"
1089 shows "real (d::int) rdvd real (c::int)*t = (real (d div g) rdvd real (c div g)*t)"
1091 assume d: "real d rdvd real c * t"
1092 from d rdvd_def obtain k where k_def: "real c * t = real d* real (k::int)" by auto
1093 from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto
1094 from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto
1095 from k_def kd_def kc_def have "real g * real kc * t = real g * real kd * real k" by simp
1096 hence "real kc * t = real kd * real k" using gp by simp
1097 hence th:"real kd rdvd real kc * t" using rdvd_def by blast
1098 from kd_def gp have th':"kd = d div g" by simp
1099 from kc_def gp have "kc = c div g" by simp
1100 with th th' show "real (d div g) rdvd real (c div g) * t" by simp
1102 assume d: "real (d div g) rdvd real (c div g) * t"
1103 from gp have gnz: "g \<noteq> 0" by simp
1104 thus "real d rdvd real c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real (c div g) * t"] real_of_int_div[OF gd] real_of_int_div[OF gc] by simp
1107 definition simpdvd :: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)" where
1108 "simpdvd d t \<equiv>
1109 (let g = numgcd t in
1110 if g > 1 then (let g' = gcd d g in
1111 if g' = 1 then (d, t)
1112 else (d div g',reducecoeffh t g'))
1115 assumes tnz: "nozerocoeff t" and dnz: "d \<noteq> 0"
1116 shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)"
1119 let ?g' = "gcd d ?g"
1120 {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)}
1122 {assume g1:"?g>1" hence g0: "?g > 0" by simp
1123 from g1 dnz have gp0: "?g' \<noteq> 0" by simp
1124 hence g'p: "?g' > 0" using gcd_ge_0_int[where x="d" and y="numgcd t"] by arith
1125 hence "?g'= 1 \<or> ?g' > 1" by arith
1126 moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)}
1127 moreover {assume g'1:"?g'>1"
1128 from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" ..
1129 let ?tt = "reducecoeffh t ?g'"
1130 let ?t = "Inum bs ?tt"
1131 have gpdg: "?g' dvd ?g" by simp
1132 have gpdd: "?g' dvd d" by simp
1133 have gpdgp: "?g' dvd ?g'" by simp
1134 from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p]
1135 have th2:"real ?g' * ?t = Inum bs t" by simp
1136 from assms g1 g0 g'1
1137 have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)"
1138 by (simp add: simpdvd_def Let_def)
1139 also have "\<dots> = (real d rdvd (Inum bs t))"
1140 using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified div_self[OF gp0]]
1141 th2[symmetric] by simp
1142 finally have ?thesis by simp }
1143 ultimately have ?thesis by blast
1145 ultimately show ?thesis by blast
1148 function (sequential) simpfm :: "fm \<Rightarrow> fm" where
1149 "simpfm (And p q) = conj (simpfm p) (simpfm q)"
1150 | "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
1151 | "simpfm (Imp p q) = imp (simpfm p) (simpfm q)"
1152 | "simpfm (Iff p q) = iff (simpfm p) (simpfm q)"
1153 | "simpfm (NOT p) = not (simpfm p)"
1154 | "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F
1155 | _ \<Rightarrow> Lt (reducecoeff a'))"
1156 | "simpfm (Le a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<le> 0) then T else F | _ \<Rightarrow> Le (reducecoeff a'))"
1157 | "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v > 0) then T else F | _ \<Rightarrow> Gt (reducecoeff a'))"
1158 | "simpfm (Ge a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<ge> 0) then T else F | _ \<Rightarrow> Ge (reducecoeff a'))"
1159 | "simpfm (Eq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v = 0) then T else F | _ \<Rightarrow> Eq (reducecoeff a'))"
1160 | "simpfm (NEq a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v \<noteq> 0) then T else F | _ \<Rightarrow> NEq (reducecoeff a'))"
1161 | "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a)
1162 else if (abs i = 1) \<and> check_int a then T
1163 else let a' = simpnum a in case a' of C v \<Rightarrow> if (i dvd v) then T else F | _ \<Rightarrow> (let (d,t) = simpdvd i a' in Dvd d t))"
1164 | "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a)
1165 else if (abs i = 1) \<and> check_int a then F
1166 else let a' = simpnum a in case a' of C v \<Rightarrow> if (\<not>(i dvd v)) then T else F | _ \<Rightarrow> (let (d,t) = simpdvd i a' in NDvd d t))"
1168 by pat_completeness auto
1169 termination by (relation "measure fmsize") auto
1171 lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p"
1172 proof(induct p rule: simpfm.induct)
1173 case (6 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
1174 {fix v assume "?sa = C v" hence ?case using sa by simp }
1175 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
1176 let ?g = "numgcd ?sa"
1177 let ?rsa = "reducecoeff ?sa"
1178 let ?r = "Inum bs ?rsa"
1179 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
1180 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
1181 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
1182 hence gp: "real ?g > 0" by simp
1183 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
1184 with sa have "Inum bs a < 0 = (real ?g * ?r < real ?g * 0)" by simp
1185 also have "\<dots> = (?r < 0)" using gp
1186 by (simp only: mult_less_cancel_left) simp
1187 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
1188 ultimately show ?case by blast
1190 case (7 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
1191 {fix v assume "?sa = C v" hence ?case using sa by simp }
1192 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
1193 let ?g = "numgcd ?sa"
1194 let ?rsa = "reducecoeff ?sa"
1195 let ?r = "Inum bs ?rsa"
1196 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
1197 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
1198 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
1199 hence gp: "real ?g > 0" by simp
1200 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
1201 with sa have "Inum bs a \<le> 0 = (real ?g * ?r \<le> real ?g * 0)" by simp
1202 also have "\<dots> = (?r \<le> 0)" using gp
1203 by (simp only: mult_le_cancel_left) simp
1204 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
1205 ultimately show ?case by blast
1207 case (8 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
1208 {fix v assume "?sa = C v" hence ?case using sa by simp }
1209 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
1210 let ?g = "numgcd ?sa"
1211 let ?rsa = "reducecoeff ?sa"
1212 let ?r = "Inum bs ?rsa"
1213 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
1214 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
1215 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
1216 hence gp: "real ?g > 0" by simp
1217 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
1218 with sa have "Inum bs a > 0 = (real ?g * ?r > real ?g * 0)" by simp
1219 also have "\<dots> = (?r > 0)" using gp
1220 by (simp only: mult_less_cancel_left) simp
1221 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
1222 ultimately show ?case by blast
1224 case (9 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
1225 {fix v assume "?sa = C v" hence ?case using sa by simp }
1226 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
1227 let ?g = "numgcd ?sa"
1228 let ?rsa = "reducecoeff ?sa"
1229 let ?r = "Inum bs ?rsa"
1230 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
1231 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
1232 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
1233 hence gp: "real ?g > 0" by simp
1234 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
1235 with sa have "Inum bs a \<ge> 0 = (real ?g * ?r \<ge> real ?g * 0)" by simp
1236 also have "\<dots> = (?r \<ge> 0)" using gp
1237 by (simp only: mult_le_cancel_left) simp
1238 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
1239 ultimately show ?case by blast
1241 case (10 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
1242 {fix v assume "?sa = C v" hence ?case using sa by simp }
1243 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
1244 let ?g = "numgcd ?sa"
1245 let ?rsa = "reducecoeff ?sa"
1246 let ?r = "Inum bs ?rsa"
1247 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
1248 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
1249 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
1250 hence gp: "real ?g > 0" by simp
1251 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
1252 with sa have "Inum bs a = 0 = (real ?g * ?r = 0)" by simp
1253 also have "\<dots> = (?r = 0)" using gp
1254 by (simp add: mult_eq_0_iff)
1255 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
1256 ultimately show ?case by blast
1258 case (11 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
1259 {fix v assume "?sa = C v" hence ?case using sa by simp }
1260 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
1261 let ?g = "numgcd ?sa"
1262 let ?rsa = "reducecoeff ?sa"
1263 let ?r = "Inum bs ?rsa"
1264 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz)
1265 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto}
1266 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto)
1267 hence gp: "real ?g > 0" by simp
1268 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff)
1269 with sa have "Inum bs a \<noteq> 0 = (real ?g * ?r \<noteq> 0)" by simp
1270 also have "\<dots> = (?r \<noteq> 0)" using gp
1271 by (simp add: mult_eq_0_iff)
1272 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)}
1273 ultimately show ?case by blast
1275 case (12 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
1276 have "i=0 \<or> (abs i = 1 \<and> check_int a) \<or> (i\<noteq>0 \<and> ((abs i \<noteq> 1) \<or> (\<not> check_int a)))" by auto
1277 {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)}
1279 {assume ai1: "abs i = 1" and ai: "check_int a"
1280 hence "i=1 \<or> i= - 1" by arith
1281 moreover {assume i1: "i = 1"
1282 from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]]
1283 have ?case using i1 ai by simp }
1284 moreover {assume i1: "i = - 1"
1285 from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]]
1286 rdvd_abs1[where d="- 1" and t="Inum bs a"]
1287 have ?case using i1 ai by simp }
1288 ultimately have ?case by blast}
1290 {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
1291 {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
1292 by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
1293 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
1294 hence th: "simpfm (Dvd i a) = Dvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond by (cases ?sa, auto simp add: Let_def split_def)
1295 from simpnum_nz have nz:"nozerocoeff ?sa" by simp
1296 from simpdvd [OF nz inz] th have ?case using sa by simp}
1297 ultimately have ?case by blast}
1298 ultimately show ?case by blast
1300 case (13 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp
1301 have "i=0 \<or> (abs i = 1 \<and> check_int a) \<or> (i\<noteq>0 \<and> ((abs i \<noteq> 1) \<or> (\<not> check_int a)))" by auto
1302 {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)}
1304 {assume ai1: "abs i = 1" and ai: "check_int a"
1305 hence "i=1 \<or> i= - 1" by arith
1306 moreover {assume i1: "i = 1"
1307 from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]]
1308 have ?case using i1 ai by simp }
1309 moreover {assume i1: "i = - 1"
1310 from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]]
1311 rdvd_abs1[where d="- 1" and t="Inum bs a"]
1312 have ?case using i1 ai by simp }
1313 ultimately have ?case by blast}
1315 {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)"
1316 {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond
1317 by (cases "abs i = 1", auto simp add: int_rdvd_iff) }
1318 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)"
1319 hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond
1320 by (cases ?sa, auto simp add: Let_def split_def)
1321 from simpnum_nz have nz:"nozerocoeff ?sa" by simp
1322 from simpdvd [OF nz inz] th have ?case using sa by simp}
1323 ultimately have ?case by blast}
1324 ultimately show ?case by blast
1325 qed (induct p rule: simpfm.induct, simp_all)
1327 lemma simpdvd_numbound0: "numbound0 t \<Longrightarrow> numbound0 (snd (simpdvd d t))"
1328 by (simp add: simpdvd_def Let_def split_def reducecoeffh_numbound0)
1330 lemma simpfm_bound0[simp]: "bound0 p \<Longrightarrow> bound0 (simpfm p)"
1331 proof(induct p rule: simpfm.induct)
1332 case (6 a) hence nb: "numbound0 a" by simp
1333 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
1334 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
1336 case (7 a) hence nb: "numbound0 a" by simp
1337 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
1338 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
1340 case (8 a) hence nb: "numbound0 a" by simp
1341 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
1342 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
1344 case (9 a) hence nb: "numbound0 a" by simp
1345 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
1346 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
1348 case (10 a) hence nb: "numbound0 a" by simp
1349 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
1350 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
1352 case (11 a) hence nb: "numbound0 a" by simp
1353 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
1354 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0)
1356 case (12 i a) hence nb: "numbound0 a" by simp
1357 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
1358 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def)
1360 case (13 i a) hence nb: "numbound0 a" by simp
1361 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb])
1362 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def)
1363 qed(auto simp add: disj_def imp_def iff_def conj_def)
1365 lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)"
1366 by (induct p rule: simpfm.induct, auto simp add: Let_def)
1367 (case_tac "simpnum a",auto simp add: split_def Let_def)+
1370 (* Generic quantifier elimination *)
1372 definition list_conj :: "fm list \<Rightarrow> fm" where
1373 "list_conj ps \<equiv> foldr conj ps T"
1374 lemma list_conj: "Ifm bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm bs p)"
1375 by (induct ps, auto simp add: list_conj_def)
1376 lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)"
1377 by (induct ps, auto simp add: list_conj_def)
1378 lemma list_conj_nb: " \<forall>p\<in> set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
1379 by (induct ps, auto simp add: list_conj_def)
1380 definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
1381 "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = List.partition bound0 cjs
1382 in conj (decr (list_conj yes)) (f (list_conj no)))"
1385 assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
1386 shows "\<forall> bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm bs ((CJNB qe p)) = Ifm bs (E p))"
1389 assume qfp: "qfree p"
1390 let ?cjs = "conjuncts p"
1391 let ?yes = "fst (List.partition bound0 ?cjs)"
1392 let ?no = "snd (List.partition bound0 ?cjs)"
1393 let ?cno = "list_conj ?no"
1394 let ?cyes = "list_conj ?yes"
1395 have part: "List.partition bound0 ?cjs = (?yes,?no)" by simp
1396 from partition_P[OF part] have "\<forall> q\<in> set ?yes. bound0 q" by blast
1397 hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb)
1398 hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf)
1399 from conjuncts_qf[OF qfp] partition_set[OF part]
1400 have " \<forall>q\<in> set ?no. qfree q" by auto
1401 hence no_qf: "qfree ?cno"by (simp add: list_conj_qf)
1402 with qe have cno_qf:"qfree (qe ?cno )"
1403 and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+
1404 from cno_qf yes_qf have qf: "qfree (CJNB qe p)"
1405 by (simp add: CJNB_def Let_def conj_qf split_def)
1407 from conjuncts have "Ifm bs p = (\<forall>q\<in> set ?cjs. Ifm bs q)" by blast
1408 also have "\<dots> = ((\<forall>q\<in> set ?yes. Ifm bs q) \<and> (\<forall>q\<in> set ?no. Ifm bs q))"
1409 using partition_set[OF part] by auto
1410 finally have "Ifm bs p = ((Ifm bs ?cyes) \<and> (Ifm bs ?cno))" using list_conj by simp}
1411 hence "Ifm bs (E p) = (\<exists>x. (Ifm (x#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))" by simp
1412 also fix y have "\<dots> = (\<exists>x. (Ifm (y#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))"
1413 using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast
1414 also have "\<dots> = (Ifm bs (decr ?cyes) \<and> Ifm bs (E ?cno))"
1415 by (auto simp add: decr[OF yes_nb] simp del: partition_filter_conv)
1416 also have "\<dots> = (Ifm bs (conj (decr ?cyes) (qe ?cno)))"
1417 using qe[rule_format, OF no_qf] by auto
1418 finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)"
1419 by (simp add: Let_def CJNB_def split_def)
1420 with qf show "qfree (CJNB qe p) \<and> Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast
1423 function (sequential) qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm" where
1424 "qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
1425 | "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
1426 | "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
1427 | "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))"
1428 | "qelim (Or p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))"
1429 | "qelim (Imp p q) = (\<lambda> qe. disj (qelim (NOT p) qe) (qelim q qe))"
1430 | "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
1431 | "qelim p = (\<lambda> y. simpfm p)"
1432 by pat_completeness auto
1433 termination by (relation "measure fmsize") auto
1436 assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))"
1437 shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm bs (qelim p qe) = Ifm bs p)"
1438 using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]]
1439 by (induct p rule: qelim.induct) (auto simp del: simpfm.simps)
1442 text {* The @{text "\<int>"} Part *}
1443 text{* Linearity for fm where Bound 0 ranges over @{text "\<int>"} *}
1445 function zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*) where
1446 "zsplit0 (C c) = (0,C c)"
1447 | "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))"
1448 | "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)"
1449 | "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)"
1450 | "zsplit0 (Neg a) = (let (i',a') = zsplit0 a in (-i', Neg a'))"
1451 | "zsplit0 (Add a b) = (let (ia,a') = zsplit0 a ;
1453 in (ia+ib, Add a' b'))"
1454 | "zsplit0 (Sub a b) = (let (ia,a') = zsplit0 a ;
1456 in (ia-ib, Sub a' b'))"
1457 | "zsplit0 (Mul i a) = (let (i',a') = zsplit0 a in (i*i', Mul i a'))"
1458 | "zsplit0 (Floor a) = (let (i',a') = zsplit0 a in (i',Floor a'))"
1459 by pat_completeness auto
1460 termination by (relation "measure num_size") auto
1463 shows "\<And> n a. zsplit0 t = (n,a) \<Longrightarrow> (Inum ((real (x::int)) #bs) (CN 0 n a) = Inum (real x #bs) t) \<and> numbound0 a"
1464 (is "\<And> n a. ?S t = (n,a) \<Longrightarrow> (?I x (CN 0 n a) = ?I x t) \<and> ?N a")
1465 proof(induct t rule: zsplit0.induct)
1466 case (1 c n a) thus ?case by auto
1468 case (2 m n a) thus ?case by (cases "m=0") auto
1470 case (3 n i a n a') thus ?case by auto
1472 case (4 c a b n a') thus ?case by auto
1475 let ?nt = "fst (zsplit0 t)"
1476 let ?at = "snd (zsplit0 t)"
1477 have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \<and> n=-?nt" using 5
1478 by (simp add: Let_def split_def)
1479 from abj 5 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
1480 from th2[simplified] th[simplified] show ?case by simp
1483 let ?ns = "fst (zsplit0 s)"
1484 let ?as = "snd (zsplit0 s)"
1485 let ?nt = "fst (zsplit0 t)"
1486 let ?at = "snd (zsplit0 t)"
1487 have abjs: "zsplit0 s = (?ns,?as)" by simp
1488 moreover have abjt: "zsplit0 t = (?nt,?at)" by simp
1489 ultimately have th: "a=Add ?as ?at \<and> n=?ns + ?nt" using 6
1490 by (simp add: Let_def split_def)
1491 from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
1492 from 6 have "(\<exists> x y. (x,y) = zsplit0 s) \<longrightarrow> (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \<and> numbound0 xb)" by blast (*FIXME*)
1493 with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
1494 from abjs 6 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
1495 from th3[simplified] th2[simplified] th[simplified] show ?case
1496 by (simp add: distrib_right)
1499 let ?ns = "fst (zsplit0 s)"
1500 let ?as = "snd (zsplit0 s)"
1501 let ?nt = "fst (zsplit0 t)"
1502 let ?at = "snd (zsplit0 t)"
1503 have abjs: "zsplit0 s = (?ns,?as)" by simp
1504 moreover have abjt: "zsplit0 t = (?nt,?at)" by simp
1505 ultimately have th: "a=Sub ?as ?at \<and> n=?ns - ?nt" using 7
1506 by (simp add: Let_def split_def)
1507 from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast
1508 from 7 have "(\<exists> x y. (x,y) = zsplit0 s) \<longrightarrow> (\<forall>xa xb. zsplit0 t = (xa, xb) \<longrightarrow> Inum (real x # bs) (CN 0 xa xb) = Inum (real x # bs) t \<and> numbound0 xb)" by blast (*FIXME*)
1509 with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
1510 from abjs 7 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast
1511 from th3[simplified] th2[simplified] th[simplified] show ?case
1512 by (simp add: left_diff_distrib)
1515 let ?nt = "fst (zsplit0 t)"
1516 let ?at = "snd (zsplit0 t)"
1517 have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \<and> n=i*?nt" using 8
1518 by (simp add: Let_def split_def)
1519 from abj 8 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
1520 hence "?I x (Mul i t) = (real i) * ?I x (CN 0 ?nt ?at)" by simp
1521 also have "\<dots> = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: distrib_left)
1522 finally show ?case using th th2 by simp
1525 let ?nt = "fst (zsplit0 t)"
1526 let ?at = "snd (zsplit0 t)"
1527 have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a= Floor ?at \<and> n=?nt" using 9
1528 by (simp add: Let_def split_def)
1529 from abj 9 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast
1530 hence na: "?N a" using th by simp
1531 have th': "(real ?nt)*(real x) = real (?nt * x)" by simp
1532 have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp
1533 also have "\<dots> = real (floor ((real ?nt)* real(x) + ?I x ?at))" by simp
1534 also have "\<dots> = real (floor (?I x ?at + real (?nt* x)))" by (simp add: add_ac)
1535 also have "\<dots> = real (floor (?I x ?at) + (?nt* x))"
1536 using floor_add[where x="?I x ?at" and a="?nt* x"] by simp
1537 also have "\<dots> = real (?nt)*(real x) + real (floor (?I x ?at))" by (simp add: add_ac)
1538 finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp
1539 with na show ?case by simp
1543 iszlfm :: "fm \<Rightarrow> real list \<Rightarrow> bool" (* Linearity test for fm *)
1544 zlfm :: "fm \<Rightarrow> fm" (* Linearity transformation for fm *)
1545 recdef iszlfm "measure size"
1546 "iszlfm (And p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)"
1547 "iszlfm (Or p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)"
1548 "iszlfm (Eq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
1549 "iszlfm (NEq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
1550 "iszlfm (Lt (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
1551 "iszlfm (Le (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
1552 "iszlfm (Gt (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
1553 "iszlfm (Ge (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)"
1554 "iszlfm (Dvd i (CN 0 c e)) =
1555 (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
1556 "iszlfm (NDvd i (CN 0 c e))=
1557 (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)"
1558 "iszlfm p = (\<lambda> bs. isatom p \<and> (bound0 p))"
1560 lemma zlin_qfree: "iszlfm p bs \<Longrightarrow> qfree p"
1561 by (induct p rule: iszlfm.induct) auto
1564 assumes lp: "iszlfm p (x#bs)"
1565 shows "\<forall> y. iszlfm p (y#bs)"
1568 show "iszlfm p (y#bs)"
1570 by(induct p rule: iszlfm.induct, simp_all add: numbound0_gen[rule_format, where x="x" and y="y"])
1573 lemma conj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (conj p q) bs"
1574 using conj_def by (cases p,auto)
1575 lemma disj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (disj p q) bs"
1576 using disj_def by (cases p,auto)
1578 recdef zlfm "measure fmsize"
1579 "zlfm (And p q) = conj (zlfm p) (zlfm q)"
1580 "zlfm (Or p q) = disj (zlfm p) (zlfm q)"
1581 "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)"
1582 "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))"
1583 "zlfm (Lt a) = (let (c,r) = zsplit0 a in
1584 if c=0 then Lt r else
1585 if c>0 then Or (Lt (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r)))
1586 else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
1587 "zlfm (Le a) = (let (c,r) = zsplit0 a in
1588 if c=0 then Le r else
1589 if c>0 then Or (Le (CN 0 c (Neg (Floor (Neg r))))) (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Lt (Add (Floor (Neg r)) r)))
1590 else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))"
1591 "zlfm (Gt a) = (let (c,r) = zsplit0 a in
1592 if c=0 then Gt r else
1593 if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r)))
1594 else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
1595 "zlfm (Ge a) = (let (c,r) = zsplit0 a in
1596 if c=0 then Ge r else
1597 if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r)))
1598 else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))"
1599 "zlfm (Eq a) = (let (c,r) = zsplit0 a in
1600 if c=0 then Eq r else
1601 if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r)))
1602 else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))"
1603 "zlfm (NEq a) = (let (c,r) = zsplit0 a in
1604 if c=0 then NEq r else
1605 if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r)))
1606 else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))"
1607 "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a)
1608 else (let (c,r) = zsplit0 a in
1609 if c=0 then Dvd (abs i) r else
1610 if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r)))
1611 else And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 (-c) (Neg (Floor r))))))"
1612 "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a)
1613 else (let (c,r) = zsplit0 a in
1614 if c=0 then NDvd (abs i) r else
1615 if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r)))
1616 else Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 (-c) (Neg (Floor r))))))"
1617 "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))"
1618 "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))"
1619 "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))"
1620 "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))"
1621 "zlfm (NOT (NOT p)) = zlfm p"
1624 "zlfm (NOT (Lt a)) = zlfm (Ge a)"
1625 "zlfm (NOT (Le a)) = zlfm (Gt a)"
1626 "zlfm (NOT (Gt a)) = zlfm (Le a)"
1627 "zlfm (NOT (Ge a)) = zlfm (Lt a)"
1628 "zlfm (NOT (Eq a)) = zlfm (NEq a)"
1629 "zlfm (NOT (NEq a)) = zlfm (Eq a)"
1630 "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)"
1631 "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)"
1632 "zlfm p = p" (hints simp add: fmsize_pos)
1634 lemma split_int_less_real:
1635 "(real (a::int) < b) = (a < floor b \<or> (a = floor b \<and> real (floor b) < b))"
1637 assume alb: "real a < b" and agb: "\<not> a < floor b"
1638 from agb have "floor b \<le> a" by simp hence th: "b < real a + 1" by (simp only: floor_le_eq)
1639 from floor_eq[OF alb th] show "a= floor b" by simp
1641 assume alb: "a < floor b"
1642 hence "real a < real (floor b)" by simp
1643 moreover have "real (floor b) \<le> b" by simp ultimately show "real a < b" by arith
1646 lemma split_int_less_real':
1647 "(real (a::int) + b < 0) = (real a - real (floor(-b)) < 0 \<or> (real a - real (floor (-b)) = 0 \<and> real (floor (-b)) + b < 0))"
1649 have "(real a + b <0) = (real a < -b)" by arith
1650 with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith
1653 lemma split_int_gt_real':
1654 "(real (a::int) + b > 0) = (real a + real (floor b) > 0 \<or> (real a + real (floor b) = 0 \<and> real (floor b) - b < 0))"
1656 have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith
1657 show ?thesis using myless[of _ "real (floor b)"]
1658 by (simp only:th split_int_less_real'[where a="-a" and b="-b"])
1659 (simp add: algebra_simps diff_minus[symmetric],arith)
1662 lemma split_int_le_real:
1663 "(real (a::int) \<le> b) = (a \<le> floor b \<or> (a = floor b \<and> real (floor b) < b))"
1665 assume alb: "real a \<le> b" and agb: "\<not> a \<le> floor b"
1666 from alb have "floor (real a) \<le> floor b " by (simp only: floor_mono)
1667 hence "a \<le> floor b" by simp with agb show "False" by simp
1669 assume alb: "a \<le> floor b"
1670 hence "real a \<le> real (floor b)" by (simp only: floor_mono)
1671 also have "\<dots>\<le> b" by simp finally show "real a \<le> b" .
1674 lemma split_int_le_real':
1675 "(real (a::int) + b \<le> 0) = (real a - real (floor(-b)) \<le> 0 \<or> (real a - real (floor (-b)) = 0 \<and> real (floor (-b)) + b < 0))"
1677 have "(real a + b \<le>0) = (real a \<le> -b)" by arith
1678 with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith
1681 lemma split_int_ge_real':
1682 "(real (a::int) + b \<ge> 0) = (real a + real (floor b) \<ge> 0 \<or> (real a + real (floor b) = 0 \<and> real (floor b) - b < 0))"
1684 have th: "(real a + b \<ge>0) = (real (-a) + (-b) \<le> 0)" by arith
1685 show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"])
1686 (simp add: algebra_simps diff_minus[symmetric],arith)
1689 lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \<and> b = real (floor b))" (is "?l = ?r")
1692 lemma split_int_eq_real': "(real (a::int) + b = 0) = ( a - floor (-b) = 0 \<and> real (floor (-b)) + b = 0)" (is "?l = ?r")
1694 have "?l = (real a = -b)" by arith
1695 with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith
1699 assumes qfp: "qfree p"
1700 shows "(Ifm (real i #bs) (zlfm p) = Ifm (real i# bs) p) \<and> iszlfm (zlfm p) (real (i::int) #bs)"
1701 (is "(?I (?l p) = ?I p) \<and> ?L (?l p)")
1703 proof(induct p rule: zlfm.induct)
1705 let ?c = "fst (zsplit0 a)"
1706 let ?r = "snd (zsplit0 a)"
1707 have spl: "zsplit0 a = (?c,?r)" by simp
1708 from zsplit0_I[OF spl, where x="i" and bs="bs"]
1709 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
1710 let ?N = "\<lambda> t. Inum (real i#bs) t"
1711 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
1713 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
1714 by (cases "?r", simp_all add: Let_def split_def,case_tac "nat", simp_all)}
1716 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))"
1717 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1718 have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
1719 also have "\<dots> = (?I (?l (Lt a)))" apply (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) by (simp add: Ia cp cnz Let_def split_def diff_minus)
1720 finally have ?case using l by simp}
1722 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))"
1723 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1724 have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def)
1725 also from cn cnz have "\<dots> = (?I (?l (Lt a)))" by (simp only: split_int_less_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
1726 finally have ?case using l by simp}
1727 ultimately show ?case by blast
1730 let ?c = "fst (zsplit0 a)"
1731 let ?r = "snd (zsplit0 a)"
1732 have spl: "zsplit0 a = (?c,?r)" by simp
1733 from zsplit0_I[OF spl, where x="i" and bs="bs"]
1734 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
1735 let ?N = "\<lambda> t. Inum (real i#bs) t"
1736 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
1738 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
1739 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat",simp_all)}
1741 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))"
1742 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1743 have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
1744 also have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
1745 finally have ?case using l by simp}
1747 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))"
1748 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1749 have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def)
1750 also from cn cnz have "\<dots> = (?I (?l (Le a)))" by (simp only: split_int_le_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac ,arith)
1751 finally have ?case using l by simp}
1752 ultimately show ?case by blast
1755 let ?c = "fst (zsplit0 a)"
1756 let ?r = "snd (zsplit0 a)"
1757 have spl: "zsplit0 a = (?c,?r)" by simp
1758 from zsplit0_I[OF spl, where x="i" and bs="bs"]
1759 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
1760 let ?N = "\<lambda> t. Inum (real i#bs) t"
1761 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
1763 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
1764 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
1766 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))"
1767 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1768 have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
1769 also have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
1770 finally have ?case using l by simp}
1772 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))"
1773 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1774 have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def)
1775 also from cn cnz have "\<dots> = (?I (?l (Gt a)))" by (simp only: split_int_gt_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
1776 finally have ?case using l by simp}
1777 ultimately show ?case by blast
1780 let ?c = "fst (zsplit0 a)"
1781 let ?r = "snd (zsplit0 a)"
1782 have spl: "zsplit0 a = (?c,?r)" by simp
1783 from zsplit0_I[OF spl, where x="i" and bs="bs"]
1784 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
1785 let ?N = "\<lambda> t. Inum (real i#bs) t"
1786 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
1788 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
1789 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
1791 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))"
1792 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1793 have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
1794 also have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia cp cnz Let_def split_def diff_minus)
1795 finally have ?case using l by simp}
1797 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))"
1798 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1799 have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def)
1800 also from cn cnz have "\<dots> = (?I (?l (Ge a)))" by (simp only: split_int_ge_real'[where a="?c*i" and b="?N ?r"]) (simp add: Ia Let_def split_def diff_minus[symmetric] add_ac, arith)
1801 finally have ?case using l by simp}
1802 ultimately show ?case by blast
1805 let ?c = "fst (zsplit0 a)"
1806 let ?r = "snd (zsplit0 a)"
1807 have spl: "zsplit0 a = (?c,?r)" by simp
1808 from zsplit0_I[OF spl, where x="i" and bs="bs"]
1809 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
1810 let ?N = "\<lambda> t. Inum (real i#bs) t"
1811 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
1813 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
1814 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
1816 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))"
1817 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1818 have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
1819 also have "\<dots> = (?I (?l (Eq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult)
1820 finally have ?case using l by simp}
1822 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))"
1823 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1824 have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def)
1825 also from cn cnz have "\<dots> = (?I (?l (Eq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith)
1826 finally have ?case using l by simp}
1827 ultimately show ?case by blast
1830 let ?c = "fst (zsplit0 a)"
1831 let ?r = "snd (zsplit0 a)"
1832 have spl: "zsplit0 a = (?c,?r)" by simp
1833 from zsplit0_I[OF spl, where x="i" and bs="bs"]
1834 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
1835 let ?N = "\<lambda> t. Inum (real i#bs) t"
1836 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith
1838 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"]
1839 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
1841 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))"
1842 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1843 have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
1844 also have "\<dots> = (?I (?l (NEq a)))" using cp cnz by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult)
1845 finally have ?case using l by simp}
1847 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))"
1848 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1849 have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def)
1850 also from cn cnz have "\<dots> = (?I (?l (NEq a)))" by (simp only: split_int_eq_real'[where a="?c*i" and b="?N ?r"]) (simp add: Let_def split_def Ia real_of_int_mult[symmetric] del: real_of_int_mult,arith)
1851 finally have ?case using l by simp}
1852 ultimately show ?case by blast
1855 let ?c = "fst (zsplit0 a)"
1856 let ?r = "snd (zsplit0 a)"
1857 have spl: "zsplit0 a = (?c,?r)" by simp
1858 from zsplit0_I[OF spl, where x="i" and bs="bs"]
1859 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
1860 let ?N = "\<lambda> t. Inum (real i#bs) t"
1861 have "j=0 \<or> (j\<noteq>0 \<and> ?c = 0) \<or> (j\<noteq>0 \<and> ?c >0 \<and> ?c\<noteq>0) \<or> (j\<noteq> 0 \<and> ?c<0 \<and> ?c\<noteq>0)" by arith
1863 { assume j: "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def)
1864 hence ?case using 11 j by (simp del: zlfm.simps add: rdvd_left_0_eq)}
1866 {assume "?c=0" and "j\<noteq>0" hence ?case
1867 using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
1868 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
1870 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))"
1871 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1872 have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))"
1873 using Ia by (simp add: Let_def split_def)
1874 also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))"
1875 by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
1876 also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and>
1877 (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))"
1878 by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
1879 also have "\<dots> = (?I (?l (Dvd j a)))" using cp cnz jnz
1880 by (simp add: Let_def split_def int_rdvd_iff[symmetric]
1881 del: real_of_int_mult) (auto simp add: add_ac)
1882 finally have ?case using l jnz by simp }
1884 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))"
1885 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1886 have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))"
1887 using Ia by (simp add: Let_def split_def)
1888 also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))"
1889 by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
1890 also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and>
1891 (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))"
1892 by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
1893 also have "\<dots> = (?I (?l (Dvd j a)))" using cn cnz jnz
1894 using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
1895 by (simp add: Let_def split_def int_rdvd_iff[symmetric]
1896 del: real_of_int_mult) (auto simp add: add_ac)
1897 finally have ?case using l jnz by blast }
1898 ultimately show ?case by blast
1901 let ?c = "fst (zsplit0 a)"
1902 let ?r = "snd (zsplit0 a)"
1903 have spl: "zsplit0 a = (?c,?r)" by simp
1904 from zsplit0_I[OF spl, where x="i" and bs="bs"]
1905 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto
1906 let ?N = "\<lambda> t. Inum (real i#bs) t"
1907 have "j=0 \<or> (j\<noteq>0 \<and> ?c = 0) \<or> (j\<noteq>0 \<and> ?c >0 \<and> ?c\<noteq>0) \<or> (j\<noteq> 0 \<and> ?c<0 \<and> ?c\<noteq>0)" by arith
1909 {assume j: "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def)
1910 hence ?case using 12 j by (simp del: zlfm.simps add: rdvd_left_0_eq)}
1912 {assume "?c=0" and "j\<noteq>0" hence ?case
1913 using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"]
1914 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)}
1916 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))"
1917 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1918 have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))"
1919 using Ia by (simp add: Let_def split_def)
1920 also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))"
1921 by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
1922 also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and>
1923 (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))"
1924 by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
1925 also have "\<dots> = (?I (?l (NDvd j a)))" using cp cnz jnz
1926 by (simp add: Let_def split_def int_rdvd_iff[symmetric]
1927 del: real_of_int_mult) (auto simp add: add_ac)
1928 finally have ?case using l jnz by simp }
1930 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))"
1931 by (simp add: nb Let_def split_def isint_Floor isint_neg)
1932 have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))"
1933 using Ia by (simp add: Let_def split_def)
1934 also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))"
1935 by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp
1936 also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and>
1937 (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))"
1938 by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac)
1939 also have "\<dots> = (?I (?l (NDvd j a)))" using cn cnz jnz
1940 using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric]
1941 by (simp add: Let_def split_def int_rdvd_iff[symmetric]
1942 del: real_of_int_mult) (auto simp add: add_ac)
1943 finally have ?case using l jnz by blast }
1944 ultimately show ?case by blast
1947 text{* plusinf : Virtual substitution of @{text "+\<infinity>"}
1948 minusinf: Virtual substitution of @{text "-\<infinity>"}
1949 @{text "\<delta>"} Compute lcm @{text "d| Dvd d c*x+t \<in> p"}
1950 @{text "d_\<delta>"} checks if a given l divides all the ds above*}
1952 fun minusinf:: "fm \<Rightarrow> fm" where
1953 "minusinf (And p q) = conj (minusinf p) (minusinf q)"
1954 | "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
1955 | "minusinf (Eq (CN 0 c e)) = F"
1956 | "minusinf (NEq (CN 0 c e)) = T"
1957 | "minusinf (Lt (CN 0 c e)) = T"
1958 | "minusinf (Le (CN 0 c e)) = T"
1959 | "minusinf (Gt (CN 0 c e)) = F"
1960 | "minusinf (Ge (CN 0 c e)) = F"
1963 lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)"
1964 by (induct p rule: minusinf.induct, auto)
1966 fun plusinf:: "fm \<Rightarrow> fm" where
1967 "plusinf (And p q) = conj (plusinf p) (plusinf q)"
1968 | "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
1969 | "plusinf (Eq (CN 0 c e)) = F"
1970 | "plusinf (NEq (CN 0 c e)) = T"
1971 | "plusinf (Lt (CN 0 c e)) = F"
1972 | "plusinf (Le (CN 0 c e)) = F"
1973 | "plusinf (Gt (CN 0 c e)) = T"
1974 | "plusinf (Ge (CN 0 c e)) = T"
1977 fun \<delta> :: "fm \<Rightarrow> int" where
1978 "\<delta> (And p q) = lcm (\<delta> p) (\<delta> q)"
1979 | "\<delta> (Or p q) = lcm (\<delta> p) (\<delta> q)"
1980 | "\<delta> (Dvd i (CN 0 c e)) = i"
1981 | "\<delta> (NDvd i (CN 0 c e)) = i"
1984 fun d_\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool" where
1985 "d_\<delta> (And p q) = (\<lambda> d. d_\<delta> p d \<and> d_\<delta> q d)"
1986 | "d_\<delta> (Or p q) = (\<lambda> d. d_\<delta> p d \<and> d_\<delta> q d)"
1987 | "d_\<delta> (Dvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
1988 | "d_\<delta> (NDvd i (CN 0 c e)) = (\<lambda> d. i dvd d)"
1989 | "d_\<delta> p = (\<lambda> d. True)"
1992 assumes lin: "iszlfm p bs"
1994 and ad: "d_\<delta> p d"
1995 shows "d_\<delta> p d'"
1997 proof(induct p rule: iszlfm.induct)
1998 case (9 i c e) thus ?case using d
1999 by (simp add: dvd_trans[of "i" "d" "d'"])
2001 case (10 i c e) thus ?case using d
2002 by (simp add: dvd_trans[of "i" "d" "d'"])
2005 lemma \<delta> : assumes lin:"iszlfm p bs"
2006 shows "d_\<delta> p (\<delta> p) \<and> \<delta> p >0"
2008 proof (induct p rule: iszlfm.induct)
2010 let ?d = "\<delta> (And p q)"
2011 from 1 lcm_pos_int have dp: "?d >0" by simp
2012 have d1: "\<delta> p dvd \<delta> (And p q)" using 1 by simp
2013 hence th: "d_\<delta> p ?d"
2014 using delta_mono 1 by (simp only: iszlfm.simps) blast
2015 have "\<delta> q dvd \<delta> (And p q)" using 1 by simp
2016 hence th': "d_\<delta> q ?d" using delta_mono 1 by (simp only: iszlfm.simps) blast
2017 from th th' dp show ?case by simp
2020 let ?d = "\<delta> (And p q)"
2021 from 2 lcm_pos_int have dp: "?d >0" by simp
2022 have "\<delta> p dvd \<delta> (And p q)" using 2 by simp
2023 hence th: "d_\<delta> p ?d" using delta_mono 2 by (simp only: iszlfm.simps) blast
2024 have "\<delta> q dvd \<delta> (And p q)" using 2 by simp
2025 hence th': "d_\<delta> q ?d" using delta_mono 2 by (simp only: iszlfm.simps) blast
2026 from th th' dp show ?case by simp
2031 assumes linp: "iszlfm p (a # bs)"
2032 shows "\<exists> (z::int). \<forall> x < z. Ifm ((real x)#bs) (minusinf p) = Ifm ((real x)#bs) p"
2033 (is "?P p" is "\<exists> (z::int). \<forall> x < z. ?I x (?M p) = ?I x p")
2035 proof (induct p rule: minusinf.induct)
2037 then have "?P f" by simp
2038 then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast
2039 with 1 have "?P g" by simp
2040 then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast
2041 let ?z = "min z1 z2"
2042 from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (And f g)) = ?I x (And f g)" by simp
2046 then have "?P f" by simp
2047 then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast
2048 with 2 have "?P g" by simp
2049 then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast
2050 let ?z = "min z1 z2"
2051 from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp
2055 then have "c > 0" by simp
2056 hence rcpos: "real c > 0" by simp
2057 from 3 have nbe: "numbound0 e" by simp
2059 have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))"
2060 proof (simp add: less_floor_eq , rule allI, rule impI)
2062 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
2063 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
2064 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))"
2065 by (simp only: mult_strict_left_mono [OF th1 rcpos])
2066 hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos by simp
2067 thus "real c * real x + Inum (real x # bs) e \<noteq> 0"
2068 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp
2073 then have "c > 0" by simp hence rcpos: "real c > 0" by simp
2074 from 4 have nbe: "numbound0 e" by simp
2076 have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))"
2077 proof (simp add: less_floor_eq , rule allI, rule impI)
2079 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
2080 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
2081 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))"
2082 by (simp only: mult_strict_left_mono [OF th1 rcpos])
2083 hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos by simp
2084 thus "real c * real x + Inum (real x # bs) e \<noteq> 0"
2085 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp
2090 then have "c > 0" by simp hence rcpos: "real c > 0" by simp
2091 from 5 have nbe: "numbound0 e" by simp
2093 have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))"
2094 proof (simp add: less_floor_eq , rule allI, rule impI)
2096 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
2097 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
2098 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))"
2099 by (simp only: mult_strict_left_mono [OF th1 rcpos])
2100 thus "real c * real x + Inum (real x # bs) e < 0"
2101 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
2106 then have "c > 0" by simp hence rcpos: "real c > 0" by simp
2107 from 6 have nbe: "numbound0 e" by simp
2109 have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))"
2110 proof (simp add: less_floor_eq , rule allI, rule impI)
2112 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
2113 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
2114 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))"
2115 by (simp only: mult_strict_left_mono [OF th1 rcpos])
2116 thus "real c * real x + Inum (real x # bs) e \<le> 0"
2117 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
2122 then have "c > 0" by simp hence rcpos: "real c > 0" by simp
2123 from 7 have nbe: "numbound0 e" by simp
2125 have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))"
2126 proof (simp add: less_floor_eq , rule allI, rule impI)
2128 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
2129 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
2130 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))"
2131 by (simp only: mult_strict_left_mono [OF th1 rcpos])
2132 thus "\<not> (real c * real x + Inum (real x # bs) e>0)"
2133 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
2138 then have "c > 0" by simp hence rcpos: "real c > 0" by simp
2139 from 8 have nbe: "numbound0 e" by simp
2141 have "\<forall> x < (floor (- (Inum (y#bs) e) / (real c))). ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))"
2142 proof (simp add: less_floor_eq , rule allI, rule impI)
2144 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)"
2145 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp
2146 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))"
2147 by (simp only: mult_strict_left_mono [OF th1 rcpos])
2148 thus "\<not> real c * real x + Inum (real x # bs) e \<ge> 0"
2149 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp
2154 lemma minusinf_repeats:
2155 assumes d: "d_\<delta> p d" and linp: "iszlfm p (a # bs)"
2156 shows "Ifm ((real(x - k*d))#bs) (minusinf p) = Ifm (real x #bs) (minusinf p)"
2158 proof(induct p rule: iszlfm.induct)
2159 case (9 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+
2160 hence "\<exists> k. d=i*k" by (simp add: dvd_def)
2161 then obtain "di" where di_def: "d=i*di" by blast
2163 proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI)
2165 "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
2166 (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
2167 hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
2168 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))"
2169 by (simp add: algebra_simps di_def)
2170 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
2171 by (simp add: algebra_simps)
2172 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
2173 thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
2176 "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
2177 hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
2178 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
2179 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def)
2180 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps)
2181 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
2183 thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
2186 case (10 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+
2187 hence "\<exists> k. d=i*k" by (simp add: dvd_def)
2188 then obtain "di" where di_def: "d=i*di" by blast
2190 proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real x - real k * real d" and b'="real x"] right_diff_distrib, rule iffI)
2192 "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e"
2193 (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt")
2194 hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def)
2195 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))"
2196 by (simp add: algebra_simps di_def)
2197 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))"
2198 by (simp add: algebra_simps)
2199 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast
2200 thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp
2203 "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e")
2204 hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def)
2205 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp
2206 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real i * real di)" by (simp add: di_def)
2207 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps)
2208 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)"
2210 thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp
2212 qed (auto simp add: numbound0_I[where bs="bs" and b="real(x - k*d)" and b'="real x"] simp del: real_of_int_mult real_of_int_diff)
2215 assumes lin: "iszlfm p (real (a::int) #bs)"
2216 and exmi: "\<exists> (x::int). Ifm (real x#bs) (minusinf p)" (is "\<exists> x. ?P1 x")
2217 shows "\<exists> (x::int). Ifm (real x#bs) p" (is "\<exists> x. ?P x")
2219 let ?d = "\<delta> p"
2220 from \<delta> [OF lin] have dpos: "?d >0" by simp
2221 from \<delta> [OF lin] have alld: "d_\<delta> p ?d" by simp
2222 from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P1 x = ?P1 (x - (k * ?d))" by simp
2223 from minusinf_inf[OF lin] have th2:"\<exists> z. \<forall> x. x<z \<longrightarrow> (?P x = ?P1 x)" by blast
2224 from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast
2228 assumes lin: "iszlfm p (real (a::int) #bs)"
2229 shows "(\<exists> (x::int). Ifm (real x#bs) (minusinf p)) =
2230 (\<exists> (x::int)\<in> {1..\<delta> p}. Ifm (real x#bs) (minusinf p))"
2231 (is "(\<exists> x. ?P x) = _")
2233 let ?d = "\<delta> p"
2234 from \<delta> [OF lin] have dpos: "?d >0" by simp
2235 from \<delta> [OF lin] have alld: "d_\<delta> p ?d" by simp
2236 from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P x = ?P (x - (k * ?d))" by simp
2237 from periodic_finite_ex[OF dpos th1] show ?thesis by blast
2240 lemma dvd1_eq1: "x >0 \<Longrightarrow> (x::int) dvd 1 = (x = 1)" by auto
2243 a_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coeffitients of a formula *)
2244 d_\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*)
2245 \<zeta> :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*)
2246 \<beta> :: "fm \<Rightarrow> num list"
2247 \<alpha> :: "fm \<Rightarrow> num list"
2249 recdef a_\<beta> "measure size"
2250 "a_\<beta> (And p q) = (\<lambda> k. And (a_\<beta> p k) (a_\<beta> q k))"
2251 "a_\<beta> (Or p q) = (\<lambda> k. Or (a_\<beta> p k) (a_\<beta> q k))"
2252 "a_\<beta> (Eq (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))"
2253 "a_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))"
2254 "a_\<beta> (Lt (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))"
2255 "a_\<beta> (Le (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))"
2256 "a_\<beta> (Gt (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))"
2257 "a_\<beta> (Ge (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))"
2258 "a_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
2259 "a_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))"
2260 "a_\<beta> p = (\<lambda> k. p)"
2262 recdef d_\<beta> "measure size"
2263 "d_\<beta> (And p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
2264 "d_\<beta> (Or p q) = (\<lambda> k. (d_\<beta> p k) \<and> (d_\<beta> q k))"
2265 "d_\<beta> (Eq (CN 0 c e)) = (\<lambda> k. c dvd k)"
2266 "d_\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)"
2267 "d_\<beta> (Lt (CN 0 c e)) = (\<lambda> k. c dvd k)"
2268 "d_\<beta> (Le (CN 0 c e)) = (\<lambda> k. c dvd k)"
2269 "d_\<beta> (Gt (CN 0 c e)) = (\<lambda> k. c dvd k)"
2270 "d_\<beta> (Ge (CN 0 c e)) = (\<lambda> k. c dvd k)"
2271 "d_\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)"
2272 "d_\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)"
2273 "d_\<beta> p = (\<lambda> k. True)"
2275 recdef \<zeta> "measure size"
2276 "\<zeta> (And p q) = lcm (\<zeta> p) (\<zeta> q)"
2277 "\<zeta> (Or p q) = lcm (\<zeta> p) (\<zeta> q)"
2278 "\<zeta> (Eq (CN 0 c e)) = c"
2279 "\<zeta> (NEq (CN 0 c e)) = c"
2280 "\<zeta> (Lt (CN 0 c e)) = c"
2281 "\<zeta> (Le (CN 0 c e)) = c"
2282 "\<zeta> (Gt (CN 0 c e)) = c"
2283 "\<zeta> (Ge (CN 0 c e)) = c"
2284 "\<zeta> (Dvd i (CN 0 c e)) = c"
2285 "\<zeta> (NDvd i (CN 0 c e))= c"
2288 recdef \<beta> "measure size"
2289 "\<beta> (And p q) = (\<beta> p @ \<beta> q)"
2290 "\<beta> (Or p q) = (\<beta> p @ \<beta> q)"
2291 "\<beta> (Eq (CN 0 c e)) = [Sub (C -1) e]"
2292 "\<beta> (NEq (CN 0 c e)) = [Neg e]"
2293 "\<beta> (Lt (CN 0 c e)) = []"
2294 "\<beta> (Le (CN 0 c e)) = []"
2295 "\<beta> (Gt (CN 0 c e)) = [Neg e]"
2296 "\<beta> (Ge (CN 0 c e)) = [Sub (C -1) e]"
2299 recdef \<alpha> "measure size"
2300 "\<alpha> (And p q) = (\<alpha> p @ \<alpha> q)"
2301 "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)"
2302 "\<alpha> (Eq (CN 0 c e)) = [Add (C -1) e]"
2303 "\<alpha> (NEq (CN 0 c e)) = [e]"
2304 "\<alpha> (Lt (CN 0 c e)) = [e]"
2305 "\<alpha> (Le (CN 0 c e)) = [Add (C -1) e]"
2306 "\<alpha> (Gt (CN 0 c e)) = []"
2307 "\<alpha> (Ge (CN 0 c e)) = []"
2309 consts mirror :: "fm \<Rightarrow> fm"
2310 recdef mirror "measure size"
2311 "mirror (And p q) = And (mirror p) (mirror q)"
2312 "mirror (Or p q) = Or (mirror p) (mirror q)"
2313 "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))"
2314 "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))"
2315 "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))"
2316 "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))"
2317 "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))"
2318 "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))"
2319 "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))"
2320 "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))"
2323 lemma mirror_\<alpha>_\<beta>:
2324 assumes lp: "iszlfm p (a#bs)"
2325 shows "(Inum (real (i::int)#bs)) ` set (\<alpha> p) = (Inum (real i#bs)) ` set (\<beta> (mirror p))"
2327 by (induct p rule: mirror.induct, auto)
2330 assumes lp: "iszlfm p (a#bs)"
2331 shows "Ifm (real (x::int)#bs) (mirror p) = Ifm (real (- x)#bs) p"
2333 proof(induct p rule: iszlfm.induct)
2335 have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
2336 (real j rdvd - (real c * real x - Inum (real x # bs) e))"
2337 by (simp only: rdvd_minus[symmetric])
2338 from 9 th show ?case
2339 by (simp add: algebra_simps
2340 numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
2343 have th: "(real j rdvd real c * real x - Inum (real x # bs) e) =
2344 (real j rdvd - (real c * real x - Inum (real x # bs) e))"
2345 by (simp only: rdvd_minus[symmetric])
2346 from 10 th show ?case
2347 by (simp add: algebra_simps
2348 numbound0_I[where bs="bs" and b'="real x" and b="- real x"])
2349 qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"])
2351 lemma mirror_l: "iszlfm p (a#bs) \<Longrightarrow> iszlfm (mirror p) (a#bs)"
2352 by (induct p rule: mirror.induct, auto simp add: isint_neg)
2354 lemma mirror_d_\<beta>: "iszlfm p (a#bs) \<and> d_\<beta> p 1
2355 \<Longrightarrow> iszlfm (mirror p) (a#bs) \<and> d_\<beta> (mirror p) 1"
2356 by (induct p rule: mirror.induct, auto simp add: isint_neg)
2358 lemma mirror_\<delta>: "iszlfm p (a#bs) \<Longrightarrow> \<delta> (mirror p) = \<delta> p"
2359 by (induct p rule: mirror.induct,auto)
2363 assumes lp: "iszlfm p (real (i::int)#bs)"
2364 shows "(\<exists> (x::int). Ifm (real x#bs) (mirror p)) = (\<exists> (x::int). Ifm (real x#bs) p)"
2365 (is "(\<exists> x. ?I x ?mp) = (\<exists> x. ?I x p)")
2367 fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast
2368 thus "\<exists> x. ?I x p" by blast
2370 fix x assume "?I x p" hence "?I (- x) ?mp"
2371 using mirror[OF lp, where x="- x", symmetric] by auto
2372 thus "\<exists> x. ?I x ?mp" by blast
2375 lemma \<beta>_numbound0: assumes lp: "iszlfm p bs"
2376 shows "\<forall> b\<in> set (\<beta> p). numbound0 b"
2377 using lp by (induct p rule: \<beta>.induct,auto)
2379 lemma d_\<beta>_mono:
2380 assumes linp: "iszlfm p (a #bs)"
2381 and dr: "d_\<beta> p l"
2383 shows "d_\<beta> p l'"
2384 using dr linp dvd_trans[of _ "l" "l'", simplified d]
2385 by (induct p rule: iszlfm.induct) simp_all
2387 lemma \<alpha>_l: assumes lp: "iszlfm p (a#bs)"
2388 shows "\<forall> b\<in> set (\<alpha> p). numbound0 b \<and> isint b (a#bs)"
2390 by(induct p rule: \<alpha>.induct, auto simp add: isint_add isint_c)
2393 assumes linp: "iszlfm p (a #bs)"
2394 shows "\<zeta> p > 0 \<and> d_\<beta> p (\<zeta> p)"
2396 proof(induct p rule: iszlfm.induct)
2398 then have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
2399 from 1 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
2400 from 1 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
2401 d_\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
2402 dl1 dl2 show ?case by (auto simp add: lcm_pos_int)
2405 then have dl1: "\<zeta> p dvd lcm (\<zeta> p) (\<zeta> q)" by simp
2406 from 2 have dl2: "\<zeta> q dvd lcm (\<zeta> p) (\<zeta> q)" by simp
2407 from 2 d_\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="lcm (\<zeta> p) (\<zeta> q)"]
2408 d_\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="lcm (\<zeta> p) (\<zeta> q)"]
2409 dl1 dl2 show ?case by (auto simp add: lcm_pos_int)
2410 qed (auto simp add: lcm_pos_int)
2412 lemma a_\<beta>: assumes linp: "iszlfm p (a #bs)" and d: "d_\<beta> p l" and lp: "l > 0"
2413 shows "iszlfm (a_\<beta> p l) (a #bs) \<and> d_\<beta> (a_\<beta> p l) 1 \<and> (Ifm (real (l * x) #bs) (a_\<beta> p l) = Ifm ((real x)#bs) p)"
2415 proof (induct p rule: iszlfm.induct)
2416 case (5 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
2417 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
2418 from cp have cnz: "c \<noteq> 0" by simp
2419 have "c div c\<le> l div c"
2420 by (simp add: zdiv_mono1[OF clel cp])
2421 then have ldcp:"0 < l div c"
2422 by (simp add: div_self[OF cnz])
2423 have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
2424 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
2426 hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\<Colon>real)) =
2427 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)"
2429 also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) < (real (l div c)) * 0)" by (simp add: algebra_simps)
2430 also have "\<dots> = (real c * real x + Inum (real x # bs) e < 0)"
2431 using mult_less_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
2432 finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp
2434 case (6 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
2435 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
2436 from cp have cnz: "c \<noteq> 0" by simp
2437 have "c div c\<le> l div c"
2438 by (simp add: zdiv_mono1[OF clel cp])
2439 then have ldcp:"0 < l div c"
2440 by (simp add: div_self[OF cnz])
2441 have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
2442 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
2444 hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<le> (0\<Colon>real)) =
2445 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<le> 0)"
2447 also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<le> (real (l div c)) * 0)" by (simp add: algebra_simps)
2448 also have "\<dots> = (real c * real x + Inum (real x # bs) e \<le> 0)"
2449 using mult_le_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
2450 finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp
2452 case (7 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
2453 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
2454 from cp have cnz: "c \<noteq> 0" by simp
2455 have "c div c\<le> l div c"
2456 by (simp add: zdiv_mono1[OF clel cp])
2457 then have ldcp:"0 < l div c"
2458 by (simp add: div_self[OF cnz])
2459 have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
2460 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
2462 hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\<Colon>real)) =
2463 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)"
2465 also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) > (real (l div c)) * 0)" by (simp add: algebra_simps)
2466 also have "\<dots> = (real c * real x + Inum (real x # bs) e > 0)"
2467 using zero_less_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
2468 finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp
2470 case (8 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
2471 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
2472 from cp have cnz: "c \<noteq> 0" by simp
2473 have "c div c\<le> l div c"
2474 by (simp add: zdiv_mono1[OF clel cp])
2475 then have ldcp:"0 < l div c"
2476 by (simp add: div_self[OF cnz])
2477 have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
2478 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
2480 hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<ge> (0\<Colon>real)) =
2481 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<ge> 0)"
2483 also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<ge> (real (l div c)) * 0)" by (simp add: algebra_simps)
2484 also have "\<dots> = (real c * real x + Inum (real x # bs) e \<ge> 0)"
2485 using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
2486 finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp
2488 case (3 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
2489 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
2490 from cp have cnz: "c \<noteq> 0" by simp
2491 have "c div c\<le> l div c"
2492 by (simp add: zdiv_mono1[OF clel cp])
2493 then have ldcp:"0 < l div c"
2494 by (simp add: div_self[OF cnz])
2495 have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
2496 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
2498 hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\<Colon>real)) =
2499 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)"
2501 also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) = (real (l div c)) * 0)" by (simp add: algebra_simps)
2502 also have "\<dots> = (real c * real x + Inum (real x # bs) e = 0)"
2503 using mult_eq_0_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
2504 finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp
2506 case (4 c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and d': "c dvd l" by simp+
2507 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
2508 from cp have cnz: "c \<noteq> 0" by simp
2509 have "c div c\<le> l div c"
2510 by (simp add: zdiv_mono1[OF clel cp])
2511 then have ldcp:"0 < l div c"
2512 by (simp add: div_self[OF cnz])
2513 have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
2514 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
2516 hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<noteq> (0\<Colon>real)) =
2517 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<noteq> 0)"
2519 also have "\<dots> = (real (l div c) * (real c * real x + Inum (real x # bs) e) \<noteq> (real (l div c)) * 0)" by (simp add: algebra_simps)
2520 also have "\<dots> = (real c * real x + Inum (real x # bs) e \<noteq> 0)"
2521 using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e"] ldcp by simp
2522 finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] be isint_Mul[OF ei] by simp
2524 case (9 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+
2525 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
2526 from cp have cnz: "c \<noteq> 0" by simp
2527 have "c div c\<le> l div c"
2528 by (simp add: zdiv_mono1[OF clel cp])
2529 then have ldcp:"0 < l div c"
2530 by (simp add: div_self[OF cnz])
2531 have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
2532 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
2534 hence "(\<exists> (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\<exists> (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)" by simp
2535 also have "\<dots> = (\<exists> (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: algebra_simps)
2536 also fix k have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
2537 using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp
2538 also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
2539 finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp
2541 case (10 j c e) hence cp: "c>0" and be: "numbound0 e" and ei:"isint e (a#bs)" and jp: "j > 0" and d': "c dvd l" by simp+
2542 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp])
2543 from cp have cnz: "c \<noteq> 0" by simp
2544 have "c div c\<le> l div c"
2545 by (simp add: zdiv_mono1[OF clel cp])
2546 then have ldcp:"0 < l div c"
2547 by (simp add: div_self[OF cnz])
2548 have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp
2549 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric]
2551 hence "(\<exists> (k::int). real l * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k) = (\<exists> (k::int). real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = (real (l div c) * real j) * real k)" by simp
2552 also have "\<dots> = (\<exists> (k::int). real (l div c) * (real c * real x + Inum (real x # bs) e - real j * real k) = real (l div c)*0)" by (simp add: algebra_simps)
2553 also fix k have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)"
2554 using zero_le_mult_iff [where a="real (l div c)" and b="real c * real x + Inum (real x # bs) e - real j * real k"] ldcp by simp
2555 also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp
2556 finally show ?case using numbound0_I[OF be,where b="real (l * x)" and b'="real x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp
2557 qed (simp_all add: numbound0_I[where bs="bs" and b="real (l * x)" and b'="real x"] isint_Mul del: real_of_int_mult)
2559 lemma a_\<beta>_ex: assumes linp: "iszlfm p (a#bs)" and d: "d_\<beta> p l" and lp: "l>0"
2560 shows "(\<exists> x. l dvd x \<and> Ifm (real x #bs) (a_\<beta> p l)) = (\<exists> (x::int). Ifm (real x#bs) p)"
2561 (is "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> x. ?P' x)")
2563 have "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> (x::int). ?P (l*x))"
2564 using unity_coeff_ex[where l="l" and P="?P", simplified] by simp
2565 also have "\<dots> = (\<exists> (x::int). ?P' x)" using a_\<beta>[OF linp d lp] by simp
2566 finally show ?thesis .
2570 assumes lp: "iszlfm p (a#bs)"
2571 and u: "d_\<beta> p 1"
2572 and d: "d_\<delta> p d"
2574 and nob: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)"
2575 and p: "Ifm (real x#bs) p" (is "?P x")
2577 using lp u d dp nob p
2578 proof(induct p rule: iszlfm.induct)
2579 case (5 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all
2580 with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] 5
2581 show ?case by (simp del: real_of_int_minus)
2583 case (6 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp_all
2584 with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] 6
2585 show ?case by (simp del: real_of_int_minus)
2587 case (7 c e) hence p: "Ifm (real x #bs) (Gt (CN 0 c e))" and c1: "c=1"
2588 and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp_all
2589 let ?e = "Inum (real x # bs) e"
2590 from ie1 have ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"]
2591 numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]
2592 by (simp add: isint_iff)
2593 {assume "real (x-d) +?e > 0" hence ?case using c1
2594 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
2595 by (simp del: real_of_int_minus)}
2597 {assume H: "\<not> real (x-d) + ?e > 0"
2599 have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" by simp
2600 from 7(5)[simplified simp_thms Inum.simps \<beta>.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]]
2601 have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x = - ?e + real j)" by auto
2602 from H p have "real x + ?e > 0 \<and> real x + ?e \<le> real d" by (simp add: c1)
2603 hence "real (x + floor ?e) > real (0::int) \<and> real (x + floor ?e) \<le> real d"
2605 hence "x + floor ?e \<ge> 1 \<and> x + floor ?e \<le> d" by simp
2606 hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e" by simp
2607 hence "\<exists> (j::int) \<in> {1 .. d}. real x = real (- floor ?e + j)"
2608 by (simp only: real_of_int_inject) (simp add: algebra_simps)
2609 hence "\<exists> (j::int) \<in> {1 .. d}. real x = - ?e + real j"
2610 by (simp add: ie[simplified isint_iff])
2611 with nob have ?case by auto}
2612 ultimately show ?case by blast
2614 case (8 c e) hence p: "Ifm (real x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e"
2615 and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
2616 let ?e = "Inum (real x # bs) e"
2617 from ie1 have ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"]
2618 by (simp add: isint_iff)
2619 {assume "real (x-d) +?e \<ge> 0" hence ?case using c1
2620 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"]
2621 by (simp del: real_of_int_minus)}
2623 {assume H: "\<not> real (x-d) + ?e \<ge> 0"
2624 let ?v="Sub (C -1) e"
2625 have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" by simp
2626 from 8(5)[simplified simp_thms Inum.simps \<beta>.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"]]
2627 have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x = - ?e - 1 + real j)" by auto
2628 from H p have "real x + ?e \<ge> 0 \<and> real x + ?e < real d" by (simp add: c1)
2629 hence "real (x + floor ?e) \<ge> real (0::int) \<and> real (x + floor ?e) < real d"
2631 hence "x + floor ?e +1 \<ge> 1 \<and> x + floor ?e + 1 \<le> d" by simp
2632 hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e + 1" by simp
2633 hence "\<exists> (j::int) \<in> {1 .. d}. x= - floor ?e - 1 + j" by (simp add: algebra_simps)
2634 hence "\<exists> (j::int) \<in> {1 .. d}. real x= real (- floor ?e - 1 + j)"
2635 by (simp only: real_of_int_inject)
2636 hence "\<exists> (j::int) \<in> {1 .. d}. real x= - ?e - 1 + real j"
2637 by (simp add: ie[simplified isint_iff])
2638 with nob have ?case by simp }
2639 ultimately show ?case by blast
2641 case (3 c e) hence p: "Ifm (real x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1"
2642 and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
2643 let ?e = "Inum (real x # bs) e"
2644 let ?v="(Sub (C -1) e)"
2645 have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp
2646 from p have "real x= - ?e" by (simp add: c1) with 3(5) show ?case using dp
2647 by simp (erule ballE[where x="1"],
2648 simp_all add:algebra_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"])
2650 case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1"
2651 and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+
2652 let ?e = "Inum (real x # bs) e"
2654 have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))" by simp
2655 {assume "real x - real d + Inum ((real (x -d)) # bs) e \<noteq> 0"
2656 hence ?case by (simp add: c1)}
2658 {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0"
2659 hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp
2660 hence "real x = - Inum (a # bs) e + real d"
2661 by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"])
2662 with 4(5) have ?case using dp by simp}
2663 ultimately show ?case by blast
2665 case (9 j c e) hence p: "Ifm (real x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1"
2666 and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
2667 let ?e = "Inum (real x # bs) e"
2668 from 9 have "isint e (a #bs)" by simp
2669 hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real x)#bs"] numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"]
2670 by (simp add: isint_iff)
2671 from 9 have id: "j dvd d" by simp
2672 from c1 ie[symmetric] have "?p x = (real j rdvd real (x+ floor ?e))" by simp
2673 also have "\<dots> = (j dvd x + floor ?e)"
2674 using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp
2675 also have "\<dots> = (j dvd x - d + floor ?e)"
2676 using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp
2677 also have "\<dots> = (real j rdvd real (x - d + floor ?e))"
2678 using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified]
2680 also have "\<dots> = (real j rdvd real x - real d + ?e)"
2683 using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp
2685 case (10 j c e) hence p: "Ifm (real x #bs) (NDvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+
2686 let ?e = "Inum (real x # bs) e"
2687 from 10 have "isint e (a#bs)" by simp
2688 hence ie: "real (floor ?e) = ?e" using numbound0_I[OF bn,where b="real x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real x)#bs"]
2689 by (simp add: isint_iff)
2690 from 10 have id: "j dvd d" by simp
2691 from c1 ie[symmetric] have "?p x = (\<not> real j rdvd real (x+ floor ?e))" by simp
2692 also have "\<dots> = (\<not> j dvd x + floor ?e)"
2693 using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp
2694 also have "\<dots> = (\<not> j dvd x - d + floor ?e)"
2695 using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp
2696 also have "\<dots> = (\<not> real j rdvd real (x - d + floor ?e))"
2697 using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified]
2699 also have "\<dots> = (\<not> real j rdvd real x - real d + ?e)"
2702 using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp
2703 qed (auto simp add: numbound0_I[where bs="bs" and b="real (x - d)" and b'="real x"]
2704 simp del: real_of_int_diff)
2707 assumes lp: "iszlfm p (a #bs)"
2708 and u: "d_\<beta> p 1"
2709 and d: "d_\<delta> p d"
2711 shows "\<forall> x. \<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> set(\<beta> p). Ifm ((Inum (a#bs) b + real j) #bs) p) \<longrightarrow> Ifm (real x#bs) p \<longrightarrow> Ifm (real (x - d)#bs) p" (is "\<forall> x. ?b \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)")
2714 assume nb:"?b" and px: "?P x"
2715 hence nb2: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)"
2717 from \<beta>[OF lp u d dp nb2 px] show "?P (x -d )" .
2720 lemma \<beta>_int: assumes lp: "iszlfm p bs"
2721 shows "\<forall> b\<in> set (\<beta> p). isint b bs"
2722 using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub)
2724 lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x))
2725 ==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D)
2726 ==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D))))
2727 ==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))"
2730 apply(drule minusinfinity)
2734 apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)")
2735 apply(frule_tac x = x and z=z in decr_lemma)
2736 apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)")
2738 apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)")
2739 prefer 2 apply arith
2741 apply(drule (1) periodic_finite_ex)
2743 apply(blast dest:decr_mult_lemma)
2748 assumes lp: "iszlfm p (a #bs)"
2749 and u: "d_\<beta> p 1"
2750 and d: "d_\<delta> p d"
2752 shows "(\<exists> (x::int). Ifm (real x #bs) p) = (\<exists> j\<in> {1.. d}. Ifm (real j #bs) (minusinf p) \<or> (\<exists> b \<in> set (\<beta> p). Ifm ((Inum (a#bs) b + real j) #bs) p))"
2753 (is "(\<exists> (x::int). ?P (real x)) = (\<exists> j\<in> ?D. ?M j \<or> (\<exists> b\<in> ?B. ?P (?I b + real j)))")
2755 from minusinf_inf[OF lp]
2756 have th: "\<exists>(z::int). \<forall>x<z. ?P (real x) = ?M x" by blast
2757 let ?B' = "{floor (?I b) | b. b\<in> ?B}"
2758 from \<beta>_int[OF lp] isint_iff[where bs="a # bs"] have B: "\<forall> b\<in> ?B. real (floor (?I b)) = ?I b" by simp
2760 have "(\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (?I b + real j)) = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real (floor (?I b)) + real j))"
2762 also have "\<dots> = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real (floor (?I b) + j)))" by simp
2763 also have"\<dots> = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j)))" by blast
2765 "(\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (?I b + real j)) = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j)))"
2767 hence th2: "\<forall> x. \<not> (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j))) \<longrightarrow> ?P (real x) \<longrightarrow> ?P (real (x - d))" using \<beta>'[OF lp u d dp] by blast
2768 from minusinf_repeats[OF d lp]
2769 have th3: "\<forall> x k. ?M x = ?M (x-k*d)" by simp
2770 from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast
2773 (* Reddy and Loveland *)
2777 \<rho> :: "fm \<Rightarrow> (num \<times> int) list" (* Compute the Reddy and Loveland Bset*)
2778 \<sigma>_\<rho>:: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm" (* Performs the modified substitution of Reddy and Loveland*)
2779 \<alpha>_\<rho> :: "fm \<Rightarrow> (num\<times>int) list"
2780 a_\<rho> :: "fm \<Rightarrow> int \<Rightarrow> fm"
2781 recdef \<rho> "measure size"
2782 "\<rho> (And p q) = (\<rho> p @ \<rho> q)"
2783 "\<rho> (Or p q) = (\<rho> p @ \<rho> q)"
2784 "\<rho> (Eq (CN 0 c e)) = [(Sub (C -1) e,c)]"
2785 "\<rho> (NEq (CN 0 c e)) = [(Neg e,c)]"
2786 "\<rho> (Lt (CN 0 c e)) = []"
2787 "\<rho> (Le (CN 0 c e)) = []"
2788 "\<rho> (Gt (CN 0 c e)) = [(Neg e, c)]"
2789 "\<rho> (Ge (CN 0 c e)) = [(Sub (C (-1)) e, c)]"
2792 recdef \<sigma>_\<rho> "measure size"
2793 "\<sigma>_\<rho> (And p q) = (\<lambda> (t,k). And (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))"
2794 "\<sigma>_\<rho> (Or p q) = (\<lambda> (t,k). Or (\<sigma>_\<rho> p (t,k)) (\<sigma>_\<rho> q (t,k)))"
2795 "\<sigma>_\<rho> (Eq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e))
2796 else (Eq (Add (Mul c t) (Mul k e))))"
2797 "\<sigma>_\<rho> (NEq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e))
2798 else (NEq (Add (Mul c t) (Mul k e))))"
2799 "\<sigma>_\<rho> (Lt (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e))
2800 else (Lt (Add (Mul c t) (Mul k e))))"
2801 "\<sigma>_\<rho> (Le (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e))
2802 else (Le (Add (Mul c t) (Mul k e))))"
2803 "\<sigma>_\<rho> (Gt (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e))
2804 else (Gt (Add (Mul c t) (Mul k e))))"
2805 "\<sigma>_\<rho> (Ge (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e))
2806 else (Ge (Add (Mul c t) (Mul k e))))"
2807 "\<sigma>_\<rho> (Dvd i (CN 0 c e)) =(\<lambda> (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e))
2808 else (Dvd (i*k) (Add (Mul c t) (Mul k e))))"
2809 "\<sigma>_\<rho> (NDvd i (CN 0 c e))=(\<lambda> (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e))
2810 else (NDvd (i*k) (Add (Mul c t) (Mul k e))))"
2811 "\<sigma>_\<rho> p = (\<lambda> (t,k). p)"
2813 recdef \<alpha>_\<rho> "measure size"
2814 "\<alpha>_\<rho> (And p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)"
2815 "\<alpha>_\<rho> (Or p q) = (\<alpha>_\<rho> p @ \<alpha>_\<rho> q)"
2816 "\<alpha>_\<rho> (Eq (CN 0 c e)) = [(Add (C -1) e,c)]"
2817 "\<alpha>_\<rho> (NEq (CN 0 c e)) = [(e,c)]"
2818 "\<alpha>_\<rho> (Lt (CN 0 c e)) = [(e,c)]"
2819 "\<alpha>_\<rho> (Le (CN 0 c e)) = [(Add (C -1) e,c)]"
2820 "\<alpha>_\<rho> p = []"
2822 (* Simulates normal substituion by modifying the formula see correctness theorem *)
2824 definition \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
2825 "\<sigma> p k t \<equiv> And (Dvd k t) (\<sigma>_\<rho> p (t,k))"
2827 lemma \<sigma>_\<rho>:
2828 assumes linp: "iszlfm p (real (x::int)#bs)"
2829 and kpos: "real k > 0"
2830 and tnb: "numbound0 t"
2831 and tint: "isint t (real x#bs)"
2832 and kdt: "k dvd floor (Inum (b'#bs) t)"
2833 shows "Ifm (real x#bs) (\<sigma>_\<rho> p (t,k)) =
2834 (Ifm ((real ((floor (Inum (b'#bs) t)) div k))#bs) p)"
2835 (is "?I (real x) (?s p) = (?I (real ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)")
2837 proof(induct p rule: \<sigma>_\<rho>.induct)
2839 from 3 have cp: "c > 0" and nb: "numbound0 e" by auto
2840 { assume kdc: "k dvd c"
2841 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2842 from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
2843 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2844 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) }
2846 { assume *: "\<not> k dvd c"
2847 from kpos have knz': "real k \<noteq> 0" by simp
2848 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t"
2849 using isint_def by simp
2850 from assms * have "?I (real x) (?s (Eq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k = 0)"
2851 using real_of_int_div[OF kdt]
2852 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2853 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2854 by (simp add: ti algebra_simps)
2855 also have "\<dots> = (?I ?tk (Eq (CN 0 c e)))"
2856 using nonzero_eq_divide_eq[OF knz',
2857 where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
2858 real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2859 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2861 finally have ?case . }
2862 ultimately show ?case by blast
2865 then have cp: "c > 0" and nb: "numbound0 e" by auto
2866 { assume kdc: "k dvd c"
2867 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2868 from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
2869 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2870 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) }
2872 { assume *: "\<not> k dvd c"
2873 from kpos have knz': "real k \<noteq> 0" by simp
2874 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2875 from assms * have "?I (real x) (?s (NEq (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<noteq> 0)"
2876 using real_of_int_div[OF kdt]
2877 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2878 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2879 by (simp add: ti algebra_simps)
2880 also have "\<dots> = (?I ?tk (NEq (CN 0 c e)))"
2881 using nonzero_eq_divide_eq[OF knz',
2882 where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
2883 real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2884 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2886 finally have ?case . }
2887 ultimately show ?case by blast
2890 then have cp: "c > 0" and nb: "numbound0 e" by auto
2891 { assume kdc: "k dvd c"
2892 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2893 from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
2894 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2895 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) }
2897 { assume *: "\<not> k dvd c"
2898 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2899 from assms * have "?I (real x) (?s (Lt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k < 0)"
2900 using real_of_int_div[OF kdt]
2901 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2902 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2903 by (simp add: ti algebra_simps)
2904 also have "\<dots> = (?I ?tk (Lt (CN 0 c e)))"
2905 using pos_less_divide_eq[OF kpos,
2906 where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
2907 real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2908 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2910 finally have ?case . }
2911 ultimately show ?case by blast
2914 then have cp: "c > 0" and nb: "numbound0 e" by auto
2915 { assume kdc: "k dvd c"
2916 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2917 from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
2918 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2919 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) }
2921 { assume *: "\<not> k dvd c"
2922 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2923 from assms * have "?I (real x) (?s (Le (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<le> 0)"
2924 using real_of_int_div[OF kdt]
2925 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2926 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2927 by (simp add: ti algebra_simps)
2928 also have "\<dots> = (?I ?tk (Le (CN 0 c e)))"
2929 using pos_le_divide_eq[OF kpos,
2930 where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
2931 real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2932 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2934 finally have ?case . }
2935 ultimately show ?case by blast
2938 then have cp: "c > 0" and nb: "numbound0 e" by auto
2939 { assume kdc: "k dvd c"
2940 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2941 from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
2942 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2943 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) }
2945 { assume *: "\<not> k dvd c"
2946 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2947 from assms * have "?I (real x) (?s (Gt (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k > 0)"
2948 using real_of_int_div[OF kdt]
2949 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2950 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2951 by (simp add: ti algebra_simps)
2952 also have "\<dots> = (?I ?tk (Gt (CN 0 c e)))"
2953 using pos_divide_less_eq[OF kpos,
2954 where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
2955 real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2956 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2958 finally have ?case . }
2959 ultimately show ?case by blast
2962 then have cp: "c > 0" and nb: "numbound0 e" by auto
2963 { assume kdc: "k dvd c"
2964 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2965 from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
2966 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2967 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) }
2969 { assume *: "\<not> k dvd c"
2970 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2971 from assms * have "?I (real x) (?s (Ge (CN 0 c e))) = ((real c * (?N (real x) t / real k) + ?N (real x) e)* real k \<ge> 0)"
2972 using real_of_int_div[OF kdt]
2973 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2974 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2975 by (simp add: ti algebra_simps)
2976 also have "\<dots> = (?I ?tk (Ge (CN 0 c e)))"
2977 using pos_divide_le_eq[OF kpos,
2978 where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric]
2979 real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2980 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
2982 finally have ?case . }
2983 ultimately show ?case by blast
2986 then have cp: "c > 0" and nb: "numbound0 e" by auto
2987 { assume kdc: "k dvd c"
2988 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2989 from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
2990 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2991 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) }
2993 { assume *: "\<not> k dvd c"
2994 from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
2995 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
2996 from assms * have "?I (real x) (?s (Dvd i (CN 0 c e))) = (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k)"
2997 using real_of_int_div[OF kdt]
2998 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
2999 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
3000 by (simp add: ti algebra_simps)
3001 also have "\<dots> = (?I ?tk (Dvd i (CN 0 c e)))"
3002 using rdvd_mult[OF knz, where n="i"]
3003 real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
3004 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
3006 finally have ?case . }
3007 ultimately show ?case by blast
3010 then have cp: "c > 0" and nb: "numbound0 e" by auto
3011 { assume kdc: "k dvd c"
3012 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
3013 from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt]
3014 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
3015 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) }
3017 { assume *: "\<not> k dvd c"
3018 from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp
3019 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp
3020 from assms * have "?I (real x) (?s (NDvd i (CN 0 c e))) = (\<not> (real i * real k rdvd (real c * (?N (real x) t / real k) + ?N (real x) e)* real k))"
3021 using real_of_int_div[OF kdt]
3022 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
3023 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
3024 by (simp add: ti algebra_simps)
3025 also have "\<dots> = (?I ?tk (NDvd i (CN 0 c e)))"
3026 using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF kdt]
3027 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"]
3028 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"]
3030 finally have ?case . }
3031 ultimately show ?case by blast
3032 qed (simp_all add: bound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"]
3033 numbound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"])
3036 lemma \<sigma>_\<rho>_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t"
3037 shows "bound0 (\<sigma>_\<rho> p (t,k))"
3039 by (induct p rule: iszlfm.induct, auto simp add: nb)
3042 assumes lp: "iszlfm p (real (i::int)#bs)"
3043 shows "\<forall> (b,k) \<in> set (\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)"
3044 using lp by (induct p rule: \<rho>.induct, auto simp add: isint_sub isint_neg)
3046 lemma \<alpha>_\<rho>_l:
3047 assumes lp: "iszlfm p (real (i::int)#bs)"
3048 shows "\<forall> (b,k) \<in> set (\<alpha>_\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)"
3049 using lp isint_add [OF isint_c[where j="- 1"],where bs="real i#bs"]
3050 by (induct p rule: \<alpha>_\<rho>.induct, auto)
3052 lemma \<rho>: assumes lp: "iszlfm p (real (i::int) #bs)"
3053 and pi: "Ifm (real i#bs) p"
3054 and d: "d_\<delta> p d"
3056 and nob: "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> Inum (real i#bs) e + real j"
3057 (is "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. _ \<noteq> ?N i e + _")
3058 shows "Ifm (real(i - d)#bs) p"
3060 proof(induct p rule: iszlfm.induct)
3061 case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
3062 and pi: "real (c*i) = - 1 - ?N i e + real (1::int)" and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> -1 - ?N i e + real j"
3064 from mult_strict_left_mono[OF dp cp] have one:"1 \<in> {1 .. c*d}" by auto
3065 from nob[rule_format, where j="1", OF one] pi show ?case by simp
3068 hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
3069 and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
3071 {assume "real (c*i) \<noteq> - ?N i e + real (c*d)"
3072 with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"]
3073 have ?case by (simp add: algebra_simps)}
3075 {assume pi: "real (c*i) = - ?N i e + real (c*d)"
3076 from mult_strict_left_mono[OF dp cp] have d: "(c*d) \<in> {1 .. c*d}" by simp
3077 from nob[rule_format, where j="c*d", OF d] pi have ?case by simp }
3078 ultimately show ?case by blast
3080 case (5 c e) hence cp: "c > 0" by simp
3081 from 5 mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric]
3083 show ?case using 5 dp
3084 by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"]
3087 case (6 c e) hence cp: "c > 0" by simp
3088 from 6 mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric]
3090 show ?case using 6 dp
3091 by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"]
3094 case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
3095 and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j"
3096 and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0"
3098 let ?fe = "floor (?N i e)"
3099 from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: algebra_simps)
3100 from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp
3101 hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric])
3102 have "real (c*i) + ?N i e > real (c*d) \<or> real (c*i) + ?N i e \<le> real (c*d)" by auto
3104 {assume "real (c*i) + ?N i e > real (c*d)" hence ?case
3105 by (simp add: algebra_simps
3106 numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])}
3108 {assume H:"real (c*i) + ?N i e \<le> real (c*d)"
3109 with ei[simplified isint_iff] have "real (c*i + ?fe) \<le> real (c*d)" by simp
3110 hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
3111 with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto
3112 hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1"
3113 by (simp only: diff_minus[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps)
3114 with nob have ?case by blast }
3115 ultimately show ?case by blast
3117 case (8 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)"
3118 and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - 1 - ?N i e + real j"
3119 and pi: "real (c*i) + ?N i e \<ge> 0" and cp': "real c >0"
3121 let ?fe = "floor (?N i e)"
3122 from pi cp have th:"(real i +?N i e / real c)*real c \<ge> 0" by (simp add: algebra_simps)
3123 from pi ei[simplified isint_iff] have "real (c*i + ?fe) \<ge> real (0::int)" by simp
3124 hence pi': "c*i + 1 + ?fe \<ge> 1" by (simp only: real_of_int_le_iff[symmetric])
3125 have "real (c*i) + ?N i e \<ge> real (c*d) \<or> real (c*i) + ?N i e < real (c*d)" by auto
3127 {assume "real (c*i) + ?N i e \<ge> real (c*d)" hence ?case
3128 by (simp add: algebra_simps
3129 numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])}
3131 {assume H:"real (c*i) + ?N i e < real (c*d)"
3132 with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp
3133 hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff)
3134 with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto
3135 hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1"
3136 by (simp only: diff_minus[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps real_of_one)
3137 hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1"
3138 by (simp only: algebra_simps diff_minus[symmetric])
3139 hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1"
3140 by (simp only: add_ac diff_minus)
3141 with nob have ?case by blast }
3142 ultimately show ?case by blast
3144 case (9 j c e) hence p: "real j rdvd real (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+
3145 let ?e = "Inum (real i # bs) e"
3146 from 9 have "isint e (real i #bs)" by simp
3147 hence ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"]
3148 by (simp add: isint_iff)
3149 from 9 have id: "j dvd d" by simp
3150 from ie[symmetric] have "?p i = (real j rdvd real (c*i+ floor ?e))" by simp
3151 also have "\<dots> = (j dvd c*i + floor ?e)"
3152 using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp
3153 also have "\<dots> = (j dvd c*i - c*d + floor ?e)"
3154 using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp
3155 also have "\<dots> = (real j rdvd real (c*i - c*d + floor ?e))"
3156 using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
3158 also have "\<dots> = (real j rdvd real (c*(i - d)) + ?e)"
3159 using ie by (simp add:algebra_simps)
3161 using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p
3162 by (simp add: algebra_simps)
3165 hence p: "\<not> (real j rdvd real (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e"
3167 let ?e = "Inum (real i # bs) e"
3168 from 10 have "isint e (real i #bs)" by simp
3169 hence ie: "real (floor ?e) = ?e"
3170 using isint_iff[where n="e" and bs="(real i)#bs"] numbound0_I[OF bn,where b="real i" and b'="real i" and bs="bs"]
3171 by (simp add: isint_iff)
3172 from 10 have id: "j dvd d" by simp
3173 from ie[symmetric] have "?p i = (\<not> (real j rdvd real (c*i+ floor ?e)))" by simp
3174 also have "\<dots> = Not (j dvd c*i + floor ?e)"
3175 using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp
3176 also have "\<dots> = Not (j dvd c*i - c*d + floor ?e)"
3177 using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp
3178 also have "\<dots> = Not (real j rdvd real (c*i - c*d + floor ?e))"
3179 using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified]
3181 also have "\<dots> = Not (real j rdvd real (c*(i - d)) + ?e)"
3182 using ie by (simp add:algebra_simps)
3184 using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p
3185 by (simp add: algebra_simps)
3186 qed (auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"])
3188 lemma \<sigma>_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t"
3189 shows "bound0 (\<sigma> p k t)"
3190 using \<sigma>_\<rho>_nb[OF lp nb] nb by (simp add: \<sigma>_def)
3192 lemma \<rho>': assumes lp: "iszlfm p (a #bs)"
3193 and d: "d_\<delta> p d"
3195 shows "\<forall> x. \<not>(\<exists> (e,c) \<in> set(\<rho> p). \<exists>(j::int) \<in> {1 .. c*d}. Ifm (a #bs) (\<sigma> p c (Add e (C j)))) \<longrightarrow> Ifm (real x#bs) p \<longrightarrow> Ifm (real (x - d)#bs) p" (is "\<forall> x. ?b x \<longrightarrow> ?P x \<longrightarrow> ?P (x - d)")
3198 assume nob1:"?b x" and px: "?P x"
3199 from iszlfm_gen[OF lp, rule_format, where y="real x"] have lp': "iszlfm p (real x#bs)".
3200 have nob: "\<forall>(e, c)\<in>set (\<rho> p). \<forall>j\<in>{1..c * d}. real (c * x) \<noteq> Inum (real x # bs) e + real j"
3202 fix e c j assume ecR: "(e,c) \<in> set (\<rho> p)" and jD: "j\<in> {1 .. c*d}"
3203 and cx: "real (c*x) = Inum (real x#bs) e + real j"
3204 let ?e = "Inum (real x#bs) e"
3205 let ?fe = "floor ?e"
3206 from \<rho>_l[OF lp'] ecR have ei:"isint e (real x#bs)" and cp:"c>0" and nb:"numbound0 e"
3208 from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" .
3209 from cx ei[simplified isint_iff] have "real (c*x) = real (?fe + j)" by simp
3210 hence cx: "c*x = ?fe + j" by (simp only: real_of_int_inject)
3211 hence cdej:"c dvd ?fe + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp)
3212 hence "real c rdvd real (?fe + j)" by (simp only: int_rdvd_iff)
3213 hence rcdej: "real c rdvd ?e + real j" by (simp add: ei[simplified isint_iff])
3214 from cx have "(c*x) div c = (?fe + j) div c" by simp
3215 with cp have "x = (?fe + j) div c" by simp
3216 with px have th: "?P ((?fe + j) div c)" by auto
3217 from cp have cp': "real c > 0" by simp
3218 from cdej have cdej': "c dvd floor (Inum (real x#bs) (Add e (C j)))" by simp
3219 from nb have nb': "numbound0 (Add e (C j))" by simp
3220 have ji: "isint (C j) (real x#bs)" by (simp add: isint_def)
3221 from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real x#bs)" .
3222 from th \<sigma>_\<rho>[where b'="real x", OF lp' cp' nb' ei' cdej',symmetric]
3223 have "Ifm (real x#bs) (\<sigma>_\<rho> p (Add e (C j), c))" by simp
3224 with rcdej have th: "Ifm (real x#bs) (\<sigma> p c (Add e (C j)))" by (simp add: \<sigma>_def)
3225 from th bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"],where bs="bs" and b="real x" and b'="a"]
3226 have "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" by blast
3227 with ecR jD nob1 show "False" by blast
3229 from \<rho>[OF lp' px d dp nob] show "?P (x -d )" .
3234 assumes lp: "iszlfm p (real (i::int)#bs)"
3235 shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. \<delta> p}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> (e,c) \<in> set (\<rho> p). \<exists> j\<in> {1 .. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j)))))"
3236 (is "(\<exists>(x::int). ?P x) = ((\<exists> j\<in> {1.. \<delta> p}. ?MP j)\<or>(\<exists> (e,c) \<in> ?R. \<exists> j\<in> _. ?SP c e j))"
3237 is "?lhs = (?MD \<or> ?RD)" is "?lhs = ?rhs")
3239 let ?d= "\<delta> p"
3240 from \<delta>[OF lp] have d:"d_\<delta> p ?d" and dp: "?d > 0" by auto
3241 { assume H:"?MD" hence th:"\<exists> (x::int). ?MP x" by blast
3242 from H minusinf_ex[OF lp th] have ?thesis by blast}
3244 { fix e c j assume exR:"(e,c) \<in> ?R" and jD:"j\<in> {1 .. c*?d}" and spx:"?SP c e j"
3245 from exR \<rho>_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real i#bs)" and cp: "c > 0"
3247 have "isint (C j) (real i#bs)" by (simp add: isint_iff)
3248 with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real i"]]
3249 have eji:"isint (Add e (C j)) (real i#bs)" by simp
3250 from nb have nb': "numbound0 (Add e (C j))" by simp
3251 from spx bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real i"]
3252 have spx': "Ifm (real i # bs) (\<sigma> p c (Add e (C j)))" by blast
3253 from spx' have rcdej:"real c rdvd (Inum (real i#bs) (Add e (C j)))"
3254 and sr:"Ifm (real i#bs) (\<sigma>_\<rho> p (Add e (C j),c))" by (simp add: \<sigma>_def)+
3255 from rcdej eji[simplified isint_iff]
3256 have "real c rdvd real (floor (Inum (real i#bs) (Add e (C j))))" by simp
3257 hence cdej:"c dvd floor (Inum (real i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff)
3258 from cp have cp': "real c > 0" by simp
3259 from \<sigma>_\<rho>[OF lp cp' nb' eji cdej] spx' have "?P (\<lfloor>Inum (real i # bs) (Add e (C j))\<rfloor> div c)"
3260 by (simp add: \<sigma>_def)
3262 with exR jD spx have ?thesis by blast}
3264 { fix x assume px: "?P x" and nob: "\<not> ?RD"
3265 from iszlfm_gen [OF lp,rule_format, where y="a"] have lp':"iszlfm p (a#bs)" .
3266 from \<rho>'[OF lp' d dp, rule_format, OF nob] have th:"\<forall> x. ?P x \<longrightarrow> ?P (x - ?d)" by blast
3267 from minusinf_inf[OF lp] obtain z where z:"\<forall> x<z. ?MP x = ?P x" by blast
3268 have zp: "abs (x - z) + 1 \<ge> 0" by arith
3269 from decr_lemma[OF dp,where x="x" and z="z"]
3270 decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\<exists> x. ?MP x" by auto
3271 with minusinf_bex[OF lp] px nob have ?thesis by blast}
3272 ultimately show ?thesis by blast
3275 lemma mirror_\<alpha>_\<rho>: assumes lp: "iszlfm p (a#bs)"
3276 shows "(\<lambda> (t,k). (Inum (a#bs) t, k)) ` set (\<alpha>_\<rho> p) = (\<lambda> (t,k). (Inum (a#bs) t,k)) ` set (\<rho> (mirror p))"
3278 by (induct p rule: mirror.induct, simp_all add: split_def image_Un )
3280 text {* The @{text "\<real>"} part*}
3282 text{* Linearity for fm where Bound 0 ranges over @{text "\<real>"}*}
3284 isrlfm :: "fm \<Rightarrow> bool" (* Linearity test for fm *)
3285 recdef isrlfm "measure size"
3286 "isrlfm (And p q) = (isrlfm p \<and> isrlfm q)"
3287 "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)"
3288 "isrlfm (Eq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
3289 "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)"
3290 "isrlfm (Lt (CN 0 c e)) = (c>0 \<and> numbound0 e)"
3291 "isrlfm (Le (CN 0 c e)) = (c>0 \<and> numbound0 e)"
3292 "isrlfm (Gt (CN 0 c e)) = (c>0 \<and> numbound0 e)"
3293 "isrlfm (Ge (CN 0 c e)) = (c>0 \<and> numbound0 e)"
3294 "isrlfm p = (isatom p \<and> (bound0 p))"
3296 definition fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm" where
3297 "fp p n s j \<equiv> (if n > 0 then
3298 (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j)))))
3299 (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1))))))))
3301 (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j)))))
3302 (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))"
3304 (* splits the bounded from the unbounded part*)
3305 function (sequential) rsplit0 :: "num \<Rightarrow> (fm \<times> int \<times> num) list" where
3306 "rsplit0 (Bound 0) = [(T,1,C 0)]"
3307 | "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b
3308 in map (\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s)) [(a,b). a\<leftarrow>acs,b\<leftarrow>bcs])"
3309 | "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))"
3310 | "rsplit0 (Neg a) = map (\<lambda> (p,n,s). (p,-n,Neg s)) (rsplit0 a)"
3311 | "rsplit0 (Floor a) = concat (map
3312 (\<lambda> (p,n,s). if n=0 then [(p,0,Floor s)]
3313 else (map (\<lambda> j. (fp p n s j, 0, Add (Floor s) (C j))) (if n > 0 then [0 .. n] else [n .. 0])))
3315 | "rsplit0 (CN 0 c a) = map (\<lambda> (p,n,s). (p,n+c,s)) (rsplit0 a)"
3316 | "rsplit0 (CN m c a) = map (\<lambda> (p,n,s). (p,n,CN m c s)) (rsplit0 a)"
3317 | "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)"
3318 | "rsplit0 (Mul c a) = map (\<lambda> (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)"
3319 | "rsplit0 t = [(T,0,t)]"
3320 by pat_completeness auto
3321 termination by (relation "measure num_size") auto
3323 lemma conj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (conj p q)"
3324 using conj_def by (cases p, auto)
3325 lemma disj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (disj p q)"
3326 using disj_def by (cases p, auto)
3330 shows "\<forall> (p,n,s) \<in> set (rsplit0 t).
3331 (Ifm (x#bs) p \<longrightarrow> (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p"
3332 (is "\<forall> (p,n,s) \<in> ?SS t. (?I p \<longrightarrow> ?N t = ?N (CN 0 n s)) \<and> _ \<and> _ ")
3333 proof(induct t rule: rsplit0.induct)
3335 let ?p = "\<lambda> (p,n,s) j. fp p n s j"
3336 let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))"
3337 let ?J = "\<lambda> n. if n>0 then [0..n] else [n..0]"
3338 let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))"
3339 have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
3340 have U1: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) =
3341 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)]))" by auto
3342 have U2': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0}.
3343 ?ff (p,n,s) = map (?f(p,n,s)) [0..n]" by auto
3344 hence U2: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) =
3345 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s).
3346 set (map (?f(p,n,s)) [0..n])))"
3348 fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
3349 assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
3350 thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
3351 by (auto simp add: split_def)
3353 have U3': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0}. ?ff (p,n,s) = map (?f(p,n,s)) [n..0]"
3355 hence U3: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) =
3356 (UNION {(p,n,s). (p,n,s)\<in> ?SS a\<and>n<0} (\<lambda>(p,n,s). set (map (?f(p,n,s)) [n..0])))"
3358 fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
3359 assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
3360 thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
3361 by (auto simp add: split_def)
3363 have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))"
3365 also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by blast
3366 also have "\<dots> =
3367 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un
3368 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un
3369 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))"
3370 using int_cases[rule_format] by blast
3371 also have "\<dots> =
3372 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un
3373 (UNION {(p,n,s). (p,n,s)\<in> ?SS a\<and>n>0} (\<lambda>(p,n,s). set(map(?f(p,n,s)) [0..n]))) Un
3374 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s).
3375 set (map (?f(p,n,s)) [n..0]))))" by (simp only: U1 U2 U3)
3376 also have "\<dots> =
3377 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
3378 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un
3379 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
3380 by (simp only: set_map set_upto set.simps)
3381 also have "\<dots> =
3382 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
3383 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un
3384 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {n .. 0}})))" by blast
3386 have FS: "?SS (Floor a) =
3387 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
3388 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un
3389 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {n .. 0}})))" by blast
3391 proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -)
3393 let ?ths = "(?I p \<longrightarrow> (?N (Floor a) = ?N (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p"
3394 assume "(\<exists>ba. (p, 0, ba) \<in> set (rsplit0 a) \<and> n = 0 \<and> s = Floor ba) \<or>
3396 (ab, ac, ba) \<in> set (rsplit0 a) \<and>
3398 (\<exists>j. p = fp ab ac ba j \<and>
3399 n = 0 \<and> s = Add (Floor ba) (C j) \<and> 0 \<le> j \<and> j \<le> ac)) \<or>
3401 (ab, ac, ba) \<in> set (rsplit0 a) \<and>
3403 (\<exists>j. p = fp ab ac ba j \<and>
3404 n = 0 \<and> s = Add (Floor ba) (C j) \<and> ac \<le> j \<and> j \<le> 0))"
3407 assume "(p, 0, s') \<in> ?SS a" and "n = 0" and "s = Floor s'"
3408 hence ?ths using 5(1) by auto }
3411 assume pns: "(p', n', s') \<in> ?SS a"
3413 and p_def: "p = ?p (p',n',s') j"
3415 and s_def: "s = (Add (Floor s') (C j))"
3416 and jp: "0 \<le> j" and jn: "j \<le> n'"
3417 from 5 pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
3418 Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
3419 numbound0 s' \<and> isrlfm p'" by blast
3420 hence nb: "numbound0 s'" by simp
3421 from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numsub_nb)
3422 let ?nxs = "CN 0 n' s'"
3423 let ?l = "floor (?N s') + j"
3425 have "?I (?p (p',n',s') j) \<longrightarrow>
3426 (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))"
3427 by (simp add: fp_def np algebra_simps numsub numadd numfloor)
3428 also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
3429 using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
3431 have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp
3432 ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
3434 with s_def n0 p_def nb nf have ?ths by auto}
3437 assume pns: "(p', n', s') \<in> ?SS a"
3439 and p_def: "p = ?p (p',n',s') j"
3441 and s_def: "s = (Add (Floor s') (C j))"
3442 and jp: "n' \<le> j" and jn: "j \<le> 0"
3443 from 5 pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow>
3444 Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and>
3445 numbound0 s' \<and> isrlfm p'" by blast
3446 hence nb: "numbound0 s'" by simp
3447 from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numneg_nb)
3448 let ?nxs = "CN 0 n' s'"
3449 let ?l = "floor (?N s') + j"
3451 have "?I (?p (p',n',s') j) \<longrightarrow>
3452 (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))"
3453 by (simp add: np fp_def algebra_simps numneg numfloor numadd numsub)
3454 also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))"
3455 using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp
3457 have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp
3458 ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))"
3460 with s_def n0 p_def nb nf have ?ths by auto}
3461 ultimately show ?ths by auto
3464 case (3 a b) then show ?case
3466 apply (erule_tac x = "(aa, aaa, ba)" in ballE) apply simp_all
3467 apply (erule_tac x = "(ab, ac, baa)" in ballE) apply simp_all
3469 qed (auto simp add: Let_def split_def algebra_simps conj_rl)
3471 lemma real_in_int_intervals:
3472 assumes xb: "real m \<le> x \<and> x < real ((n::int) + 1)"
3473 shows "\<exists> j\<in> {m.. n}. real j \<le> x \<and> x < real (j+1)" (is "\<exists> j\<in> ?N. ?P j")
3474 by (rule bexI[where P="?P" and x="floor x" and A="?N"])
3475 (auto simp add: floor_less_eq[where x="x" and a="n+1", simplified] xb[simplified] floor_mono[where x="real m" and y="x", OF conjunct1[OF xb], simplified floor_real_of_int[where n="m"]])
3477 lemma rsplit0_complete:
3478 assumes xp:"0 \<le> x" and x1:"x < 1"
3479 shows "\<exists> (p,n,s) \<in> set (rsplit0 t). Ifm (x#bs) p" (is "\<exists> (p,n,s) \<in> ?SS t. ?I p")
3480 proof(induct t rule: rsplit0.induct)
3482 then have "\<exists> (pa,na,sa) \<in> ?SS a. ?I pa" by auto
3483 then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\<in> ?SS a \<and> ?I pa" by blast
3484 with 2 have "\<exists> (pb,nb,sb) \<in> ?SS b. ?I pb" by blast
3485 then obtain "pb" "nb" "sb" where pb: "(pb,nb,sb)\<in> ?SS b \<and> ?I pb" by blast
3486 from pa pb have th: "((pa,na,sa),(pb,nb,sb)) \<in> set[(x,y). x\<leftarrow>rsplit0 a, y\<leftarrow>rsplit0 b]"
3488 let ?f="(\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s))"
3489 from imageI[OF th, where f="?f"] have "?f ((pa,na,sa),(pb,nb,sb)) \<in> ?SS (Add a b)"
3490 by (simp add: Let_def)
3491 hence "(And pa pb, na +nb, Add sa sb) \<in> ?SS (Add a b)" by simp
3492 moreover from pa pb have "?I (And pa pb)" by simp
3493 ultimately show ?case by blast
3496 let ?p = "\<lambda> (p,n,s) j. fp p n s j"
3497 let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))"
3498 let ?J = "\<lambda> n. if n>0 then [0..n] else [n..0]"
3499 let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))"
3500 have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith
3501 have U1: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)]))" by auto
3502 have U2': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0}. ?ff (p,n,s) = map (?f(p,n,s)) [0..n]"
3504 hence U2: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) [0..n])))"
3506 fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
3507 assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
3508 thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
3509 by (auto simp add: split_def)
3511 have U3': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0}. ?ff (p,n,s) = map (?f(p,n,s)) [n..0]"
3513 hence U3: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) [n..0])))"
3515 fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g
3516 assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c"
3517 thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))"
3518 by (auto simp add: split_def)
3521 have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))" by auto
3522 also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by blast
3523 also have "\<dots> =
3524 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un
3525 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un
3526 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))"
3527 using int_cases[rule_format] by blast
3528 also have "\<dots> =
3529 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un
3530 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) [0..n]))) Un
3531 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) [n..0]))))" by (simp only: U1 U2 U3)
3532 also have "\<dots> =
3533 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
3534 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un
3535 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))"
3536 by (simp only: set_map set_upto set.simps)
3537 also have "\<dots> =
3538 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
3539 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un
3540 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {n .. 0}})))" by blast
3542 have FS: "?SS (Floor a) =
3543 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un
3544 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {0 .. n}})) Un
3545 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). {?f(p,n,s) j| j. j\<in> {n .. 0}})))"
3547 from 5 have "\<exists> (p,n,s) \<in> ?SS a. ?I p" by auto
3548 then obtain "p" "n" "s" where pns: "(p,n,s) \<in> ?SS a \<and> ?I p" by blast
3549 let ?N = "\<lambda> t. Inum (x#bs) t"
3550 from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \<and> numbound0 s \<and> isrlfm p"
3553 have "n=0 \<or> n >0 \<or> n <0" by arith
3554 moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto }
3558 from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) \<le> ?N s" by simp
3559 also from mult_left_mono[OF xp] np have "?N s \<le> real n * x + ?N s" by simp
3560 finally have "?N (Floor s) \<le> real n * x + ?N s" .
3562 {from x1 np have "real n *x + ?N s < real n + ?N s" by simp
3563 also from real_of_int_floor_add_one_gt[where r="?N s"]
3564 have "\<dots> < real n + ?N (Floor s) + 1" by simp
3565 finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp}
3566 ultimately have "?N (Floor s) \<le> real n *x + ?N s\<and> real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp
3567 hence th: "0 \<le> real n *x + ?N s - ?N (Floor s) \<and> real n *x + ?N s - ?N (Floor s) < real (n+1)" by simp
3568 from real_in_int_intervals th have "\<exists> j\<in> {0 .. n}. real j \<le> real n *x + ?N s - ?N (Floor s)\<and> real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp
3570 hence "\<exists> j\<in> {0 .. n}. 0 \<le> real n *x + ?N s - ?N (Floor s) - real j \<and> real n *x + ?N s - ?N (Floor s) - real (j+1) < 0"
3571 by(simp only: myle[of _ "real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"])
3572 hence "\<exists> j\<in> {0.. n}. ?I (?p (p,n,s) j)"
3573 using pns by (simp add: fp_def np algebra_simps numsub numadd)
3574 then obtain "j" where j_def: "j\<in> {0 .. n} \<and> ?I (?p (p,n,s) j)" by blast
3575 hence "\<exists>x \<in> {?p (p,n,s) j |j. 0\<le> j \<and> j \<le> n }. ?I x" by auto
3576 hence ?case using pns
3577 by (simp only: FS,simp add: bex_Un)
3578 (rule disjI2, rule disjI1,rule exI [where x="p"],
3579 rule exI [where x="n"],rule exI [where x="s"],simp_all add: np)
3582 { assume nn: "n < 0" hence np: "-n >0" by simp
3583 from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) + 1 > ?N s" by simp
3584 moreover from mult_left_mono_neg[OF xp] nn have "?N s \<ge> real n * x + ?N s" by simp
3585 ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith
3587 {from x1 nn have "real n *x + ?N s \<ge> real n + ?N s" by simp
3588 moreover from real_of_int_floor_le[where r="?N s"] have "real n + ?N s \<ge> real n + ?N (Floor s)" by simp
3589 ultimately have "real n *x + ?N s \<ge> ?N (Floor s) + real n"
3590 by (simp only: algebra_simps)}
3591 ultimately have "?N (Floor s) + real n \<le> real n *x + ?N s\<and> real n *x + ?N s < ?N (Floor s) + real (1::int)" by simp
3592 hence th: "real n \<le> real n *x + ?N s - ?N (Floor s) \<and> real n *x + ?N s - ?N (Floor s) < real (1::int)" by simp
3593 have th1: "\<forall> (a::real). (- a > 0) = (a < 0)" by auto
3594 have th2: "\<forall> (a::real). (0 \<ge> - a) = (a \<ge> 0)" by auto
3595 from real_in_int_intervals th have "\<exists> j\<in> {n .. 0}. real j \<le> real n *x + ?N s - ?N (Floor s)\<and> real n *x + ?N s - ?N (Floor s) < real (j+1)" by simp
3597 hence "\<exists> j\<in> {n .. 0}. 0 \<le> real n *x + ?N s - ?N (Floor s) - real j \<and> real n *x + ?N s - ?N (Floor s) - real (j+1) < 0"
3598 by(simp only: myle[of _ "real n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real n *x + ?N s - ?N (Floor s)"])
3599 hence "\<exists> j\<in> {n .. 0}. 0 \<ge> - (real n *x + ?N s - ?N (Floor s) - real j) \<and> - (real n *x + ?N s - ?N (Floor s) - real (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format])
3600 hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)"
3601 using pns by (simp add: fp_def nn diff_minus add_ac mult_ac numfloor numadd numneg
3602 del: diff_less_0_iff_less diff_le_0_iff_le)
3603 then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast
3604 hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto
3605 hence ?case using pns
3606 by (simp only: FS,simp add: bex_Un)
3607 (rule disjI2, rule disjI2,rule exI [where x="p"],
3608 rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn)
3610 ultimately show ?case by blast
3611 qed (auto simp add: Let_def split_def)
3613 (* Linearize a formula where Bound 0 ranges over [0,1) *)
3615 definition rsplit :: "(int \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm" where
3616 "rsplit f a \<equiv> foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) (rsplit0 a)) F"
3618 lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\<exists> x \<in> set xs. Ifm bs (f x))"
3619 by(induct xs, simp_all)
3621 lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\<forall> x \<in> set xs. Ifm bs (f x))"
3622 by(induct xs, simp_all)
3624 lemma foldr_disj_map_rlfm:
3625 assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)"
3626 and \<phi>: "\<forall> (\<phi>,n,s) \<in> set xs. numbound0 s \<and> isrlfm \<phi>"
3627 shows "isrlfm (foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) xs) F)"
3628 using lf \<phi> by (induct xs, auto)
3630 lemma rsplit_ex: "Ifm bs (rsplit f a) = (\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). Ifm bs (conj \<phi> (f n s)))"
3631 using foldr_disj_map[where xs="rsplit0 a"] rsplit_def by (simp add: split_def)
3633 lemma rsplit_l: assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)"
3634 shows "isrlfm (rsplit f a)"
3636 from rsplit0_cs[where t="a"] have th: "\<forall> (\<phi>,n,s) \<in> set (rsplit0 a). numbound0 s \<and> isrlfm \<phi>" by blast
3637 from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp
3641 assumes xp: "x \<ge> 0" and x1: "x < 1"
3642 and f: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> (Ifm (x#bs) (f n s) = Ifm (x#bs) (g a))"
3643 shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)"
3645 let ?I = "\<lambda>x p. Ifm (x#bs) p"
3646 let ?N = "\<lambda> x t. Inum (x#bs) t"
3647 assume "?I x (rsplit f a)"
3648 hence "\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). ?I x (And \<phi> (f n s))" using rsplit_ex by simp
3649 then obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and "?I x (And \<phi> (f n s))" by blast
3650 hence \<phi>: "?I x \<phi>" and fns: "?I x (f n s)" by auto
3651 from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \<phi>
3652 have th: "(?N x a = ?N x (CN 0 n s)) \<and> numbound0 s" by auto
3653 from f[rule_format, OF th] fns show "?I x (g a)" by simp
3655 let ?I = "\<lambda>x p. Ifm (x#bs) p"
3656 let ?N = "\<lambda> x t. Inum (x#bs) t"
3657 assume ga: "?I x (g a)"
3658 from rsplit0_complete[OF xp x1, where bs="bs" and t="a"]
3659 obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and fx: "?I x \<phi>" by blast
3660 from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx
3661 have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto
3662 with ga f have "?I x (f n s)" by auto
3663 with rsplit_ex fnsS fx show "?I x (rsplit f a)" by auto
3666 definition lt :: "int \<Rightarrow> num \<Rightarrow> fm" where
3667 lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t))
3668 else (Gt (CN 0 (-c) (Neg t))))"
3670 definition le :: "int \<Rightarrow> num \<Rightarrow> fm" where
3671 le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t))
3672 else (Ge (CN 0 (-c) (Neg t))))"
3674 definition gt :: "int \<Rightarrow> num \<Rightarrow> fm" where
3675 gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t))
3676 else (Lt (CN 0 (-c) (Neg t))))"
3678 definition ge :: "int \<Rightarrow> num \<Rightarrow> fm" where
3679 ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t))
3680 else (Le (CN 0 (-c) (Neg t))))"
3682 definition eq :: "int \<Rightarrow> num \<Rightarrow> fm" where
3683 eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t))
3684 else (Eq (CN 0 (-c) (Neg t))))"
3686 definition neq :: "int \<Rightarrow> num \<Rightarrow> fm" where
3687 neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t))
3688 else (NEq (CN 0 (-c) (Neg t))))"
3690 lemma lt_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (lt n s) = Ifm (x#bs) (Lt a)"
3691 (is "\<forall> a n s . ?N a = ?N (CN 0 n s) \<and> _\<longrightarrow> ?I (lt n s) = ?I (Lt a)")
3694 assume H: "?N a = ?N (CN 0 n s)"
3695 show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def))
3696 (cases "n > 0", simp_all add: lt_def algebra_simps myless[of _ "0"])
3699 lemma lt_l: "isrlfm (rsplit lt a)"
3700 by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def,
3701 case_tac s, simp_all, case_tac "nat", simp_all)
3703 lemma le_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (le n s) = Ifm (x#bs) (Le a)" (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (le n s) = ?I (Le a)")
3706 assume H: "?N a = ?N (CN 0 n s)"
3707 show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def))
3708 (cases "n > 0", simp_all add: le_def algebra_simps myle[of _ "0"])
3711 lemma le_l: "isrlfm (rsplit le a)"
3712 by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def)
3713 (case_tac s, simp_all, case_tac "nat",simp_all)
3715 lemma gt_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (gt n s) = Ifm (x#bs) (Gt a)" (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (gt n s) = ?I (Gt a)")
3718 assume H: "?N a = ?N (CN 0 n s)"
3719 show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def))
3720 (cases "n > 0", simp_all add: gt_def algebra_simps myless[of _ "0"])
3722 lemma gt_l: "isrlfm (rsplit gt a)"
3723 by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def)
3724 (case_tac s, simp_all, case_tac "nat", simp_all)
3726 lemma ge_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (ge n s) = Ifm (x#bs) (Ge a)" (is "\<forall> a n s . ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (ge n s) = ?I (Ge a)")
3729 assume H: "?N a = ?N (CN 0 n s)"
3730 show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def))
3731 (cases "n > 0", simp_all add: ge_def algebra_simps myle[of _ "0"])
3733 lemma ge_l: "isrlfm (rsplit ge a)"
3734 by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def)
3735 (case_tac s, simp_all, case_tac "nat", simp_all)
3737 lemma eq_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (eq n s) = Ifm (x#bs) (Eq a)" (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (eq n s) = ?I (Eq a)")
3740 assume H: "?N a = ?N (CN 0 n s)"
3741 show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def algebra_simps)
3743 lemma eq_l: "isrlfm (rsplit eq a)"
3744 by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def)
3745 (case_tac s, simp_all, case_tac"nat", simp_all)
3747 lemma neq_mono: "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (neq n s) = Ifm (x#bs) (NEq a)" (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (neq n s) = ?I (NEq a)")
3750 assume H: "?N a = ?N (CN 0 n s)"
3751 show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def algebra_simps)
3754 lemma neq_l: "isrlfm (rsplit neq a)"
3755 by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def)
3756 (case_tac s, simp_all, case_tac"nat", simp_all)
3759 assumes u0:"0 \<le> u" and u1: "u < 1"
3760 shows "(-u \<le> real (n::int)) = (0 \<le> n)"
3764 assumes u0:"0 \<le> u" and u1: "u < 1"
3765 shows "(real (n::int) < real (m::int) - u) = (n < m)"
3769 assumes up: "u \<ge> 0" and u1: "u<1" and np: "real n > 0"
3770 shows "(real (i::int) rdvd real (n::int) * u - s) = (\<exists> j\<in> {0 .. n - 1}. real n * u = s - real (floor s) + real j \<and> real i rdvd real (j - floor s))" (is "?lhs = ?rhs")
3772 let ?ss = "s - real (floor s)"
3773 from real_of_int_floor_add_one_gt[where r="s", simplified myless[of "s"]]
3774 real_of_int_floor_le[where r="s"] have ss0:"?ss \<ge> 0" and ss1:"?ss < 1"
3775 by (auto simp add: myle[of _ "s", symmetric] myless[of "?ss"])
3776 from np have n0: "real n \<ge> 0" by simp
3777 from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np]
3778 have nu0:"real n * u - s \<ge> -s" and nun:"real n * u -s < real n - s" by auto
3779 from int_rdvd_real[where i="i" and x="real (n::int) * u - s"]
3780 have "real i rdvd real n * u - s =
3781 (i dvd floor (real n * u -s) \<and> (real (floor (real n * u - s)) = real n * u - s ))"
3782 (is "_ = (?DE)" is "_ = (?D \<and> ?E)") by simp
3783 also have "\<dots> = (?DE \<and> real(floor (real n * u - s) + floor s)\<ge> -?ss
3784 \<and> real(floor (real n * u - s) + floor s)< real n - ?ss)" (is "_=(?DE \<and>real ?a \<ge> _ \<and> real ?a < _)")
3785 using nu0 nun by auto
3786 also have "\<dots> = (?DE \<and> ?a \<ge> 0 \<and> ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1])
3787 also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. ?a = j))" by simp
3788 also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. real (\<lfloor>real n * u - s\<rfloor>) = real j - real \<lfloor>s\<rfloor> ))"
3789 by (simp only: algebra_simps real_of_int_diff[symmetric] real_of_int_inject del: real_of_int_diff)
3790 also have "\<dots> = ((\<exists> j\<in> {0 .. (n - 1)}. real n * u - s = real j - real \<lfloor>s\<rfloor> \<and> real i rdvd real n * u - s))" using int_rdvd_iff[where i="i" and t="\<lfloor>real n * u - s\<rfloor>"]
3791 by (auto cong: conj_cong)
3792 also have "\<dots> = ?rhs" by(simp cong: conj_cong) (simp add: algebra_simps )
3793 finally show ?thesis .
3797 DVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
3799 DVDJ_def: "DVDJ i n s = (foldr disj (map (\<lambda> j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))) [0..n - 1]) F)"
3802 NDVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
3804 NDVDJ_def: "NDVDJ i n s = (foldr conj (map (\<lambda> j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor (Neg s))))) [0..n - 1]) T)"
3807 assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0"
3808 shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))"
3810 let ?f = "\<lambda> j. conj (Eq(CN 0 n (Add s (Sub(Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))"
3811 let ?s= "Inum (x#bs) s"
3812 from foldr_disj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"]
3813 have "Ifm (x#bs) (DVDJ i n s) = (\<exists> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))"
3814 by (simp add: np DVDJ_def)
3815 also have "\<dots> = (\<exists> j\<in> {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \<and> real i rdvd real (j - floor (- ?s)))" by (simp add: algebra_simps diff_minus[symmetric])
3816 also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"]
3817 have "\<dots> = (real i rdvd real n * x - (-?s))" by simp
3818 finally show ?thesis by simp
3822 assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0"
3823 shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))"
3825 let ?f = "\<lambda> j. disj(NEq(CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor(Neg s))))"
3826 let ?s= "Inum (x#bs) s"
3827 from foldr_conj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"]
3828 have "Ifm (x#bs) (NDVDJ i n s) = (\<forall> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))"
3829 by (simp add: np NDVDJ_def)
3830 also have "\<dots> = (\<not> (\<exists> j\<in> {0 .. (n - 1)}. real n * x = (- ?s) - real (floor (- ?s)) + real j \<and> real i rdvd real (j - floor (- ?s))))" by (simp add: algebra_simps diff_minus[symmetric])
3831 also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"]
3832 have "\<dots> = (\<not> (real i rdvd real n * x - (-?s)))" by simp
3833 finally show ?thesis by simp
3836 lemma foldr_disj_map_rlfm2:
3837 assumes lf: "\<forall> n . isrlfm (f n)"
3838 shows "isrlfm (foldr disj (map f xs) F)"
3839 using lf by (induct xs, auto)
3840 lemma foldr_And_map_rlfm2:
3841 assumes lf: "\<forall> n . isrlfm (f n)"
3842 shows "isrlfm (foldr conj (map f xs) T)"
3843 using lf by (induct xs, auto)
3845 lemma DVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s"
3846 shows "isrlfm (DVDJ i n s)"
3848 let ?f="\<lambda>j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j)))))
3849 (Dvd i (Sub (C j) (Floor (Neg s))))"
3850 have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
3851 from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp
3854 lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s"
3855 shows "isrlfm (NDVDJ i n s)"
3857 let ?f="\<lambda>j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j)))))
3858 (NDvd i (Sub (C j) (Floor (Neg s))))"
3859 have th: "\<forall> j. isrlfm (?f j)" using nb np by auto
3860 from NDVDJ_def foldr_And_map_rlfm2[OF th] show ?thesis by auto
3863 definition DVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
3864 DVD_def: "DVD i c t =
3865 (if i=0 then eq c t else
3866 if c = 0 then (Dvd i t) else if c >0 then DVDJ (abs i) c t else DVDJ (abs i) (-c) (Neg t))"
3868 definition NDVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
3870 (if i=0 then neq c t else
3871 if c = 0 then (NDvd i t) else if c >0 then NDVDJ (abs i) c t else NDVDJ (abs i) (-c) (Neg t))"
3874 assumes xp: "0\<le> x" and x1: "x < 1"
3875 shows "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (DVD i n s) = Ifm (x#bs) (Dvd i a)"
3876 (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (DVD i n s) = ?I (Dvd i a)")
3879 assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s"
3880 let ?th = "?I (DVD i n s) = ?I (Dvd i a)"
3881 have "i=0 \<or> (i\<noteq>0 \<and> n=0) \<or> (i\<noteq>0 \<and> n < 0) \<or> (i\<noteq>0 \<and> n > 0)" by arith
3882 moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]]
3883 by (simp add: DVD_def rdvd_left_0_eq)}
3884 moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H DVD_def) }
3885 moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th
3886 by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1
3887 rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) }
3888 moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)}
3889 ultimately show ?th by blast
3892 lemma NDVD_mono: assumes xp: "0\<le> x" and x1: "x < 1"
3893 shows "\<forall> a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \<and> numbound0 s \<longrightarrow> Ifm (x#bs) (NDVD i n s) = Ifm (x#bs) (NDvd i a)"
3894 (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (NDVD i n s) = ?I (NDvd i a)")
3897 assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s"
3898 let ?th = "?I (NDVD i n s) = ?I (NDvd i a)"
3899 have "i=0 \<or> (i\<noteq>0 \<and> n=0) \<or> (i\<noteq>0 \<and> n < 0) \<or> (i\<noteq>0 \<and> n > 0)" by arith
3900 moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]]
3901 by (simp add: NDVD_def rdvd_left_0_eq)}
3902 moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H NDVD_def) }
3903 moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th
3904 by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1
3905 rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) }
3906 moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th
3907 by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)}
3908 ultimately show ?th by blast
3911 lemma DVD_l: "isrlfm (rsplit (DVD i) a)"
3912 by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l)
3913 (case_tac s, simp_all, case_tac "nat", simp_all)
3915 lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)"
3916 by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l)
3917 (case_tac s, simp_all, case_tac "nat", simp_all)
3919 consts rlfm :: "fm \<Rightarrow> fm"
3920 recdef rlfm "measure fmsize"
3921 "rlfm (And p q) = conj (rlfm p) (rlfm q)"
3922 "rlfm (Or p q) = disj (rlfm p) (rlfm q)"
3923 "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)"
3924 "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))"
3925 "rlfm (Lt a) = rsplit lt a"
3926 "rlfm (Le a) = rsplit le a"
3927 "rlfm (Gt a) = rsplit gt a"
3928 "rlfm (Ge a) = rsplit ge a"
3929 "rlfm (Eq a) = rsplit eq a"
3930 "rlfm (NEq a) = rsplit neq a"
3931 "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a"
3932 "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a"
3933 "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))"
3934 "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))"
3935 "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))"
3936 "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))"
3937 "rlfm (NOT (NOT p)) = rlfm p"
3940 "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))"
3941 "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))"
3942 "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))"
3943 "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))"
3944 "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))"
3945 "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))"
3946 "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))"
3947 "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))"
3948 "rlfm p = p" (hints simp add: fmsize_pos)
3950 lemma bound0at_l : "\<lbrakk>isatom p ; bound0 p\<rbrakk> \<Longrightarrow> isrlfm p"
3951 by (induct p rule: isrlfm.induct, auto)
3953 lemma simpfm_rl: "isrlfm p \<Longrightarrow> isrlfm (simpfm p)"
3956 hence "bound0 (Lt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
3957 by (cases a,simp_all, case_tac "nat", simp_all)
3959 {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))"
3960 using simpfm_bound0 by blast
3961 have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def)
3962 with bn bound0at_l have ?case by blast}
3964 { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
3965 { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
3966 with numgcd_pos[where t="CN 0 c (simpnum e)"]
3967 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
3968 from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
3969 by (simp add: numgcd_def)
3970 from `c > 0` have th': "c\<noteq>0" by auto
3971 from `c > 0` have cp: "c \<ge> 0" by simp
3972 from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
3973 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
3975 with Lt a have ?case
3976 by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)}
3977 ultimately show ?case by blast
3980 hence "bound0 (Le a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
3981 by (cases a,simp_all, case_tac "nat", simp_all)
3983 { assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))"
3984 using simpfm_bound0 by blast
3985 have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def)
3986 with bn bound0at_l have ?case by blast}
3988 { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
3989 { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
3990 with numgcd_pos[where t="CN 0 c (simpnum e)"]
3991 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
3992 from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
3993 by (simp add: numgcd_def)
3994 from `c > 0` have th': "c\<noteq>0" by auto
3995 from `c > 0` have cp: "c \<ge> 0" by simp
3996 from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
3997 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
3999 with Le a have ?case
4000 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
4001 ultimately show ?case by blast
4004 hence "bound0 (Gt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
4005 by (cases a,simp_all, case_tac "nat", simp_all)
4007 {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))"
4008 using simpfm_bound0 by blast
4009 have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def)
4010 with bn bound0at_l have ?case by blast}
4012 { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
4013 { assume cn1: "numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
4014 with numgcd_pos[where t="CN 0 c (simpnum e)"]
4015 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
4016 from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
4017 by (simp add: numgcd_def)
4018 from `c > 0` have th': "c\<noteq>0" by auto
4019 from `c > 0` have cp: "c \<ge> 0" by simp
4020 from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
4021 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
4023 with Gt a have ?case
4024 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
4025 ultimately show ?case by blast
4028 hence "bound0 (Ge a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
4029 by (cases a,simp_all, case_tac "nat", simp_all)
4031 { assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))"
4032 using simpfm_bound0 by blast
4033 have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def)
4034 with bn bound0at_l have ?case by blast}
4036 { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
4037 { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
4038 with numgcd_pos[where t="CN 0 c (simpnum e)"]
4039 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
4040 from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
4041 by (simp add: numgcd_def)
4042 from `c > 0` have th': "c\<noteq>0" by auto
4043 from `c > 0` have cp: "c \<ge> 0" by simp
4044 from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
4045 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
4047 with Ge a have ?case
4048 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
4049 ultimately show ?case by blast
4052 hence "bound0 (Eq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
4053 by (cases a,simp_all, case_tac "nat", simp_all)
4055 { assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))"
4056 using simpfm_bound0 by blast
4057 have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def)
4058 with bn bound0at_l have ?case by blast}
4060 { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
4061 { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
4062 with numgcd_pos[where t="CN 0 c (simpnum e)"]
4063 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
4064 from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
4065 by (simp add: numgcd_def)
4066 from `c > 0` have th': "c\<noteq>0" by auto
4067 from `c > 0` have cp: "c \<ge> 0" by simp
4068 from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
4069 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
4071 with Eq a have ?case
4072 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
4073 ultimately show ?case by blast
4076 hence "bound0 (NEq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)"
4077 by (cases a,simp_all, case_tac "nat", simp_all)
4079 {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))"
4080 using simpfm_bound0 by blast
4081 have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def)
4082 with bn bound0at_l have ?case by blast}
4084 { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e"
4085 { assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0"
4086 with numgcd_pos[where t="CN 0 c (simpnum e)"]
4087 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp
4088 from `c > 0` have th:"numgcd (CN 0 c (simpnum e)) \<le> c"
4089 by (simp add: numgcd_def)
4090 from `c > 0` have th': "c\<noteq>0" by auto
4091 from `c > 0` have cp: "c \<ge> 0" by simp
4092 from zdiv_mono2[OF cp th1 th, simplified div_self[OF th']]
4093 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp
4095 with NEq a have ?case
4096 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)}
4097 ultimately show ?case by blast
4099 case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))"
4100 using simpfm_bound0 by blast
4101 have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def)
4102 with bn bound0at_l show ?case by blast
4104 case (NDvd i a) hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))"
4105 using simpfm_bound0 by blast
4106 have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def)
4107 with bn bound0at_l show ?case by blast
4108 qed(auto simp add: conj_def imp_def disj_def iff_def Let_def simpfm_bound0 numadd_nb numneg_nb)
4111 assumes qfp: "qfree p"
4112 and xp: "0 \<le> x" and x1: "x < 1"
4113 shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \<and> isrlfm (rlfm p)"
4115 by (induct p rule: rlfm.induct)
4116 (auto simp add: rsplit[OF xp x1 lt_mono] lt_l rsplit[OF xp x1 le_mono] le_l rsplit[OF xp x1 gt_mono] gt_l
4117 rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l
4118 rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl)
4120 assumes qfp: "qfree p"
4121 shows "isrlfm (rlfm p)"
4122 using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l
4123 by (induct p rule: rlfm.induct,auto simp add: simpfm_rl)
4125 (* Operations needed for Ferrante and Rackoff *)
4126 lemma rminusinf_inf:
4127 assumes lp: "isrlfm p"
4128 shows "\<exists> z. \<forall> x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p")
4130 proof (induct p rule: minusinf.induct)
4131 case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto
4133 case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto
4136 from 3 have nb: "numbound0 e" by simp
4137 from 3 have cp: "real c > 0" by simp
4139 let ?e="Inum (a#bs) e"
4140 let ?z = "(- ?e) / real c"
4143 hence "(real c * x < - ?e)"
4144 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac)
4145 hence "real c * x + ?e < 0" by arith
4146 hence "real c * x + ?e \<noteq> 0" by simp
4147 with xz have "?P ?z x (Eq (CN 0 c e))"
4148 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4149 hence "\<forall> x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp
4153 from 4 have nb: "numbound0 e" by simp
4154 from 4 have cp: "real c > 0" by simp
4156 let ?e="Inum (a#bs) e"
4157 let ?z = "(- ?e) / real c"
4160 hence "(real c * x < - ?e)"
4161 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac)
4162 hence "real c * x + ?e < 0" by arith
4163 hence "real c * x + ?e \<noteq> 0" by simp
4164 with xz have "?P ?z x (NEq (CN 0 c e))"
4165 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4166 hence "\<forall> x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp
4170 from 5 have nb: "numbound0 e" by simp
4171 from 5 have cp: "real c > 0" by simp
4173 let ?e="Inum (a#bs) e"
4174 let ?z = "(- ?e) / real c"
4177 hence "(real c * x < - ?e)"
4178 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac)
4179 hence "real c * x + ?e < 0" by arith
4180 with xz have "?P ?z x (Lt (CN 0 c e))"
4181 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4182 hence "\<forall> x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp
4186 from 6 have nb: "numbound0 e" by simp
4187 from 6 have cp: "real c > 0" by simp
4189 let ?e="Inum (a#bs) e"
4190 let ?z = "(- ?e) / real c"
4193 hence "(real c * x < - ?e)"
4194 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac)
4195 hence "real c * x + ?e < 0" by arith
4196 with xz have "?P ?z x (Le (CN 0 c e))"
4197 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4198 hence "\<forall> x < ?z. ?P ?z x (Le (CN 0 c e))" by simp
4202 from 7 have nb: "numbound0 e" by simp
4203 from 7 have cp: "real c > 0" by simp
4205 let ?e="Inum (a#bs) e"
4206 let ?z = "(- ?e) / real c"
4209 hence "(real c * x < - ?e)"
4210 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac)
4211 hence "real c * x + ?e < 0" by arith
4212 with xz have "?P ?z x (Gt (CN 0 c e))"
4213 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4214 hence "\<forall> x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp
4218 from 8 have nb: "numbound0 e" by simp
4219 from 8 have cp: "real c > 0" by simp
4221 let ?e="Inum (a#bs) e"
4222 let ?z = "(- ?e) / real c"
4225 hence "(real c * x < - ?e)"
4226 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac)
4227 hence "real c * x + ?e < 0" by arith
4228 with xz have "?P ?z x (Ge (CN 0 c e))"
4229 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4230 hence "\<forall> x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp
4235 assumes lp: "isrlfm p"
4236 shows "\<exists> z. \<forall> x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p")
4238 proof (induct p rule: isrlfm.induct)
4239 case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto
4241 case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto
4244 from 3 have nb: "numbound0 e" by simp
4245 from 3 have cp: "real c > 0" by simp
4247 let ?e="Inum (a#bs) e"
4248 let ?z = "(- ?e) / real c"
4251 with mult_strict_right_mono [OF xz cp] cp
4252 have "(real c * x > - ?e)" by (simp add: mult_ac)
4253 hence "real c * x + ?e > 0" by arith
4254 hence "real c * x + ?e \<noteq> 0" by simp
4255 with xz have "?P ?z x (Eq (CN 0 c e))"
4256 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4257 hence "\<forall> x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp
4261 from 4 have nb: "numbound0 e" by simp
4262 from 4 have cp: "real c > 0" by simp
4264 let ?e="Inum (a#bs) e"
4265 let ?z = "(- ?e) / real c"
4268 with mult_strict_right_mono [OF xz cp] cp
4269 have "(real c * x > - ?e)" by (simp add: mult_ac)
4270 hence "real c * x + ?e > 0" by arith
4271 hence "real c * x + ?e \<noteq> 0" by simp
4272 with xz have "?P ?z x (NEq (CN 0 c e))"
4273 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4274 hence "\<forall> x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp
4278 from 5 have nb: "numbound0 e" by simp
4279 from 5 have cp: "real c > 0" by simp
4281 let ?e="Inum (a#bs) e"
4282 let ?z = "(- ?e) / real c"
4285 with mult_strict_right_mono [OF xz cp] cp
4286 have "(real c * x > - ?e)" by (simp add: mult_ac)
4287 hence "real c * x + ?e > 0" by arith
4288 with xz have "?P ?z x (Lt (CN 0 c e))"
4289 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4290 hence "\<forall> x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp
4294 from 6 have nb: "numbound0 e" by simp
4295 from 6 have cp: "real c > 0" by simp
4297 let ?e="Inum (a#bs) e"
4298 let ?z = "(- ?e) / real c"
4301 with mult_strict_right_mono [OF xz cp] cp
4302 have "(real c * x > - ?e)" by (simp add: mult_ac)
4303 hence "real c * x + ?e > 0" by arith
4304 with xz have "?P ?z x (Le (CN 0 c e))"
4305 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4306 hence "\<forall> x > ?z. ?P ?z x (Le (CN 0 c e))" by simp
4310 from 7 have nb: "numbound0 e" by simp
4311 from 7 have cp: "real c > 0" by simp
4313 let ?e="Inum (a#bs) e"
4314 let ?z = "(- ?e) / real c"
4317 with mult_strict_right_mono [OF xz cp] cp
4318 have "(real c * x > - ?e)" by (simp add: mult_ac)
4319 hence "real c * x + ?e > 0" by arith
4320 with xz have "?P ?z x (Gt (CN 0 c e))"
4321 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4322 hence "\<forall> x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp
4326 from 8 have nb: "numbound0 e" by simp
4327 from 8 have cp: "real c > 0" by simp
4329 let ?e="Inum (a#bs) e"
4330 let ?z = "(- ?e) / real c"
4333 with mult_strict_right_mono [OF xz cp] cp
4334 have "(real c * x > - ?e)" by (simp add: mult_ac)
4335 hence "real c * x + ?e > 0" by arith
4336 with xz have "?P ?z x (Ge (CN 0 c e))"
4337 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp }
4338 hence "\<forall> x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp
4342 lemma rminusinf_bound0:
4343 assumes lp: "isrlfm p"
4344 shows "bound0 (minusinf p)"
4346 by (induct p rule: minusinf.induct) simp_all
4348 lemma rplusinf_bound0:
4349 assumes lp: "isrlfm p"
4350 shows "bound0 (plusinf p)"
4352 by (induct p rule: plusinf.induct) simp_all
4355 assumes lp: "isrlfm p"
4356 and ex: "Ifm (a#bs) (minusinf p)"
4357 shows "\<exists> x. Ifm (x#bs) p"
4359 from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
4360 have th: "\<forall> x. Ifm (x#bs) (minusinf p)" by auto
4361 from rminusinf_inf[OF lp, where bs="bs"]
4362 obtain z where z_def: "\<forall>x<z. Ifm (x # bs) (minusinf p) = Ifm (x # bs) p" by blast
4363 from th have "Ifm ((z - 1)#bs) (minusinf p)" by simp
4364 moreover have "z - 1 < z" by simp
4365 ultimately show ?thesis using z_def by auto
4369 assumes lp: "isrlfm p"
4370 and ex: "Ifm (a#bs) (plusinf p)"
4371 shows "\<exists> x. Ifm (x#bs) p"
4373 from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex
4374 have th: "\<forall> x. Ifm (x#bs) (plusinf p)" by auto
4375 from rplusinf_inf[OF lp, where bs="bs"]
4376 obtain z where z_def: "\<forall>x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast
4377 from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp
4378 moreover have "z + 1 > z" by simp
4379 ultimately show ?thesis using z_def by auto
4383 \<Upsilon>:: "fm \<Rightarrow> (num \<times> int) list"
4384 \<upsilon> :: "fm \<Rightarrow> (num \<times> int) \<Rightarrow> fm "
4385 recdef \<Upsilon> "measure size"
4386 "\<Upsilon> (And p q) = (\<Upsilon> p @ \<Upsilon> q)"
4387 "\<Upsilon> (Or p q) = (\<Upsilon> p @ \<Upsilon> q)"
4388 "\<Upsilon> (Eq (CN 0 c e)) = [(Neg e,c)]"
4389 "\<Upsilon> (NEq (CN 0 c e)) = [(Neg e,c)]"
4390 "\<Upsilon> (Lt (CN 0 c e)) = [(Neg e,c)]"
4391 "\<Upsilon> (Le (CN 0 c e)) = [(Neg e,c)]"
4392 "\<Upsilon> (Gt (CN 0 c e)) = [(Neg e,c)]"
4393 "\<Upsilon> (Ge (CN 0 c e)) = [(Neg e,c)]"
4396 recdef \<upsilon> "measure size"
4397 "\<upsilon> (And p q) = (\<lambda> (t,n). And (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
4398 "\<upsilon> (Or p q) = (\<lambda> (t,n). Or (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))"
4399 "\<upsilon> (Eq (CN 0 c e)) = (\<lambda> (t,n). Eq (Add (Mul c t) (Mul n e)))"
4400 "\<upsilon> (NEq (CN 0 c e)) = (\<lambda> (t,n). NEq (Add (Mul c t) (Mul n e)))"
4401 "\<upsilon> (Lt (CN 0 c e)) = (\<lambda> (t,n). Lt (Add (Mul c t) (Mul n e)))"
4402 "\<upsilon> (Le (CN 0 c e)) = (\<lambda> (t,n). Le (Add (Mul c t) (Mul n e)))"
4403 "\<upsilon> (Gt (CN 0 c e)) = (\<lambda> (t,n). Gt (Add (Mul c t) (Mul n e)))"
4404 "\<upsilon> (Ge (CN 0 c e)) = (\<lambda> (t,n). Ge (Add (Mul c t) (Mul n e)))"
4405 "\<upsilon> p = (\<lambda> (t,n). p)"
4407 lemma \<upsilon>_I: assumes lp: "isrlfm p"
4408 and np: "real n > 0" and nbt: "numbound0 t"
4409 shows "(Ifm (x#bs) (\<upsilon> p (t,n)) = Ifm (((Inum (x#bs) t)/(real n))#bs) p) \<and> bound0 (\<upsilon> p (t,n))" (is "(?I x (\<upsilon> p (t,n)) = ?I ?u p) \<and> ?B p" is "(_ = ?I (?t/?n) p) \<and> _" is "(_ = ?I (?N x t /_) p) \<and> _")
4411 proof(induct p rule: \<upsilon>.induct)
4413 from 5 have cp: "c >0" and nb: "numbound0 e" by simp_all
4414 have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)"
4415 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
4416 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)"
4417 by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)"
4418 and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
4419 also have "\<dots> = (real c *?t + ?n* (?N x e) < 0)"
4421 finally show ?case using nbt nb by (simp add: algebra_simps)
4424 from 6 have cp: "c >0" and nb: "numbound0 e" by simp_all
4425 have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<le> 0)"
4426 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
4427 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<le> 0)"
4428 by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)"
4429 and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
4430 also have "\<dots> = (real c *?t + ?n* (?N x e) \<le> 0)"
4432 finally show ?case using nbt nb by (simp add: algebra_simps)
4435 from 7 have cp: "c >0" and nb: "numbound0 e" by simp_all
4436 have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)"
4437 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
4438 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)"
4439 by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)"
4440 and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
4441 also have "\<dots> = (real c *?t + ?n* (?N x e) > 0)"
4443 finally show ?case using nbt nb by (simp add: algebra_simps)
4446 from 8 have cp: "c >0" and nb: "numbound0 e" by simp_all
4447 have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<ge> 0)"
4448 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
4449 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<ge> 0)"
4450 by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)"
4451 and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
4452 also have "\<dots> = (real c *?t + ?n* (?N x e) \<ge> 0)"
4454 finally show ?case using nbt nb by (simp add: algebra_simps)
4457 from 3 have cp: "c >0" and nb: "numbound0 e" by simp_all
4458 from np have np: "real n \<noteq> 0" by simp
4459 have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)"
4460 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
4461 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)"
4462 by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)"
4463 and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
4464 also have "\<dots> = (real c *?t + ?n* (?N x e) = 0)"
4466 finally show ?case using nbt nb by (simp add: algebra_simps)
4469 from 4 have cp: "c >0" and nb: "numbound0 e" by simp_all
4470 from np have np: "real n \<noteq> 0" by simp
4471 have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<noteq> 0)"
4472 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp
4473 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<noteq> 0)"
4474 by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)"
4475 and b="0", simplified divide_zero_left]) (simp only: algebra_simps)
4476 also have "\<dots> = (real c *?t + ?n* (?N x e) \<noteq> 0)"
4478 finally show ?case using nbt nb by (simp add: algebra_simps)
4479 qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"])
4482 assumes lp: "isrlfm p"
4483 shows "\<forall> (t,k) \<in> set (\<Upsilon> p). numbound0 t \<and> k >0"
4485 by(induct p rule: \<Upsilon>.induct) auto
4487 lemma rminusinf_\<Upsilon>:
4488 assumes lp: "isrlfm p"
4489 and nmi: "\<not> (Ifm (a#bs) (minusinf p))" (is "\<not> (Ifm (a#bs) (?M p))")
4490 and ex: "Ifm (x#bs) p" (is "?I x p")
4491 shows "\<exists> (s,m) \<in> set (\<Upsilon> p). x \<ge> Inum (a#bs) s / real m" (is "\<exists> (s,m) \<in> ?U p. x \<ge> ?N a s / real m")
4493 have "\<exists> (s,m) \<in> set (\<Upsilon> p). real m * x \<ge> Inum (a#bs) s " (is "\<exists> (s,m) \<in> ?U p. real m *x \<ge> ?N a s")
4495 by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"])
4496 then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<ge> ?N a s" by blast
4497 from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto
4498 from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \<ge> ?N a s / real m"
4499 by (auto simp add: mult_commute)
4500 thus ?thesis using smU by auto
4503 lemma rplusinf_\<Upsilon>:
4504 assumes lp: "isrlfm p"
4505 and nmi: "\<not> (Ifm (a#bs) (plusinf p))" (is "\<not> (Ifm (a#bs) (?M p))")
4506 and ex: "Ifm (x#bs) p" (is "?I x p")
4507 shows "\<exists> (s,m) \<in> set (\<Upsilon> p). x \<le> Inum (a#bs) s / real m" (is "\<exists> (s,m) \<in> ?U p. x \<le> ?N a s / real m")
4509 have "\<exists> (s,m) \<in> set (\<Upsilon> p). real m * x \<le> Inum (a#bs) s " (is "\<exists> (s,m) \<in> ?U p. real m *x \<le> ?N a s")
4511 by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"])
4512 then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<le> ?N a s" by blast
4513 from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto
4514 from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \<le> ?N a s / real m"
4515 by (auto simp add: mult_commute)
4516 thus ?thesis using smU by auto
4520 assumes lp: "isrlfm p"
4521 and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (t,n). Inum (x#bs) t / real n) ` set (\<Upsilon> p)"
4522 (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (t,n). ?N x t / real n ) ` (?U p)")
4523 and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p"
4524 and ly: "l < y" and yu: "y < u"
4525 shows "Ifm (y#bs) p"
4527 proof (induct p rule: isrlfm.induct)
4528 case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
4529 from 5 have "x * real c + ?N x e < 0" by (simp add: algebra_simps)
4530 hence pxc: "x < (- ?N x e) / real c"
4531 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"])
4532 from 5 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
4533 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
4534 hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
4535 moreover {assume y: "y < (-?N x e)/ real c"
4536 hence "y * real c < - ?N x e"
4537 by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
4538 hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
4539 hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
4540 moreover {assume y: "y > (- ?N x e) / real c"
4541 with yu have eu: "u > (- ?N x e) / real c" by auto
4542 with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto)
4543 with lx pxc have "False" by auto
4544 hence ?case by simp }
4545 ultimately show ?case by blast
4547 case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
4548 from 6 have "x * real c + ?N x e \<le> 0" by (simp add: algebra_simps)
4549 hence pxc: "x \<le> (- ?N x e) / real c"
4550 by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"])
4551 from 6 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
4552 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
4553 hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
4554 moreover {assume y: "y < (-?N x e)/ real c"
4555 hence "y * real c < - ?N x e"
4556 by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric])
4557 hence "real c * y + ?N x e < 0" by (simp add: algebra_simps)
4558 hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
4559 moreover {assume y: "y > (- ?N x e) / real c"
4560 with yu have eu: "u > (- ?N x e) / real c" by auto
4561 with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto)
4562 with lx pxc have "False" by auto
4563 hence ?case by simp }
4564 ultimately show ?case by blast
4566 case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
4567 from 7 have "x * real c + ?N x e > 0" by (simp add: algebra_simps)
4568 hence pxc: "x > (- ?N x e) / real c"
4569 by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"])
4570 from 7 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
4571 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
4572 hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
4573 moreover {assume y: "y > (-?N x e)/ real c"
4574 hence "y * real c > - ?N x e"
4575 by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
4576 hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
4577 hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
4578 moreover {assume y: "y < (- ?N x e) / real c"
4579 with ly have eu: "l < (- ?N x e) / real c" by auto
4580 with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto)
4581 with xu pxc have "False" by auto
4582 hence ?case by simp }
4583 ultimately show ?case by blast
4585 case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
4586 from 8 have "x * real c + ?N x e \<ge> 0" by (simp add: algebra_simps)
4587 hence pxc: "x \<ge> (- ?N x e) / real c"
4588 by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"])
4589 from 8 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
4590 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
4591 hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto
4592 moreover {assume y: "y > (-?N x e)/ real c"
4593 hence "y * real c > - ?N x e"
4594 by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric])
4595 hence "real c * y + ?N x e > 0" by (simp add: algebra_simps)
4596 hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp}
4597 moreover {assume y: "y < (- ?N x e) / real c"
4598 with ly have eu: "l < (- ?N x e) / real c" by auto
4599 with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto)
4600 with xu pxc have "False" by auto
4601 hence ?case by simp }
4602 ultimately show ?case by blast
4604 case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
4605 from cp have cnz: "real c \<noteq> 0" by simp
4606 from 3 have "x * real c + ?N x e = 0" by (simp add: algebra_simps)
4607 hence pxc: "x = (- ?N x e) / real c"
4608 by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"])
4609 from 3 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
4610 with lx xu have yne: "x \<noteq> - ?N x e / real c" by auto
4611 with pxc show ?case by simp
4613 case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp_all
4614 from cp have cnz: "real c \<noteq> 0" by simp
4615 from 4 have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto
4616 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto
4617 hence "y* real c \<noteq> -?N x e"
4618 by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp
4619 hence "y* real c + ?N x e \<noteq> 0" by (simp add: algebra_simps)
4620 thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"]
4621 by (simp add: algebra_simps)
4622 qed (auto simp add: numbound0_I[where bs="bs" and b="y" and b'="x"])
4624 lemma rinf_\<Upsilon>:
4625 assumes lp: "isrlfm p"
4626 and nmi: "\<not> (Ifm (x#bs) (minusinf p))" (is "\<not> (Ifm (x#bs) (?M p))")
4627 and npi: "\<not> (Ifm (x#bs) (plusinf p))" (is "\<not> (Ifm (x#bs) (?P p))")
4628 and ex: "\<exists> x. Ifm (x#bs) p" (is "\<exists> x. ?I x p")
4629 shows "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p).
4630 ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p"
4632 let ?N = "\<lambda> x t. Inum (x#bs) t"
4633 let ?U = "set (\<Upsilon> p)"
4634 from ex obtain a where pa: "?I a p" by blast
4635 from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi
4636 have nmi': "\<not> (?I a (?M p))" by simp
4637 from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi
4638 have npi': "\<not> (?I a (?P p))" by simp
4639 have "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). ?I ((?N a l/real n + ?N a s /real m) / 2) p"
4641 let ?M = "(\<lambda> (t,c). ?N a t / real c) ` ?U"
4642 have fM: "finite ?M" by auto
4643 from rminusinf_\<Upsilon>[OF lp nmi pa] rplusinf_\<Upsilon>[OF lp npi pa]
4644 have "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). a \<le> ?N x l / real n \<and> a \<ge> ?N x s / real m" by blast
4645 then obtain "t" "n" "s" "m" where
4646 tnU: "(t,n) \<in> ?U" and smU: "(s,m) \<in> ?U"
4647 and xs1: "a \<le> ?N x s / real m" and tx1: "a \<ge> ?N x t / real n" by blast
4648 from \<Upsilon>_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \<le> ?N a s / real m" and tx: "a \<ge> ?N a t / real n" by auto
4649 from tnU have Mne: "?M \<noteq> {}" by auto
4650 hence Une: "?U \<noteq> {}" by simp
4653 have linM: "?l \<in> ?M" using fM Mne by simp
4654 have uinM: "?u \<in> ?M" using fM Mne by simp
4655 have tnM: "?N a t / real n \<in> ?M" using tnU by auto
4656 have smM: "?N a s / real m \<in> ?M" using smU by auto
4657 have lM: "\<forall> t\<in> ?M. ?l \<le> t" using Mne fM by auto
4658 have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto
4659 have "?l \<le> ?N a t / real n" using tnM Mne by simp hence lx: "?l \<le> a" using tx by simp
4660 have "?N a s / real m \<le> ?u" using smM Mne by simp hence xu: "a \<le> ?u" using xs by simp
4661 from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
4662 have "(\<exists> s\<in> ?M. ?I s p) \<or>
4663 (\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p)" .
4664 moreover { fix u assume um: "u\<in> ?M" and pu: "?I u p"
4665 hence "\<exists> (tu,nu) \<in> ?U. u = ?N a tu / real nu" by auto
4666 then obtain "tu" "nu" where tuU: "(tu,nu) \<in> ?U" and tuu:"u= ?N a tu / real nu" by blast
4667 have "(u + u) / 2 = u" by auto with pu tuu
4668 have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp
4669 with tuU have ?thesis by blast}
4671 assume "\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
4672 then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M"
4673 and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
4675 from t1M have "\<exists> (t1u,t1n) \<in> ?U. t1 = ?N a t1u / real t1n" by auto
4676 then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast
4677 from t2M have "\<exists> (t2u,t2n) \<in> ?U. t2 = ?N a t2u / real t2n" by auto
4678 then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \<in> ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast
4679 from t1x xt2 have t1t2: "t1 < t2" by simp
4680 let ?u = "(t1 + t2) / 2"
4681 from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto
4682 from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" .
4683 with t1uU t2uU t1u t2u have ?thesis by blast}
4684 ultimately show ?thesis by blast
4686 then obtain "l" "n" "s" "m" where lnU: "(l,n) \<in> ?U" and smU:"(s,m) \<in> ?U"
4687 and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast
4688 from lnU smU \<Upsilon>_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto
4689 from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"]
4690 numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
4691 have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp
4693 show ?thesis by auto
4695 (* The Ferrante - Rackoff Theorem *)
4698 assumes lp: "isrlfm p"
4699 shows "(\<exists> x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \<or> (Ifm (x#bs) (plusinf p)) \<or> (\<exists> (t,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). Ifm ((((Inum (x#bs) t)/ real n + (Inum (x#bs) s) / real m) /2)#bs) p))"
4700 (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
4702 assume px: "\<exists> x. ?I x p"
4703 have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
4704 moreover {assume "?M \<or> ?P" hence "?D" by blast}
4705 moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
4706 from rinf_\<Upsilon>[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast}
4707 ultimately show "?D" by blast
4710 moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .}
4711 moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
4712 moreover {assume f:"?F" hence "?E" by blast}
4713 ultimately show "?E" by blast
4717 lemma fr_eq_\<upsilon>:
4718 assumes lp: "isrlfm p"
4719 shows "(\<exists> x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \<or> (Ifm (x#bs) (plusinf p)) \<or> (\<exists> (t,k) \<in> set (\<Upsilon> p). \<exists> (s,l) \<in> set (\<Upsilon> p). Ifm (x#bs) (\<upsilon> p (Add(Mul l t) (Mul k s) , 2*k*l))))"
4720 (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
4722 assume px: "\<exists> x. ?I x p"
4723 have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
4724 moreover {assume "?M \<or> ?P" hence "?D" by blast}
4725 moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
4726 let ?f ="\<lambda> (t,n). Inum (x#bs) t / real n"
4727 let ?N = "\<lambda> t. Inum (x#bs) t"
4728 {fix t n s m assume "(t,n)\<in> set (\<Upsilon> p)" and "(s,m) \<in> set (\<Upsilon> p)"
4729 with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0"
4731 let ?st = "Add (Mul m t) (Mul n s)"
4732 from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0"
4733 by (simp add: mult_commute)
4734 from tnb snb have st_nb: "numbound0 ?st" by simp
4735 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
4736 using mnp mp np by (simp add: algebra_simps add_divide_distrib)
4737 from \<upsilon>_I[OF lp mnp st_nb, where x="x" and bs="bs"]
4738 have "?I x (\<upsilon> p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])}
4739 with rinf_\<Upsilon>[OF lp nmi npi px] have "?F" by blast hence "?D" by blast}
4740 ultimately show "?D" by blast
4743 moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .}
4744 moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . }
4745 moreover {fix t k s l assume "(t,k) \<in> set (\<Upsilon> p)" and "(s,l) \<in> set (\<Upsilon> p)"
4746 and px:"?I x (\<upsilon> p (Add (Mul l t) (Mul k s), 2*k*l))"
4747 with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto
4748 let ?st = "Add (Mul l t) (Mul k s)"
4749 from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0"
4750 by (simp add: mult_commute)
4751 from tnb snb have st_nb: "numbound0 ?st" by simp
4752 from \<upsilon>_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto}
4753 ultimately show "?E" by blast
4756 text{* The overall Part *}
4758 lemma real_ex_int_real01:
4759 shows "(\<exists> (x::real). P x) = (\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))"
4764 let ?u = "x - real ?i"
4765 have "x = real ?i + ?u" by simp
4766 hence "P (real ?i + ?u)" using Px by simp
4767 moreover have "real ?i \<le> x" using real_of_int_floor_le by simp hence "0 \<le> ?u" by arith
4768 moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith
4769 ultimately show "(\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))" by blast
4772 fun exsplitnum :: "num \<Rightarrow> num" where
4773 "exsplitnum (C c) = (C c)"
4774 | "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)"
4775 | "exsplitnum (Bound n) = Bound (n+1)"
4776 | "exsplitnum (Neg a) = Neg (exsplitnum a)"
4777 | "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) "
4778 | "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) "
4779 | "exsplitnum (Mul c a) = Mul c (exsplitnum a)"
4780 | "exsplitnum (Floor a) = Floor (exsplitnum a)"
4781 | "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))"
4782 | "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)"
4783 | "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)"
4785 fun exsplit :: "fm \<Rightarrow> fm" where
4786 "exsplit (Lt a) = Lt (exsplitnum a)"
4787 | "exsplit (Le a) = Le (exsplitnum a)"
4788 | "exsplit (Gt a) = Gt (exsplitnum a)"
4789 | "exsplit (Ge a) = Ge (exsplitnum a)"
4790 | "exsplit (Eq a) = Eq (exsplitnum a)"
4791 | "exsplit (NEq a) = NEq (exsplitnum a)"
4792 | "exsplit (Dvd i a) = Dvd i (exsplitnum a)"
4793 | "exsplit (NDvd i a) = NDvd i (exsplitnum a)"
4794 | "exsplit (And p q) = And (exsplit p) (exsplit q)"
4795 | "exsplit (Or p q) = Or (exsplit p) (exsplit q)"
4796 | "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)"
4797 | "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)"
4798 | "exsplit (NOT p) = NOT (exsplit p)"
4802 "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t"
4803 by(induct t rule: exsplitnum.induct) (simp_all add: algebra_simps)
4806 assumes qfp: "qfree p"
4807 shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p"
4808 using qfp exsplitnum[where x="x" and y="y" and bs="bs"]
4809 by(induct p rule: exsplit.induct) simp_all
4812 assumes qf: "qfree p"
4813 shows "(Ifm bs (E p)) = (\<exists> (i::int). Ifm (real i#bs) (E (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (exsplit p))))" (is "?lhs = ?rhs")
4815 have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real i)#bs) (exsplit p))"
4816 by (simp add: myless[of _ "1"] myless[of _ "0"] add_ac diff_minus)
4817 also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real i + x) #bs) p)"
4818 by (simp only: exsplit[OF qf] add_ac)
4819 also have "\<dots> = (\<exists> x. Ifm (x#bs) p)"
4820 by (simp only: real_ex_int_real01[where P="\<lambda> x. Ifm (x#bs) p"])
4821 finally show ?thesis by simp
4824 (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *)
4826 definition ferrack01 :: "fm \<Rightarrow> fm" where
4827 "ferrack01 p \<equiv> (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p);
4828 U = remdups(map simp_num_pair
4829 (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m))
4830 (alluopairs (\<Upsilon> p'))))
4831 in decr (evaldjf (\<upsilon> p') U ))"
4834 assumes qf: "qfree p"
4835 shows "(\<exists> x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (\<exists> (t,n) \<in> set (\<Upsilon> (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). \<exists> (s,m) \<in> set (\<Upsilon> (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). Ifm (x#bs) (\<upsilon> (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) (Add (Mul m t) (Mul n s), 2*n*m)))"
4836 (is "(\<exists> x. ?I x ?q) = ?F")
4839 let ?M = "?I x (minusinf ?rq)"
4840 let ?P = "?I x (plusinf ?rq)"
4841 have MF: "?M = False"
4842 apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
4843 by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
4844 have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def)
4845 by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all)
4846 have "(\<exists> x. ?I x ?q ) =
4847 ((?I x (minusinf ?rq)) \<or> (?I x (plusinf ?rq )) \<or> (\<exists> (t,n) \<in> set (\<Upsilon> ?rq). \<exists> (s,m) \<in> set (\<Upsilon> ?rq ). ?I x (\<upsilon> ?rq (Add (Mul m t) (Mul n s), 2*n*m))))"
4848 (is "(\<exists> x. ?I x ?q) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
4850 assume "\<exists> x. ?I x ?q"
4851 then obtain x where qx: "?I x ?q" by blast
4852 hence xp: "0\<le> x" and x1: "x< 1" and px: "?I x p"
4853 by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf])
4854 from qx have "?I x ?rq "
4855 by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1])
4856 hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto
4857 from qf have qfq:"isrlfm ?rq"
4858 by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1])
4859 with lqx fr_eq_\<upsilon>[OF qfq] show "?M \<or> ?P \<or> ?F" by blast
4862 let ?U = "set (\<Upsilon> ?rq )"
4863 from MF PF D have "?F" by auto
4864 then obtain t n s m where aU:"(t,n) \<in> ?U" and bU:"(s,m)\<in> ?U" and rqx: "?I x (\<upsilon> ?rq (Add (Mul m t) (Mul n s), 2*n*m))" by blast
4865 from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf]
4866 by (auto simp add: rsplit_def lt_def ge_def)
4867 from aU bU \<Upsilon>_l[OF lrq] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by (auto simp add: split_def)
4868 let ?st = "Add (Mul m t) (Mul n s)"
4869 from tnb snb have stnb: "numbound0 ?st" by simp
4870 from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0"
4871 by (simp add: mult_commute)
4872 from conjunct1[OF \<upsilon>_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx
4873 have "\<exists> x. ?I x ?rq" by auto
4875 using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def)
4877 with MF PF show ?thesis by blast
4880 lemma \<Upsilon>_cong_aux:
4881 assumes Ul: "\<forall> (t,n) \<in> set U. numbound0 t \<and> n >0"
4882 shows "((\<lambda> (t,n). Inum (x#bs) t /real n) ` (set (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((\<lambda> ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (set U \<times> set U))"
4886 assume "((t,n),(s,m)) \<in> set (alluopairs U)"
4887 hence th: "((t,n),(s,m)) \<in> (set U \<times> set U)"
4888 using alluopairs_set1[where xs="U"] by blast
4889 let ?N = "\<lambda> t. Inum (x#bs) t"
4890 let ?st= "Add (Mul m t) (Mul n s)"
4891 from Ul th have mnz: "m \<noteq> 0" by auto
4892 from Ul th have nnz: "n \<noteq> 0" by auto
4893 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
4894 using mnz nnz by (simp add: algebra_simps add_divide_distrib)
4896 thus "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) /
4897 (2 * real n * real m)
4898 \<in> (\<lambda>((t, n), s, m).
4899 (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) `
4900 (set U \<times> set U)"using mnz nnz th
4901 apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def)
4902 by (rule_tac x="(s,m)" in bexI,simp_all)
4903 (rule_tac x="(t,n)" in bexI,simp_all add: mult_commute)
4906 assume tnU: "(t,n) \<in> set U" and smU:"(s,m) \<in> set U"
4907 let ?N = "\<lambda> t. Inum (x#bs) t"
4908 let ?st= "Add (Mul m t) (Mul n s)"
4909 from Ul smU have mnz: "m \<noteq> 0" by auto
4910 from Ul tnU have nnz: "n \<noteq> 0" by auto
4911 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
4912 using mnz nnz by (simp add: algebra_simps add_divide_distrib)
4913 let ?P = "\<lambda> (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2"
4914 have Pc:"\<forall> a b. ?P a b = ?P b a"
4916 from Ul alluopairs_set1 have Up:"\<forall> ((t,n),(s,m)) \<in> set (alluopairs U). n \<noteq> 0 \<and> m \<noteq> 0" by blast
4917 from alluopairs_ex[OF Pc, where xs="U"] tnU smU
4918 have th':"\<exists> ((t',n'),(s',m')) \<in> set (alluopairs U). ?P (t',n') (s',m')"
4920 then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \<in> set (alluopairs U)"
4921 and Pts': "?P (t',n') (s',m')" by blast
4922 from ts'_U Up have mnz': "m' \<noteq> 0" and nnz': "n'\<noteq> 0" by auto
4923 let ?st' = "Add (Mul m' t') (Mul n' s')"
4924 have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')"
4925 using mnz' nnz' by (simp add: algebra_simps add_divide_distrib)
4927 "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp
4928 also have "\<dots> = ((\<lambda>(t, n). Inum (x # bs) t / real n) ((\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st')
4929 finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2
4930 \<in> (\<lambda>(t, n). Inum (x # bs) t / real n) `
4931 (\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) `
4933 using ts'_U by blast
4936 lemma \<Upsilon>_cong:
4937 assumes lp: "isrlfm p"
4938 and UU': "((\<lambda> (t,n). Inum (x#bs) t /real n) ` U') = ((\<lambda> ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (U \<times> U))" (is "?f ` U' = ?g ` (U\<times>U)")
4939 and U: "\<forall> (t,n) \<in> U. numbound0 t \<and> n > 0"
4940 and U': "\<forall> (t,n) \<in> U'. numbound0 t \<and> n > 0"
4941 shows "(\<exists> (t,n) \<in> U. \<exists> (s,m) \<in> U. Ifm (x#bs) (\<upsilon> p (Add (Mul m t) (Mul n s),2*n*m))) = (\<exists> (t,n) \<in> U'. Ifm (x#bs) (\<upsilon> p (t,n)))"
4945 then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and
4946 Pst: "Ifm (x#bs) (\<upsilon> p (Add (Mul m t) (Mul n s),2*n*m))" by blast
4947 let ?N = "\<lambda> t. Inum (x#bs) t"
4948 from tnU smU U have tnb: "numbound0 t" and np: "n > 0"
4949 and snb: "numbound0 s" and mp:"m > 0" by auto
4950 let ?st= "Add (Mul m t) (Mul n s)"
4951 from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0"
4952 by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
4953 from tnb snb have stnb: "numbound0 ?st" by simp
4954 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
4955 using mp np by (simp add: algebra_simps add_divide_distrib)
4956 from tnU smU UU' have "?g ((t,n),(s,m)) \<in> ?f ` U'" by blast
4957 hence "\<exists> (t',n') \<in> U'. ?g ((t,n),(s,m)) = ?f (t',n')"
4958 by auto (rule_tac x="(a,b)" in bexI, auto)
4959 then obtain t' n' where tnU': "(t',n') \<in> U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast
4960 from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
4961 from \<upsilon>_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst
4962 have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
4963 from conjunct1[OF \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x"], symmetric] th[simplified split_def fst_conv snd_conv,symmetric] Pst2[simplified st[symmetric]]
4964 have "Ifm (x # bs) (\<upsilon> p (t', n')) " by (simp only: st)
4965 then show ?rhs using tnU' by auto
4968 then obtain t' n' where tnU': "(t',n') \<in> U'" and Pt': "Ifm (x # bs) (\<upsilon> p (t', n'))"
4970 from tnU' UU' have "?f (t',n') \<in> ?g ` (U\<times>U)" by blast
4971 hence "\<exists> ((t,n),(s,m)) \<in> (U\<times>U). ?f (t',n') = ?g ((t,n),(s,m))"
4972 by auto (rule_tac x="(a,b)" in bexI, auto)
4973 then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and
4974 th: "?f (t',n') = ?g((t,n),(s,m)) "by blast
4975 let ?N = "\<lambda> t. Inum (x#bs) t"
4976 from tnU smU U have tnb: "numbound0 t" and np: "n > 0"
4977 and snb: "numbound0 s" and mp:"m > 0" by auto
4978 let ?st= "Add (Mul m t) (Mul n s)"
4979 from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0"
4980 by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult)
4981 from tnb snb have stnb: "numbound0 ?st" by simp
4982 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)"
4983 using mp np by (simp add: algebra_simps add_divide_distrib)
4984 from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto
4985 from \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt'
4986 have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp
4987 with \<upsilon>_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast
4991 assumes qf: "qfree p"
4992 shows "((\<exists> x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (Ifm bs (ferrack01 p))) \<and> qfree (ferrack01 p)" (is "(?lhs = ?rhs) \<and> _")
4994 let ?I = "\<lambda> x p. Ifm (x#bs) p"
4996 let ?N = "\<lambda> t. Inum (x#bs) t"
4997 let ?q = "rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)"
4998 let ?U = "\<Upsilon> ?q"
4999 let ?Up = "alluopairs ?U"
5000 let ?g = "\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)"
5001 let ?S = "map ?g ?Up"
5002 let ?SS = "map simp_num_pair ?S"
5003 let ?Y = "remdups ?SS"
5004 let ?f= "(\<lambda> (t,n). ?N t / real n)"
5005 let ?h = "\<lambda> ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2"
5006 let ?F = "\<lambda> p. \<exists> a \<in> set (\<Upsilon> p). \<exists> b \<in> set (\<Upsilon> p). ?I x (\<upsilon> p (?g(a,b)))"
5007 let ?ep = "evaldjf (\<upsilon> ?q) ?Y"
5008 from rlfm_l[OF qf] have lq: "isrlfm ?q"
5009 by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def)
5010 from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \<le> (set ?U \<times> set ?U)" by simp
5011 from \<Upsilon>_l[OF lq] have U_l: "\<forall> (t,n) \<in> set ?U. numbound0 t \<and> n > 0" .
5013 have "\<forall> ((t,n),(s,m)) \<in> set ?Up. numbound0 t \<and> n> 0 \<and> numbound0 s \<and> m > 0" by auto
5014 hence Snb: "\<forall> (t,n) \<in> set ?S. numbound0 t \<and> n > 0 "
5015 by (auto simp add: mult_pos_pos)
5016 have Y_l: "\<forall> (t,n) \<in> set ?Y. numbound0 t \<and> n > 0"
5018 { fix t n assume tnY: "(t,n) \<in> set ?Y"
5019 hence "(t,n) \<in> set ?SS" by simp
5020 hence "\<exists> (t',n') \<in> set ?S. simp_num_pair (t',n') = (t,n)"
5021 by (auto simp add: split_def simp del: map_map)
5022 (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all)
5023 then obtain t' n' where tn'S: "(t',n') \<in> set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast
5024 from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto
5025 from simp_num_pair_l[OF tnb np tns]
5026 have "numbound0 t \<and> n > 0" . }
5027 thus ?thesis by blast
5030 have YU: "(?f ` set ?Y) = (?h ` (set ?U \<times> set ?U))"
5032 from simp_num_pair_ci[where bs="x#bs"] have
5033 "\<forall>x. (?f o simp_num_pair) x = ?f x" by auto
5034 hence th: "?f o simp_num_pair = ?f" using ext by blast
5035 have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_compose)
5036 also have "\<dots> = (?f ` set ?S)" by (simp add: th)
5037 also have "\<dots> = ((?f o ?g) ` set ?Up)"
5038 by (simp only: set_map o_def image_compose[symmetric])
5039 also have "\<dots> = (?h ` (set ?U \<times> set ?U))"
5040 using \<Upsilon>_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast
5041 finally show ?thesis .
5043 have "\<forall> (t,n) \<in> set ?Y. bound0 (\<upsilon> ?q (t,n))"
5045 { fix t n assume tnY: "(t,n) \<in> set ?Y"
5046 with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto
5047 from \<upsilon>_I[OF lq np tnb]
5048 have "bound0 (\<upsilon> ?q (t,n))" by simp}
5049 thus ?thesis by blast
5051 hence ep_nb: "bound0 ?ep" using evaldjf_bound0[where xs="?Y" and f="\<upsilon> ?q"]
5054 from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q"
5055 by (simp only: split_def fst_conv snd_conv)
5056 also have "\<dots> = (\<exists> (t,n) \<in> set ?Y. ?I x (\<upsilon> ?q (t,n)))" using \<Upsilon>_cong[OF lq YU U_l Y_l]
5057 by (simp only: split_def fst_conv snd_conv)
5058 also have "\<dots> = (Ifm (x#bs) ?ep)"
5059 using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\<upsilon> ?q",symmetric]
5060 by (simp only: split_def pair_collapse)
5061 also have "\<dots> = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast
5062 finally have lr: "?lhs = ?rhs" by (simp only: ferrack01_def Let_def)
5063 from decr_qf[OF ep_nb] have "qfree (ferrack01 p)" by (simp only: Let_def ferrack01_def)
5064 with lr show ?thesis by blast
5068 assumes lp: "iszlfm p (real (i::int)#bs)"
5069 and up: "d_\<beta> p 1" and dd: "d_\<delta> p d" and dp: "d > 0"
5070 shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. d}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> j\<in> {1.. d}. \<exists> b\<in> (Inum (real i#bs)) ` set (\<beta> p). Ifm ((b+real j)#bs) p))"
5071 using cp_thm[OF lp up dd dp] by auto
5073 definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int" where
5074 "unit p \<equiv> (let p' = zlfm p ; l = \<zeta> p' ; q = And (Dvd l (CN 0 1 (C 0))) (a_\<beta> p' l); d = \<delta> q;
5075 B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q))
5076 in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
5078 lemma unit: assumes qf: "qfree p"
5079 shows "\<And> q B d. unit p = (q,B,d) \<Longrightarrow> ((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and> (Inum (real i#bs)) ` set B = (Inum (real i#bs)) ` set (\<beta> q) \<and> d_\<beta> q 1 \<and> d_\<delta> q d \<and> d >0 \<and> iszlfm q (real (i::int)#bs) \<and> (\<forall> b\<in> set B. numbound0 b)"
5082 assume qBd: "unit p = (q,B,d)"
5083 let ?thes = "((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and>
5084 Inum (real i#bs) ` set B = Inum (real i#bs) ` set (\<beta> q) \<and>
5085 d_\<beta> q 1 \<and> d_\<delta> q d \<and> 0 < d \<and> iszlfm q (real i # bs) \<and> (\<forall> b\<in> set B. numbound0 b)"
5086 let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
5088 let ?l = "\<zeta> ?p'"
5089 let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a_\<beta> ?p' ?l)"
5090 let ?d = "\<delta> ?q"
5091 let ?B = "set (\<beta> ?q)"
5092 let ?B'= "remdups (map simpnum (\<beta> ?q))"
5093 let ?A = "set (\<alpha> ?q)"
5094 let ?A'= "remdups (map simpnum (\<alpha> ?q))"
5095 from conjunct1[OF zlfm_I[OF qf, where bs="bs"]]
5096 have pp': "\<forall> i. ?I i ?p' = ?I i p" by auto
5097 from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]]
5098 have lp': "\<forall> (i::int). iszlfm ?p' (real i#bs)" by simp
5099 hence lp'': "iszlfm ?p' (real (i::int)#bs)" by simp
5100 from lp' \<zeta>[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d_\<beta> ?p' ?l" by auto
5101 from a_\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp'
5102 have pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by (simp add: int_rdvd_iff)
5103 from lp'' lp a_\<beta>[OF lp'' dl lp] have lq:"iszlfm ?q (real i#bs)" and uq: "d_\<beta> ?q 1"
5104 by (auto simp add: isint_def)
5105 from \<delta>[OF lq] have dp:"?d >0" and dd: "d_\<delta> ?q ?d" by blast+
5106 let ?N = "\<lambda> t. Inum (real (i::int)#bs) t"
5107 have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_compose)
5108 also have "\<dots> = ?N ` ?B" using simpnum_ci[where bs="real i #bs"] by auto
5109 finally have BB': "?N ` set ?B' = ?N ` ?B" .
5110 have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_compose)
5111 also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] by auto
5112 finally have AA': "?N ` set ?A' = ?N ` ?A" .
5113 from \<beta>_numbound0[OF lq] have B_nb:"\<forall> b\<in> set ?B'. numbound0 b"
5114 by (simp add: simpnum_numbound0)
5115 from \<alpha>_l[OF lq] have A_nb: "\<forall> b\<in> set ?A'. numbound0 b"
5116 by (simp add: simpnum_numbound0)
5117 {assume "length ?B' \<le> length ?A'"
5118 hence q:"q=?q" and "B = ?B'" and d:"d = ?d"
5119 using qBd by (auto simp add: Let_def unit_def)
5120 with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<beta> q)"
5121 and bn: "\<forall>b\<in> set B. numbound0 b" by simp+
5122 with pq_ex dp uq dd lq q d have ?thes by simp}
5124 {assume "\<not> (length ?B' \<le> length ?A')"
5125 hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"
5126 using qBd by (auto simp add: Let_def unit_def)
5127 with AA' mirror_\<alpha>_\<beta>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<beta> q)"
5128 and bn: "\<forall>b\<in> set B. numbound0 b" by simp+
5129 from mirror_ex[OF lq] pq_ex q
5130 have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
5131 from lq uq q mirror_d_\<beta> [where p="?q" and bs="bs" and a="real i"]
5132 have lq': "iszlfm q (real i#bs)" and uq: "d_\<beta> q 1" by auto
5133 from \<delta>[OF lq'] mirror_\<delta>[OF lq] q d have dq:"d_\<delta> q d " by auto
5134 from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp
5136 ultimately show ?thes by blast
5138 (* Cooper's Algorithm *)
5140 definition cooper :: "fm \<Rightarrow> fm" where
5142 (let (q,B,d) = unit p; js = [1..d];
5143 mq = simpfm (minusinf q);
5144 md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) js
5145 in if md = T then T else
5146 (let qd = evaldjf (\<lambda> t. simpfm (subst0 t q))
5147 (remdups (map (\<lambda> (b,j). simpnum (Add b (C j)))
5148 [(b,j). b\<leftarrow>B,j\<leftarrow>js]))
5149 in decr (disj md qd)))"
5150 lemma cooper: assumes qf: "qfree p"
5151 shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (cooper p))) \<and> qfree (cooper p)"
5152 (is "(?lhs = ?rhs) \<and> _")
5155 let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
5156 let ?q = "fst (unit p)"
5157 let ?B = "fst (snd(unit p))"
5158 let ?d = "snd (snd (unit p))"
5160 let ?mq = "minusinf ?q"
5161 let ?smq = "simpfm ?mq"
5162 let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
5164 let ?N = "\<lambda> t. Inum (real (i::int)#bs) t"
5165 let ?bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]"
5166 let ?sbjs = "map (\<lambda> (b,j). simpnum (Add b (C j))) ?bjs"
5167 let ?qd = "evaldjf (\<lambda> t. simpfm (subst0 t ?q)) (remdups ?sbjs)"
5168 have qbf:"unit p = (?q,?B,?d)" by simp
5169 from unit[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and
5170 B:"?N ` set ?B = ?N ` set (\<beta> ?q)" and
5171 uq:"d_\<beta> ?q 1" and dd: "d_\<delta> ?q ?d" and dp: "?d > 0" and
5172 lq: "iszlfm ?q (real i#bs)" and
5173 Bn: "\<forall> b\<in> set ?B. numbound0 b" by auto
5174 from zlin_qfree[OF lq] have qfq: "qfree ?q" .
5175 from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq".
5176 have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
5177 hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)"
5178 by (auto simp only: subst0_bound0[OF qfmq])
5179 hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
5180 by (auto simp add: simpfm_bound0)
5181 from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp
5182 from Bn jsnb have "\<forall> (b,j) \<in> set ?bjs. numbound0 (Add b (C j))"
5184 hence "\<forall> (b,j) \<in> set ?bjs. numbound0 (simpnum (Add b (C j)))"
5185 using simpnum_numbound0 by blast
5186 hence "\<forall> t \<in> set ?sbjs. numbound0 t" by simp
5187 hence "\<forall> t \<in> set (remdups ?sbjs). bound0 (subst0 t ?q)"
5188 using subst0_bound0[OF qfq] by auto
5189 hence th': "\<forall> t \<in> set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))"
5190 using simpfm_bound0 by blast
5191 from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp
5193 have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all)
5194 from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B
5195 have "?lhs = (\<exists> j\<in> {1.. ?d}. ?I j ?mq \<or> (\<exists> b\<in> ?N ` set ?B. Ifm ((b+ real j)#bs) ?q))" by auto
5196 also have "\<dots> = ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> (b,j) \<in> (?N ` set ?B \<times> set ?js). Ifm ((b+ real j)#bs) ?q))" by auto
5197 also have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> (\<lambda> (b,j). ?N (Add b (C j))) ` set ?bjs. Ifm (t #bs) ?q))" by simp
5198 also have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> (\<lambda> (b,j). ?N (simpnum (Add b (C j)))) ` set ?bjs. Ifm (t #bs) ?q))" by (simp only: simpnum_ci)
5199 also have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> set ?sbjs. Ifm (?N t #bs) ?q))"
5200 by (auto simp add: split_def)
5201 also have "\<dots> = ((\<exists> j\<in> set ?js. (\<lambda> j. ?I i (simpfm (subst0 (C j) ?smq))) j) \<or> (\<exists> t \<in> set (remdups ?sbjs). (\<lambda> t. ?I i (simpfm (subst0 t ?q))) t))" by (simp only: simpfm subst0_I[OF qfq] simpfm Inum.simps subst0_I[OF qfmq] set_remdups)
5202 also have "\<dots> = ((?I i (evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js)) \<or> (?I i (evaldjf (\<lambda> t. simpfm (subst0 t ?q)) (remdups ?sbjs))))" by (simp only: evaldjf_ex)
5203 finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by (simp add: disj)
5204 hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp
5205 {assume mdT: "?md = T"
5206 hence cT:"cooper p = T"
5207 by (simp only: cooper_def unit_def split_def Let_def if_True) simp
5208 from mdT mdqd have lhs:"?lhs" by (auto simp add: disj)
5209 from mdT have "?rhs" by (simp add: cooper_def unit_def split_def)
5210 with lhs cT have ?thesis by simp }
5212 {assume mdT: "?md \<noteq> T" hence "cooper p = decr (disj ?md ?qd)"
5213 by (simp only: cooper_def unit_def split_def Let_def if_False)
5214 with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp }
5215 ultimately show ?thesis by blast
5219 assumes qf: "qfree p"
5220 shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ cooper p))) \<and> qfree (DJ cooper p)"
5222 from cooper have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (cooper p)" by blast
5223 from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast
5224 have "Ifm bs (DJ cooper p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (cooper q))"
5225 by (simp add: DJ_def evaldjf_ex)
5226 also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs) q)"
5227 using cooper disjuncts_qf[OF qf] by blast
5228 also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto)
5229 finally show ?thesis using thqf by blast
5232 (* Redy and Loveland *)
5234 lemma \<sigma>_\<rho>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'"
5235 shows "Ifm (a#bs) (\<sigma>_\<rho> p (t,c)) = Ifm (a#bs) (\<sigma>_\<rho> p (t',c))"
5237 by (induct p rule: iszlfm.induct, auto simp add: tt')
5239 lemma \<sigma>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'"
5240 shows "Ifm (a#bs) (\<sigma> p c t) = Ifm (a#bs) (\<sigma> p c t')"
5241 by (simp add: \<sigma>_def tt' \<sigma>_\<rho>_cong[OF lp tt'])
5243 lemma \<rho>_cong: assumes lp: "iszlfm p (a#bs)"
5244 and RR: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R = (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
5245 shows "(\<exists> (e,c) \<in> R. \<exists> j\<in> {1.. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j)))) = (\<exists> (e,c) \<in> set (\<rho> p). \<exists> j\<in> {1.. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j))))"
5248 let ?d = "\<delta> p"
5249 assume ?lhs then obtain e c j where ecR: "(e,c) \<in> R" and jD:"j \<in> {1 .. c*?d}"
5250 and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast
5251 from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" by auto
5252 hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" using RR by simp
5253 hence "\<exists> (e',c') \<in> set (\<rho> p). Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto
5254 then obtain e' c' where ecRo:"(e',c') \<in> set (\<rho> p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'"
5255 and cc':"c = c'" by blast
5256 from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
5258 from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
5259 from ecRo jD px' cc' show ?rhs apply auto
5260 by (rule_tac x="(e', c')" in bexI,simp_all)
5261 (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
5263 let ?d = "\<delta> p"
5264 assume ?rhs then obtain e c j where ecR: "(e,c) \<in> set (\<rho> p)" and jD:"j \<in> {1 .. c*?d}"
5265 and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast
5266 from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" by auto
5267 hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp
5268 hence "\<exists> (e',c') \<in> R. Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto
5269 then obtain e' c' where ecRo:"(e',c') \<in> R" and ee':"Inum (a#bs) e = Inum (a#bs) e'"
5270 and cc':"c = c'" by blast
5271 from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp
5272 from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp
5273 from ecRo jD px' cc' show ?lhs apply auto
5274 by (rule_tac x="(e', c')" in bexI,simp_all)
5275 (rule_tac x="j" in bexI, simp_all add: cc'[symmetric])
5279 assumes lp: "iszlfm p (real (i::int)#bs)"
5280 and R: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R = (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)"
5281 shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. \<delta> p}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> (e,c) \<in> R. \<exists> j\<in> {1.. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j)))))"
5282 using rl_thm[OF lp] \<rho>_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp
5284 definition chooset :: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int" where
5285 "chooset p \<equiv> (let q = zlfm p ; d = \<delta> q;
5286 B = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<rho> q)) ;
5287 a = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<alpha>_\<rho> q))
5288 in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
5290 lemma chooset: assumes qf: "qfree p"
5291 shows "\<And> q B d. chooset p = (q,B,d) \<Longrightarrow> ((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and> ((\<lambda>(t,k). (Inum (real i#bs) t,k)) ` set B = (\<lambda>(t,k). (Inum (real i#bs) t,k)) ` set (\<rho> q)) \<and> (\<delta> q = d) \<and> d >0 \<and> iszlfm q (real (i::int)#bs) \<and> (\<forall> (e,c)\<in> set B. numbound0 e \<and> c>0)"
5294 assume qBd: "chooset p = (q,B,d)"
5295 let ?thes = "((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and> ((\<lambda>(t,k). (Inum (real i#bs) t,k)) ` set B = (\<lambda>(t,k). (Inum (real i#bs) t,k)) ` set (\<rho> q)) \<and> (\<delta> q = d) \<and> d >0 \<and> iszlfm q (real (i::int)#bs) \<and> (\<forall> (e,c)\<in> set B. numbound0 e \<and> c>0)"
5296 let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
5298 let ?d = "\<delta> ?q"
5299 let ?B = "set (\<rho> ?q)"
5300 let ?f = "\<lambda> (t,k). (simpnum t,k)"
5301 let ?B'= "remdups (map ?f (\<rho> ?q))"
5302 let ?A = "set (\<alpha>_\<rho> ?q)"
5303 let ?A'= "remdups (map ?f (\<alpha>_\<rho> ?q))"
5304 from conjunct1[OF zlfm_I[OF qf, where bs="bs"]]
5305 have pp': "\<forall> i. ?I i ?q = ?I i p" by auto
5306 hence pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by simp
5307 from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real i"]
5308 have lq: "iszlfm ?q (real (i::int)#bs)" .
5309 from \<delta>[OF lq] have dp:"?d >0" by blast
5310 let ?N = "\<lambda> (t,c). (Inum (real (i::int)#bs) t,c)"
5311 have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_compose)
5312 also have "\<dots> = ?N ` ?B"
5313 by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def)
5314 finally have BB': "?N ` set ?B' = ?N ` ?B" .
5315 have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_compose)
5316 also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"]
5317 by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def)
5318 finally have AA': "?N ` set ?A' = ?N ` ?A" .
5319 from \<rho>_l[OF lq] have B_nb:"\<forall> (e,c)\<in> set ?B'. numbound0 e \<and> c > 0"
5320 by (simp add: simpnum_numbound0 split_def)
5321 from \<alpha>_\<rho>_l[OF lq] have A_nb: "\<forall> (e,c)\<in> set ?A'. numbound0 e \<and> c > 0"
5322 by (simp add: simpnum_numbound0 split_def)
5323 {assume "length ?B' \<le> length ?A'"
5324 hence q:"q=?q" and "B = ?B'" and d:"d = ?d"
5325 using qBd by (auto simp add: Let_def chooset_def)
5326 with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<rho> q)"
5327 and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto
5328 with pq_ex dp lq q d have ?thes by simp}
5330 {assume "\<not> (length ?B' \<le> length ?A')"
5331 hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d"
5332 using qBd by (auto simp add: Let_def chooset_def)
5333 with AA' mirror_\<alpha>_\<rho>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<rho> q)"
5334 and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto
5335 from mirror_ex[OF lq] pq_ex q
5336 have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp
5337 from lq q mirror_l [where p="?q" and bs="bs" and a="real i"]
5338 have lq': "iszlfm q (real i#bs)" by auto
5339 from mirror_\<delta>[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp
5341 ultimately show ?thes by blast
5344 definition stage :: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm" where
5345 "stage p d \<equiv> (\<lambda> (e,c). evaldjf (\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))) [1..c*d])"
5347 shows "Ifm bs (stage p d (e,c)) = (\<exists> j\<in>{1 .. c*d}. Ifm bs (\<sigma> p c (Add e (C j))))"
5348 by (unfold stage_def split_def ,simp only: evaldjf_ex simpfm) simp
5350 lemma stage_nb: assumes lp: "iszlfm p (a#bs)" and cp: "c >0" and nb:"numbound0 e"
5351 shows "bound0 (stage p d (e,c))"
5353 let ?f = "\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))"
5354 have th: "\<forall> j\<in> set [1..c*d]. bound0 (?f j)"
5357 from nb have nb':"numbound0 (Add e (C j))" by simp
5358 from simpfm_bound0[OF \<sigma>_nb[OF lp nb', where k="c"]]
5359 show "bound0 (simpfm (\<sigma> p c (Add e (C j))))" .
5361 from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp
5364 definition redlove :: "fm \<Rightarrow> fm" where
5366 (let (q,B,d) = chooset p;
5367 mq = simpfm (minusinf q);
5368 md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) [1..d]
5369 in if md = T then T else
5370 (let qd = evaldjf (stage q d) B
5371 in decr (disj md qd)))"
5373 lemma redlove: assumes qf: "qfree p"
5374 shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (redlove p))) \<and> qfree (redlove p)"
5375 (is "(?lhs = ?rhs) \<and> _")
5378 let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p"
5379 let ?q = "fst (chooset p)"
5380 let ?B = "fst (snd(chooset p))"
5381 let ?d = "snd (snd (chooset p))"
5383 let ?mq = "minusinf ?q"
5384 let ?smq = "simpfm ?mq"
5385 let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js"
5387 let ?N = "\<lambda> (t,k). (Inum (real (i::int)#bs) t,k)"
5388 let ?qd = "evaldjf (stage ?q ?d) ?B"
5389 have qbf:"chooset p = (?q,?B,?d)" by simp
5390 from chooset[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and
5391 B:"?N ` set ?B = ?N ` set (\<rho> ?q)" and dd: "\<delta> ?q = ?d" and dp: "?d > 0" and
5392 lq: "iszlfm ?q (real i#bs)" and
5393 Bn: "\<forall> (e,c)\<in> set ?B. numbound0 e \<and> c > 0" by auto
5394 from zlin_qfree[OF lq] have qfq: "qfree ?q" .
5395 from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq".
5396 have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp
5397 hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)"
5398 by (auto simp only: subst0_bound0[OF qfmq])
5399 hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))"
5400 by (auto simp add: simpfm_bound0)
5401 from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp
5402 from Bn stage_nb[OF lq] have th:"\<forall> x \<in> set ?B. bound0 (stage ?q ?d x)" by auto
5403 from evaldjf_bound0[OF th] have qdb: "bound0 ?qd" .
5405 have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all)
5406 from trans [OF pq_ex rl_thm'[OF lq B]] dd
5407 have "?lhs = ((\<exists> j\<in> {1.. ?d}. ?I j ?mq) \<or> (\<exists> (e,c)\<in> set ?B. \<exists> j\<in> {1 .. c*?d}. Ifm (real i#bs) (\<sigma> ?q c (Add e (C j)))))" by auto
5408 also have "\<dots> = ((\<exists> j\<in> {1.. ?d}. ?I j ?smq) \<or> (\<exists> (e,c)\<in> set ?B. ?I i (stage ?q ?d (e,c) )))"
5409 by (simp add: simpfm stage split_def)
5410 also have "\<dots> = ((\<exists> j\<in> {1 .. ?d}. ?I i (subst0 (C j) ?smq)) \<or> ?I i ?qd)"
5411 by (simp add: evaldjf_ex subst0_I[OF qfmq])
5412 finally have mdqd:"?lhs = (?I i ?md \<or> ?I i ?qd)" by (simp only: evaldjf_ex set_upto simpfm)
5413 also have "\<dots> = (?I i (disj ?md ?qd))" by (simp add: disj)
5414 also have "\<dots> = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb])
5415 finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" .
5416 {assume mdT: "?md = T"
5417 hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def)
5418 from mdT have lhs:"?lhs" using mdqd by simp
5419 from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def)
5420 with lhs cT have ?thesis by simp }
5422 {assume mdT: "?md \<noteq> T" hence "redlove p = decr (disj ?md ?qd)"
5423 by (simp add: redlove_def chooset_def split_def Let_def)
5424 with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp }
5425 ultimately show ?thesis by blast
5429 assumes qf: "qfree p"
5430 shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ redlove p))) \<and> qfree (DJ redlove p)"
5432 from redlove have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (redlove p)" by blast
5433 from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast
5434 have "Ifm bs (DJ redlove p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (redlove q))"
5435 by (simp add: DJ_def evaldjf_ex)
5436 also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs) q)"
5437 using redlove disjuncts_qf[OF qf] by blast
5438 also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto)
5439 finally show ?thesis using thqf by blast
5443 lemma exsplit_qf: assumes qf: "qfree p"
5444 shows "qfree (exsplit p)"
5445 using qf by (induct p rule: exsplit.induct, auto)
5447 definition mircfr :: "fm \<Rightarrow> fm" where
5448 "mircfr = DJ cooper o ferrack01 o simpfm o exsplit"
5450 definition mirlfr :: "fm \<Rightarrow> fm" where
5451 "mirlfr = DJ redlove o ferrack01 o simpfm o exsplit"
5453 lemma mircfr: "\<forall> bs p. qfree p \<longrightarrow> qfree (mircfr p) \<and> Ifm bs (mircfr p) = Ifm bs (E p)"
5454 proof(clarsimp simp del: Ifm.simps)
5456 assume qf: "qfree p"
5457 show "qfree (mircfr p)\<and>(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)")
5459 let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
5460 have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)"
5461 using splitex[OF qf] by simp
5462 with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\<exists> (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+
5463 with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def)
5467 lemma mirlfr: "\<forall> bs p. qfree p \<longrightarrow> qfree(mirlfr p) \<and> Ifm bs (mirlfr p) = Ifm bs (E p)"
5468 proof(clarsimp simp del: Ifm.simps)
5470 assume qf: "qfree p"
5471 show "qfree (mirlfr p)\<and>(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)")
5473 let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))"
5474 have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)"
5475 using splitex[OF qf] by simp
5476 with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\<exists> (i::int). Ifm (real i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+
5477 with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def)
5481 definition mircfrqe:: "fm \<Rightarrow> fm" where
5482 "mircfrqe p = qelim (prep p) mircfr"
5484 definition mirlfrqe:: "fm \<Rightarrow> fm" where
5485 "mirlfrqe p = qelim (prep p) mirlfr"
5487 theorem mircfrqe: "(Ifm bs (mircfrqe p) = Ifm bs p) \<and> qfree (mircfrqe p)"
5488 using qelim_ci[OF mircfr] prep by (auto simp add: mircfrqe_def)
5490 theorem mirlfrqe: "(Ifm bs (mirlfrqe p) = Ifm bs p) \<and> qfree (mirlfrqe p)"
5491 using qelim_ci[OF mirlfr] prep by (auto simp add: mirlfrqe_def)
5494 "problem1 = A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0))))))"
5497 "problem2 = A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0))))"
5500 "problem3 = A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0))))))"
5503 "problem4 = E (And (Ge (Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg (Bound 0))))))"
5506 Par_List.map (fn e => e ())
5507 [fn () => @{code mircfrqe} @{code problem1},
5508 fn () => @{code mirlfrqe} @{code problem1},
5509 fn () => @{code mircfrqe} @{code problem2},
5510 fn () => @{code mirlfrqe} @{code problem2},
5511 fn () => @{code mircfrqe} @{code problem3},
5512 fn () => @{code mirlfrqe} @{code problem3},
5513 fn () => @{code mircfrqe} @{code problem4},
5514 fn () => @{code mirlfrqe} @{code problem4}]
5518 functions mircfrqe mirlfrqe
5521 oracle mirfr_oracle = {* fn (proofs, ct) =>
5524 fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t
5525 of NONE => error "Variable not found in the list!"
5526 | SOME n => @{code Bound} n)
5527 | num_of_term vs @{term "real (0::int)"} = @{code C} 0
5528 | num_of_term vs @{term "real (1::int)"} = @{code C} 1
5529 | num_of_term vs @{term "0::real"} = @{code C} 0
5530 | num_of_term vs @{term "1::real"} = @{code C} 1
5531 | num_of_term vs (Bound i) = @{code Bound} i
5532 | num_of_term vs (@{term "uminus :: real \<Rightarrow> real"} $ t') = @{code Neg} (num_of_term vs t')
5533 | num_of_term vs (@{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
5534 @{code Add} (num_of_term vs t1, num_of_term vs t2)
5535 | num_of_term vs (@{term "op - :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
5536 @{code Sub} (num_of_term vs t1, num_of_term vs t2)
5537 | num_of_term vs (@{term "op * :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) =
5538 (case (num_of_term vs t1)
5539 of @{code C} i => @{code Mul} (i, num_of_term vs t2)
5540 | _ => error "num_of_term: unsupported Multiplication")
5541 | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "numeral :: _ \<Rightarrow> int"} $ t')) =
5542 @{code C} (HOLogic.dest_num t')
5543 | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "neg_numeral :: _ \<Rightarrow> int"} $ t')) =
5544 @{code C} (~ (HOLogic.dest_num t'))
5545 | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ t')) =
5546 @{code Floor} (num_of_term vs t')
5547 | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ t')) =
5548 @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t')))
5549 | num_of_term vs (@{term "numeral :: _ \<Rightarrow> real"} $ t') =
5550 @{code C} (HOLogic.dest_num t')
5551 | num_of_term vs (@{term "neg_numeral :: _ \<Rightarrow> real"} $ t') =
5552 @{code C} (~ (HOLogic.dest_num t'))
5553 | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t);
5555 fun fm_of_term vs @{term True} = @{code T}
5556 | fm_of_term vs @{term False} = @{code F}
5557 | fm_of_term vs (@{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =
5558 @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2))
5559 | fm_of_term vs (@{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) =