|
1 (* Title: HOL/Reflection/MIR.thy |
|
2 Author: Amine Chaieb |
|
3 *) |
|
4 |
|
5 theory MIR |
|
6 imports Complex_Main Dense_Linear_Order Efficient_Nat |
|
7 uses ("mir_tac.ML") |
|
8 begin |
|
9 |
|
10 section {* Quantifier elimination for @{text "\<real> (0, 1, +, floor, <)"} *} |
|
11 |
|
12 declare real_of_int_floor_cancel [simp del] |
|
13 |
|
14 primrec alluopairs:: "'a list \<Rightarrow> ('a \<times> 'a) list" where |
|
15 "alluopairs [] = []" |
|
16 | "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)" |
|
17 |
|
18 lemma alluopairs_set1: "set (alluopairs xs) \<le> {(x,y). x\<in> set xs \<and> y\<in> set xs}" |
|
19 by (induct xs, auto) |
|
20 |
|
21 lemma alluopairs_set: |
|
22 "\<lbrakk>x\<in> set xs ; y \<in> set xs\<rbrakk> \<Longrightarrow> (x,y) \<in> set (alluopairs xs) \<or> (y,x) \<in> set (alluopairs xs) " |
|
23 by (induct xs, auto) |
|
24 |
|
25 lemma alluopairs_ex: |
|
26 assumes Pc: "\<forall> x y. P x y = P y x" |
|
27 shows "(\<exists> x \<in> set xs. \<exists> y \<in> set xs. P x y) = (\<exists> (x,y) \<in> set (alluopairs xs). P x y)" |
|
28 proof |
|
29 assume "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y" |
|
30 then obtain x y where x: "x \<in> set xs" and y:"y \<in> set xs" and P: "P x y" by blast |
|
31 from alluopairs_set[OF x y] P Pc show"\<exists>(x, y)\<in>set (alluopairs xs). P x y" |
|
32 by auto |
|
33 next |
|
34 assume "\<exists>(x, y)\<in>set (alluopairs xs). P x y" |
|
35 then obtain "x" and "y" where xy:"(x,y) \<in> set (alluopairs xs)" and P: "P x y" by blast+ |
|
36 from xy have "x \<in> set xs \<and> y\<in> set xs" using alluopairs_set1 by blast |
|
37 with P show "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y" by blast |
|
38 qed |
|
39 |
|
40 (* generate a list from i to j*) |
|
41 consts iupt :: "int \<times> int \<Rightarrow> int list" |
|
42 recdef iupt "measure (\<lambda> (i,j). nat (j-i +1))" |
|
43 "iupt (i,j) = (if j <i then [] else (i# iupt(i+1, j)))" |
|
44 |
|
45 lemma iupt_set: "set (iupt(i,j)) = {i .. j}" |
|
46 proof(induct rule: iupt.induct) |
|
47 case (1 a b) |
|
48 show ?case |
|
49 using prems by (simp add: simp_from_to) |
|
50 qed |
|
51 |
|
52 lemma nth_pos2: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)" |
|
53 using Nat.gr0_conv_Suc |
|
54 by clarsimp |
|
55 |
|
56 |
|
57 lemma myl: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a \<le> b) = (0 \<le> b - a)" |
|
58 proof(clarify) |
|
59 fix x y ::"'a" |
|
60 have "(x \<le> y) = (x - y \<le> 0)" by (simp only: le_iff_diff_le_0[where a="x" and b="y"]) |
|
61 also have "\<dots> = (- (y - x) \<le> 0)" by simp |
|
62 also have "\<dots> = (0 \<le> y - x)" by (simp only: neg_le_0_iff_le[where a="y-x"]) |
|
63 finally show "(x \<le> y) = (0 \<le> y - x)" . |
|
64 qed |
|
65 |
|
66 lemma myless: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a < b) = (0 < b - a)" |
|
67 proof(clarify) |
|
68 fix x y ::"'a" |
|
69 have "(x < y) = (x - y < 0)" by (simp only: less_iff_diff_less_0[where a="x" and b="y"]) |
|
70 also have "\<dots> = (- (y - x) < 0)" by simp |
|
71 also have "\<dots> = (0 < y - x)" by (simp only: neg_less_0_iff_less[where a="y-x"]) |
|
72 finally show "(x < y) = (0 < y - x)" . |
|
73 qed |
|
74 |
|
75 lemma myeq: "\<forall> (a::'a::{pordered_ab_group_add}) (b::'a). (a = b) = (0 = b - a)" |
|
76 by auto |
|
77 |
|
78 (* Maybe should be added to the library \<dots> *) |
|
79 lemma floor_int_eq: "(real n\<le> x \<and> x < real (n+1)) = (floor x = n)" |
|
80 proof( auto) |
|
81 assume lb: "real n \<le> x" |
|
82 and ub: "x < real n + 1" |
|
83 have "real (floor x) \<le> x" by simp |
|
84 hence "real (floor x) < real (n + 1) " using ub by arith |
|
85 hence "floor x < n+1" by simp |
|
86 moreover from lb have "n \<le> floor x" using floor_mono2[where x="real n" and y="x"] |
|
87 by simp ultimately show "floor x = n" by simp |
|
88 qed |
|
89 |
|
90 (* Periodicity of dvd *) |
|
91 lemma dvd_period: |
|
92 assumes advdd: "(a::int) dvd d" |
|
93 shows "(a dvd (x + t)) = (a dvd ((x+ c*d) + t))" |
|
94 using advdd |
|
95 proof- |
|
96 {fix x k |
|
97 from inf_period(3)[OF advdd, rule_format, where x=x and k="-k"] |
|
98 have " ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp} |
|
99 hence "\<forall>x.\<forall>k. ((a::int) dvd (x + t)) = (a dvd (x+k*d + t))" by simp |
|
100 then show ?thesis by simp |
|
101 qed |
|
102 |
|
103 (* The Divisibility relation between reals *) |
|
104 definition |
|
105 rdvd:: "real \<Rightarrow> real \<Rightarrow> bool" (infixl "rdvd" 50) |
|
106 where |
|
107 rdvd_def: "x rdvd y \<longleftrightarrow> (\<exists>k\<Colon>int. y = x * real k)" |
|
108 |
|
109 lemma int_rdvd_real: |
|
110 shows "real (i::int) rdvd x = (i dvd (floor x) \<and> real (floor x) = x)" (is "?l = ?r") |
|
111 proof |
|
112 assume "?l" |
|
113 hence th: "\<exists> k. x=real (i*k)" by (simp add: rdvd_def) |
|
114 hence th': "real (floor x) = x" by (auto simp del: real_of_int_mult) |
|
115 with th have "\<exists> k. real (floor x) = real (i*k)" by simp |
|
116 hence "\<exists> k. floor x = i*k" by (simp only: real_of_int_inject) |
|
117 thus ?r using th' by (simp add: dvd_def) |
|
118 next |
|
119 assume "?r" hence "(i\<Colon>int) dvd \<lfloor>x\<Colon>real\<rfloor>" .. |
|
120 hence "\<exists> k. real (floor x) = real (i*k)" |
|
121 by (simp only: real_of_int_inject) (simp add: dvd_def) |
|
122 thus ?l using prems by (simp add: rdvd_def) |
|
123 qed |
|
124 |
|
125 lemma int_rdvd_iff: "(real (i::int) rdvd real t) = (i dvd t)" |
|
126 by (auto simp add: rdvd_def dvd_def) (rule_tac x="k" in exI, simp only :real_of_int_mult[symmetric]) |
|
127 |
|
128 |
|
129 lemma rdvd_abs1: |
|
130 "(abs (real d) rdvd t) = (real (d ::int) rdvd t)" |
|
131 proof |
|
132 assume d: "real d rdvd t" |
|
133 from d int_rdvd_real have d2: "d dvd (floor t)" and ti: "real (floor t) = t" by auto |
|
134 |
|
135 from iffD2[OF zdvd_abs1] d2 have "(abs d) dvd (floor t)" by blast |
|
136 with ti int_rdvd_real[symmetric] have "real (abs d) rdvd t" by blast |
|
137 thus "abs (real d) rdvd t" by simp |
|
138 next |
|
139 assume "abs (real d) rdvd t" hence "real (abs d) rdvd t" by simp |
|
140 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 |
|
141 from iffD1[OF zdvd_abs1] d2 have "d dvd floor t" by blast |
|
142 with ti int_rdvd_real[symmetric] show "real d rdvd t" by blast |
|
143 qed |
|
144 |
|
145 lemma rdvd_minus: "(real (d::int) rdvd t) = (real d rdvd -t)" |
|
146 apply (auto simp add: rdvd_def) |
|
147 apply (rule_tac x="-k" in exI, simp) |
|
148 apply (rule_tac x="-k" in exI, simp) |
|
149 done |
|
150 |
|
151 lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)" |
|
152 by (auto simp add: rdvd_def) |
|
153 |
|
154 lemma rdvd_mult: |
|
155 assumes knz: "k\<noteq>0" |
|
156 shows "(real (n::int) * real (k::int) rdvd x * real k) = (real n rdvd x)" |
|
157 using knz by (simp add:rdvd_def) |
|
158 |
|
159 lemma rdvd_trans: assumes mn:"m rdvd n" and nk:"n rdvd k" |
|
160 shows "m rdvd k" |
|
161 proof- |
|
162 from rdvd_def mn obtain c where nmc:"n = m * real (c::int)" by auto |
|
163 from rdvd_def nk obtain c' where nkc:"k = n * real (c'::int)" by auto |
|
164 hence "k = m * real (c * c')" using nmc by simp |
|
165 thus ?thesis using rdvd_def by blast |
|
166 qed |
|
167 |
|
168 (*********************************************************************************) |
|
169 (**** SHADOW SYNTAX AND SEMANTICS ****) |
|
170 (*********************************************************************************) |
|
171 |
|
172 datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num |
|
173 | Mul int num | Floor num| CF int num num |
|
174 |
|
175 (* A size for num to make inductive proofs simpler*) |
|
176 primrec num_size :: "num \<Rightarrow> nat" where |
|
177 "num_size (C c) = 1" |
|
178 | "num_size (Bound n) = 1" |
|
179 | "num_size (Neg a) = 1 + num_size a" |
|
180 | "num_size (Add a b) = 1 + num_size a + num_size b" |
|
181 | "num_size (Sub a b) = 3 + num_size a + num_size b" |
|
182 | "num_size (CN n c a) = 4 + num_size a " |
|
183 | "num_size (CF c a b) = 4 + num_size a + num_size b" |
|
184 | "num_size (Mul c a) = 1 + num_size a" |
|
185 | "num_size (Floor a) = 1 + num_size a" |
|
186 |
|
187 (* Semantics of numeral terms (num) *) |
|
188 primrec Inum :: "real list \<Rightarrow> num \<Rightarrow> real" where |
|
189 "Inum bs (C c) = (real c)" |
|
190 | "Inum bs (Bound n) = bs!n" |
|
191 | "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)" |
|
192 | "Inum bs (Neg a) = -(Inum bs a)" |
|
193 | "Inum bs (Add a b) = Inum bs a + Inum bs b" |
|
194 | "Inum bs (Sub a b) = Inum bs a - Inum bs b" |
|
195 | "Inum bs (Mul c a) = (real c) * Inum bs a" |
|
196 | "Inum bs (Floor a) = real (floor (Inum bs a))" |
|
197 | "Inum bs (CF c a b) = real c * real (floor (Inum bs a)) + Inum bs b" |
|
198 definition "isint t bs \<equiv> real (floor (Inum bs t)) = Inum bs t" |
|
199 |
|
200 lemma isint_iff: "isint n bs = (real (floor (Inum bs n)) = Inum bs n)" |
|
201 by (simp add: isint_def) |
|
202 |
|
203 lemma isint_Floor: "isint (Floor n) bs" |
|
204 by (simp add: isint_iff) |
|
205 |
|
206 lemma isint_Mul: "isint e bs \<Longrightarrow> isint (Mul c e) bs" |
|
207 proof- |
|
208 let ?e = "Inum bs e" |
|
209 let ?fe = "floor ?e" |
|
210 assume be: "isint e bs" hence efe:"real ?fe = ?e" by (simp add: isint_iff) |
|
211 have "real ((floor (Inum bs (Mul c e)))) = real (floor (real (c * ?fe)))" using efe by simp |
|
212 also have "\<dots> = real (c* ?fe)" by (simp only: floor_real_of_int) |
|
213 also have "\<dots> = real c * ?e" using efe by simp |
|
214 finally show ?thesis using isint_iff by simp |
|
215 qed |
|
216 |
|
217 lemma isint_neg: "isint e bs \<Longrightarrow> isint (Neg e) bs" |
|
218 proof- |
|
219 let ?I = "\<lambda> t. Inum bs t" |
|
220 assume ie: "isint e bs" |
|
221 hence th: "real (floor (?I e)) = ?I e" by (simp add: isint_def) |
|
222 have "real (floor (?I (Neg e))) = real (floor (- (real (floor (?I e)))))" by (simp add: th) |
|
223 also have "\<dots> = - real (floor (?I e))" by(simp add: floor_minus_real_of_int) |
|
224 finally show "isint (Neg e) bs" by (simp add: isint_def th) |
|
225 qed |
|
226 |
|
227 lemma isint_sub: |
|
228 assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs" |
|
229 proof- |
|
230 let ?I = "\<lambda> t. Inum bs t" |
|
231 from ie have th: "real (floor (?I e)) = ?I e" by (simp add: isint_def) |
|
232 have "real (floor (?I (Sub (C c) e))) = real (floor ((real (c -floor (?I e)))))" by (simp add: th) |
|
233 also have "\<dots> = real (c- floor (?I e))" by(simp add: floor_minus_real_of_int) |
|
234 finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th) |
|
235 qed |
|
236 |
|
237 lemma isint_add: assumes |
|
238 ai:"isint a bs" and bi: "isint b bs" shows "isint (Add a b) bs" |
|
239 proof- |
|
240 let ?a = "Inum bs a" |
|
241 let ?b = "Inum bs b" |
|
242 from ai bi isint_iff have "real (floor (?a + ?b)) = real (floor (real (floor ?a) + real (floor ?b)))" by simp |
|
243 also have "\<dots> = real (floor ?a) + real (floor ?b)" by simp |
|
244 also have "\<dots> = ?a + ?b" using ai bi isint_iff by simp |
|
245 finally show "isint (Add a b) bs" by (simp add: isint_iff) |
|
246 qed |
|
247 |
|
248 lemma isint_c: "isint (C j) bs" |
|
249 by (simp add: isint_iff) |
|
250 |
|
251 |
|
252 (* FORMULAE *) |
|
253 datatype fm = |
|
254 T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num| |
|
255 NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm |
|
256 |
|
257 |
|
258 (* A size for fm *) |
|
259 fun fmsize :: "fm \<Rightarrow> nat" where |
|
260 "fmsize (NOT p) = 1 + fmsize p" |
|
261 | "fmsize (And p q) = 1 + fmsize p + fmsize q" |
|
262 | "fmsize (Or p q) = 1 + fmsize p + fmsize q" |
|
263 | "fmsize (Imp p q) = 3 + fmsize p + fmsize q" |
|
264 | "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)" |
|
265 | "fmsize (E p) = 1 + fmsize p" |
|
266 | "fmsize (A p) = 4+ fmsize p" |
|
267 | "fmsize (Dvd i t) = 2" |
|
268 | "fmsize (NDvd i t) = 2" |
|
269 | "fmsize p = 1" |
|
270 (* several lemmas about fmsize *) |
|
271 lemma fmsize_pos: "fmsize p > 0" |
|
272 by (induct p rule: fmsize.induct) simp_all |
|
273 |
|
274 (* Semantics of formulae (fm) *) |
|
275 primrec Ifm ::"real list \<Rightarrow> fm \<Rightarrow> bool" where |
|
276 "Ifm bs T = True" |
|
277 | "Ifm bs F = False" |
|
278 | "Ifm bs (Lt a) = (Inum bs a < 0)" |
|
279 | "Ifm bs (Gt a) = (Inum bs a > 0)" |
|
280 | "Ifm bs (Le a) = (Inum bs a \<le> 0)" |
|
281 | "Ifm bs (Ge a) = (Inum bs a \<ge> 0)" |
|
282 | "Ifm bs (Eq a) = (Inum bs a = 0)" |
|
283 | "Ifm bs (NEq a) = (Inum bs a \<noteq> 0)" |
|
284 | "Ifm bs (Dvd i b) = (real i rdvd Inum bs b)" |
|
285 | "Ifm bs (NDvd i b) = (\<not>(real i rdvd Inum bs b))" |
|
286 | "Ifm bs (NOT p) = (\<not> (Ifm bs p))" |
|
287 | "Ifm bs (And p q) = (Ifm bs p \<and> Ifm bs q)" |
|
288 | "Ifm bs (Or p q) = (Ifm bs p \<or> Ifm bs q)" |
|
289 | "Ifm bs (Imp p q) = ((Ifm bs p) \<longrightarrow> (Ifm bs q))" |
|
290 | "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)" |
|
291 | "Ifm bs (E p) = (\<exists> x. Ifm (x#bs) p)" |
|
292 | "Ifm bs (A p) = (\<forall> x. Ifm (x#bs) p)" |
|
293 |
|
294 consts prep :: "fm \<Rightarrow> fm" |
|
295 recdef prep "measure fmsize" |
|
296 "prep (E T) = T" |
|
297 "prep (E F) = F" |
|
298 "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))" |
|
299 "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))" |
|
300 "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" |
|
301 "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))" |
|
302 "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))" |
|
303 "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" |
|
304 "prep (E p) = E (prep p)" |
|
305 "prep (A (And p q)) = And (prep (A p)) (prep (A q))" |
|
306 "prep (A p) = prep (NOT (E (NOT p)))" |
|
307 "prep (NOT (NOT p)) = prep p" |
|
308 "prep (NOT (And p q)) = Or (prep (NOT p)) (prep (NOT q))" |
|
309 "prep (NOT (A p)) = prep (E (NOT p))" |
|
310 "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))" |
|
311 "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))" |
|
312 "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))" |
|
313 "prep (NOT p) = NOT (prep p)" |
|
314 "prep (Or p q) = Or (prep p) (prep q)" |
|
315 "prep (And p q) = And (prep p) (prep q)" |
|
316 "prep (Imp p q) = prep (Or (NOT p) q)" |
|
317 "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))" |
|
318 "prep p = p" |
|
319 (hints simp add: fmsize_pos) |
|
320 lemma prep: "\<And> bs. Ifm bs (prep p) = Ifm bs p" |
|
321 by (induct p rule: prep.induct, auto) |
|
322 |
|
323 |
|
324 (* Quantifier freeness *) |
|
325 fun qfree:: "fm \<Rightarrow> bool" where |
|
326 "qfree (E p) = False" |
|
327 | "qfree (A p) = False" |
|
328 | "qfree (NOT p) = qfree p" |
|
329 | "qfree (And p q) = (qfree p \<and> qfree q)" |
|
330 | "qfree (Or p q) = (qfree p \<and> qfree q)" |
|
331 | "qfree (Imp p q) = (qfree p \<and> qfree q)" |
|
332 | "qfree (Iff p q) = (qfree p \<and> qfree q)" |
|
333 | "qfree p = True" |
|
334 |
|
335 (* Boundedness and substitution *) |
|
336 primrec numbound0 :: "num \<Rightarrow> bool" (* a num is INDEPENDENT of Bound 0 *) where |
|
337 "numbound0 (C c) = True" |
|
338 | "numbound0 (Bound n) = (n>0)" |
|
339 | "numbound0 (CN n i a) = (n > 0 \<and> numbound0 a)" |
|
340 | "numbound0 (Neg a) = numbound0 a" |
|
341 | "numbound0 (Add a b) = (numbound0 a \<and> numbound0 b)" |
|
342 | "numbound0 (Sub a b) = (numbound0 a \<and> numbound0 b)" |
|
343 | "numbound0 (Mul i a) = numbound0 a" |
|
344 | "numbound0 (Floor a) = numbound0 a" |
|
345 | "numbound0 (CF c a b) = (numbound0 a \<and> numbound0 b)" |
|
346 |
|
347 lemma numbound0_I: |
|
348 assumes nb: "numbound0 a" |
|
349 shows "Inum (b#bs) a = Inum (b'#bs) a" |
|
350 using nb by (induct a) (auto simp add: nth_pos2) |
|
351 |
|
352 lemma numbound0_gen: |
|
353 assumes nb: "numbound0 t" and ti: "isint t (x#bs)" |
|
354 shows "\<forall> y. isint t (y#bs)" |
|
355 using nb ti |
|
356 proof(clarify) |
|
357 fix y |
|
358 from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def] |
|
359 show "isint t (y#bs)" |
|
360 by (simp add: isint_def) |
|
361 qed |
|
362 |
|
363 primrec bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *) where |
|
364 "bound0 T = True" |
|
365 | "bound0 F = True" |
|
366 | "bound0 (Lt a) = numbound0 a" |
|
367 | "bound0 (Le a) = numbound0 a" |
|
368 | "bound0 (Gt a) = numbound0 a" |
|
369 | "bound0 (Ge a) = numbound0 a" |
|
370 | "bound0 (Eq a) = numbound0 a" |
|
371 | "bound0 (NEq a) = numbound0 a" |
|
372 | "bound0 (Dvd i a) = numbound0 a" |
|
373 | "bound0 (NDvd i a) = numbound0 a" |
|
374 | "bound0 (NOT p) = bound0 p" |
|
375 | "bound0 (And p q) = (bound0 p \<and> bound0 q)" |
|
376 | "bound0 (Or p q) = (bound0 p \<and> bound0 q)" |
|
377 | "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))" |
|
378 | "bound0 (Iff p q) = (bound0 p \<and> bound0 q)" |
|
379 | "bound0 (E p) = False" |
|
380 | "bound0 (A p) = False" |
|
381 |
|
382 lemma bound0_I: |
|
383 assumes bp: "bound0 p" |
|
384 shows "Ifm (b#bs) p = Ifm (b'#bs) p" |
|
385 using bp numbound0_I [where b="b" and bs="bs" and b'="b'"] |
|
386 by (induct p) (auto simp add: nth_pos2) |
|
387 |
|
388 primrec numsubst0:: "num \<Rightarrow> num \<Rightarrow> num" (* substitute a num into a num for Bound 0 *) where |
|
389 "numsubst0 t (C c) = (C c)" |
|
390 | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)" |
|
391 | "numsubst0 t (CN n i a) = (if n=0 then Add (Mul i t) (numsubst0 t a) else CN n i (numsubst0 t a))" |
|
392 | "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)" |
|
393 | "numsubst0 t (Neg a) = Neg (numsubst0 t a)" |
|
394 | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)" |
|
395 | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" |
|
396 | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)" |
|
397 | "numsubst0 t (Floor a) = Floor (numsubst0 t a)" |
|
398 |
|
399 lemma numsubst0_I: |
|
400 shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t" |
|
401 by (induct t) (simp_all add: nth_pos2) |
|
402 |
|
403 lemma numsubst0_I': |
|
404 assumes nb: "numbound0 a" |
|
405 shows "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t" |
|
406 by (induct t) (simp_all add: nth_pos2 numbound0_I[OF nb, where b="b" and b'="b'"]) |
|
407 |
|
408 primrec subst0:: "num \<Rightarrow> fm \<Rightarrow> fm" (* substitue a num into a formula for Bound 0 *) where |
|
409 "subst0 t T = T" |
|
410 | "subst0 t F = F" |
|
411 | "subst0 t (Lt a) = Lt (numsubst0 t a)" |
|
412 | "subst0 t (Le a) = Le (numsubst0 t a)" |
|
413 | "subst0 t (Gt a) = Gt (numsubst0 t a)" |
|
414 | "subst0 t (Ge a) = Ge (numsubst0 t a)" |
|
415 | "subst0 t (Eq a) = Eq (numsubst0 t a)" |
|
416 | "subst0 t (NEq a) = NEq (numsubst0 t a)" |
|
417 | "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)" |
|
418 | "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)" |
|
419 | "subst0 t (NOT p) = NOT (subst0 t p)" |
|
420 | "subst0 t (And p q) = And (subst0 t p) (subst0 t q)" |
|
421 | "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)" |
|
422 | "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)" |
|
423 | "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)" |
|
424 |
|
425 lemma subst0_I: assumes qfp: "qfree p" |
|
426 shows "Ifm (b#bs) (subst0 a p) = Ifm ((Inum (b#bs) a)#bs) p" |
|
427 using qfp numsubst0_I[where b="b" and bs="bs" and a="a"] |
|
428 by (induct p) (simp_all add: nth_pos2 ) |
|
429 |
|
430 consts |
|
431 decrnum:: "num \<Rightarrow> num" |
|
432 decr :: "fm \<Rightarrow> fm" |
|
433 |
|
434 recdef decrnum "measure size" |
|
435 "decrnum (Bound n) = Bound (n - 1)" |
|
436 "decrnum (Neg a) = Neg (decrnum a)" |
|
437 "decrnum (Add a b) = Add (decrnum a) (decrnum b)" |
|
438 "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" |
|
439 "decrnum (Mul c a) = Mul c (decrnum a)" |
|
440 "decrnum (Floor a) = Floor (decrnum a)" |
|
441 "decrnum (CN n c a) = CN (n - 1) c (decrnum a)" |
|
442 "decrnum (CF c a b) = CF c (decrnum a) (decrnum b)" |
|
443 "decrnum a = a" |
|
444 |
|
445 recdef decr "measure size" |
|
446 "decr (Lt a) = Lt (decrnum a)" |
|
447 "decr (Le a) = Le (decrnum a)" |
|
448 "decr (Gt a) = Gt (decrnum a)" |
|
449 "decr (Ge a) = Ge (decrnum a)" |
|
450 "decr (Eq a) = Eq (decrnum a)" |
|
451 "decr (NEq a) = NEq (decrnum a)" |
|
452 "decr (Dvd i a) = Dvd i (decrnum a)" |
|
453 "decr (NDvd i a) = NDvd i (decrnum a)" |
|
454 "decr (NOT p) = NOT (decr p)" |
|
455 "decr (And p q) = And (decr p) (decr q)" |
|
456 "decr (Or p q) = Or (decr p) (decr q)" |
|
457 "decr (Imp p q) = Imp (decr p) (decr q)" |
|
458 "decr (Iff p q) = Iff (decr p) (decr q)" |
|
459 "decr p = p" |
|
460 |
|
461 lemma decrnum: assumes nb: "numbound0 t" |
|
462 shows "Inum (x#bs) t = Inum bs (decrnum t)" |
|
463 using nb by (induct t rule: decrnum.induct, simp_all add: nth_pos2) |
|
464 |
|
465 lemma decr: assumes nb: "bound0 p" |
|
466 shows "Ifm (x#bs) p = Ifm bs (decr p)" |
|
467 using nb |
|
468 by (induct p rule: decr.induct, simp_all add: nth_pos2 decrnum) |
|
469 |
|
470 lemma decr_qf: "bound0 p \<Longrightarrow> qfree (decr p)" |
|
471 by (induct p, simp_all) |
|
472 |
|
473 consts |
|
474 isatom :: "fm \<Rightarrow> bool" (* test for atomicity *) |
|
475 recdef isatom "measure size" |
|
476 "isatom T = True" |
|
477 "isatom F = True" |
|
478 "isatom (Lt a) = True" |
|
479 "isatom (Le a) = True" |
|
480 "isatom (Gt a) = True" |
|
481 "isatom (Ge a) = True" |
|
482 "isatom (Eq a) = True" |
|
483 "isatom (NEq a) = True" |
|
484 "isatom (Dvd i b) = True" |
|
485 "isatom (NDvd i b) = True" |
|
486 "isatom p = False" |
|
487 |
|
488 lemma numsubst0_numbound0: assumes nb: "numbound0 t" |
|
489 shows "numbound0 (numsubst0 t a)" |
|
490 using nb by (induct a, auto) |
|
491 |
|
492 lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t" |
|
493 shows "bound0 (subst0 t p)" |
|
494 using qf numsubst0_numbound0[OF nb] by (induct p, auto) |
|
495 |
|
496 lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p" |
|
497 by (induct p, simp_all) |
|
498 |
|
499 |
|
500 definition djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where |
|
501 "djf f p q = (if q=T then T else if q=F then f p else |
|
502 (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or fp q))" |
|
503 |
|
504 definition evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where |
|
505 "evaldjf f ps = foldr (djf f) ps F" |
|
506 |
|
507 lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)" |
|
508 by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) |
|
509 (cases "f p", simp_all add: Let_def djf_def) |
|
510 |
|
511 lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm bs (f p))" |
|
512 by(induct ps, simp_all add: evaldjf_def djf_Or) |
|
513 |
|
514 lemma evaldjf_bound0: |
|
515 assumes nb: "\<forall> x\<in> set xs. bound0 (f x)" |
|
516 shows "bound0 (evaldjf f xs)" |
|
517 using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) |
|
518 |
|
519 lemma evaldjf_qf: |
|
520 assumes nb: "\<forall> x\<in> set xs. qfree (f x)" |
|
521 shows "qfree (evaldjf f xs)" |
|
522 using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) |
|
523 |
|
524 consts |
|
525 disjuncts :: "fm \<Rightarrow> fm list" |
|
526 conjuncts :: "fm \<Rightarrow> fm list" |
|
527 recdef disjuncts "measure size" |
|
528 "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)" |
|
529 "disjuncts F = []" |
|
530 "disjuncts p = [p]" |
|
531 |
|
532 recdef conjuncts "measure size" |
|
533 "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)" |
|
534 "conjuncts T = []" |
|
535 "conjuncts p = [p]" |
|
536 lemma disjuncts: "(\<exists> q\<in> set (disjuncts p). Ifm bs q) = Ifm bs p" |
|
537 by(induct p rule: disjuncts.induct, auto) |
|
538 lemma conjuncts: "(\<forall> q\<in> set (conjuncts p). Ifm bs q) = Ifm bs p" |
|
539 by(induct p rule: conjuncts.induct, auto) |
|
540 |
|
541 lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). bound0 q" |
|
542 proof- |
|
543 assume nb: "bound0 p" |
|
544 hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto) |
|
545 thus ?thesis by (simp only: list_all_iff) |
|
546 qed |
|
547 lemma conjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). bound0 q" |
|
548 proof- |
|
549 assume nb: "bound0 p" |
|
550 hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto) |
|
551 thus ?thesis by (simp only: list_all_iff) |
|
552 qed |
|
553 |
|
554 lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). qfree q" |
|
555 proof- |
|
556 assume qf: "qfree p" |
|
557 hence "list_all qfree (disjuncts p)" |
|
558 by (induct p rule: disjuncts.induct, auto) |
|
559 thus ?thesis by (simp only: list_all_iff) |
|
560 qed |
|
561 lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). qfree q" |
|
562 proof- |
|
563 assume qf: "qfree p" |
|
564 hence "list_all qfree (conjuncts p)" |
|
565 by (induct p rule: conjuncts.induct, auto) |
|
566 thus ?thesis by (simp only: list_all_iff) |
|
567 qed |
|
568 |
|
569 constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" |
|
570 "DJ f p \<equiv> evaldjf f (disjuncts p)" |
|
571 |
|
572 lemma DJ: assumes fdj: "\<forall> p q. f (Or p q) = Or (f p) (f q)" |
|
573 and fF: "f F = F" |
|
574 shows "Ifm bs (DJ f p) = Ifm bs (f p)" |
|
575 proof- |
|
576 have "Ifm bs (DJ f p) = (\<exists> q \<in> set (disjuncts p). Ifm bs (f q))" |
|
577 by (simp add: DJ_def evaldjf_ex) |
|
578 also have "\<dots> = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) |
|
579 finally show ?thesis . |
|
580 qed |
|
581 |
|
582 lemma DJ_qf: assumes |
|
583 fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)" |
|
584 shows "\<forall>p. qfree p \<longrightarrow> qfree (DJ f p) " |
|
585 proof(clarify) |
|
586 fix p assume qf: "qfree p" |
|
587 have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def) |
|
588 from disjuncts_qf[OF qf] have "\<forall> q\<in> set (disjuncts p). qfree q" . |
|
589 with fqf have th':"\<forall> q\<in> set (disjuncts p). qfree (f q)" by blast |
|
590 |
|
591 from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp |
|
592 qed |
|
593 |
|
594 lemma DJ_qe: assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))" |
|
595 shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (Ifm bs ((DJ qe p)) = Ifm bs (E p))" |
|
596 proof(clarify) |
|
597 fix p::fm and bs |
|
598 assume qf: "qfree p" |
|
599 from qe have qth: "\<forall> p. qfree p \<longrightarrow> qfree (qe p)" by blast |
|
600 from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto |
|
601 have "Ifm bs (DJ qe p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (qe q))" |
|
602 by (simp add: DJ_def evaldjf_ex) |
|
603 also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto |
|
604 also have "\<dots> = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto) |
|
605 finally show "qfree (DJ qe p) \<and> Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast |
|
606 qed |
|
607 (* Simplification *) |
|
608 |
|
609 (* Algebraic simplifications for nums *) |
|
610 consts bnds:: "num \<Rightarrow> nat list" |
|
611 lex_ns:: "nat list \<times> nat list \<Rightarrow> bool" |
|
612 recdef bnds "measure size" |
|
613 "bnds (Bound n) = [n]" |
|
614 "bnds (CN n c a) = n#(bnds a)" |
|
615 "bnds (Neg a) = bnds a" |
|
616 "bnds (Add a b) = (bnds a)@(bnds b)" |
|
617 "bnds (Sub a b) = (bnds a)@(bnds b)" |
|
618 "bnds (Mul i a) = bnds a" |
|
619 "bnds (Floor a) = bnds a" |
|
620 "bnds (CF c a b) = (bnds a)@(bnds b)" |
|
621 "bnds a = []" |
|
622 recdef lex_ns "measure (\<lambda> (xs,ys). length xs + length ys)" |
|
623 "lex_ns ([], ms) = True" |
|
624 "lex_ns (ns, []) = False" |
|
625 "lex_ns (n#ns, m#ms) = (n<m \<or> ((n = m) \<and> lex_ns (ns,ms))) " |
|
626 constdefs lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" |
|
627 "lex_bnd t s \<equiv> lex_ns (bnds t, bnds s)" |
|
628 |
|
629 consts |
|
630 numgcdh:: "num \<Rightarrow> int \<Rightarrow> int" |
|
631 reducecoeffh:: "num \<Rightarrow> int \<Rightarrow> num" |
|
632 dvdnumcoeff:: "num \<Rightarrow> int \<Rightarrow> bool" |
|
633 consts maxcoeff:: "num \<Rightarrow> int" |
|
634 recdef maxcoeff "measure size" |
|
635 "maxcoeff (C i) = abs i" |
|
636 "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)" |
|
637 "maxcoeff (CF c t s) = max (abs c) (maxcoeff s)" |
|
638 "maxcoeff t = 1" |
|
639 |
|
640 lemma maxcoeff_pos: "maxcoeff t \<ge> 0" |
|
641 apply (induct t rule: maxcoeff.induct, auto) |
|
642 done |
|
643 |
|
644 recdef numgcdh "measure size" |
|
645 "numgcdh (C i) = (\<lambda>g. zgcd i g)" |
|
646 "numgcdh (CN n c t) = (\<lambda>g. zgcd c (numgcdh t g))" |
|
647 "numgcdh (CF c s t) = (\<lambda>g. zgcd c (numgcdh t g))" |
|
648 "numgcdh t = (\<lambda>g. 1)" |
|
649 |
|
650 definition |
|
651 numgcd :: "num \<Rightarrow> int" |
|
652 where |
|
653 numgcd_def: "numgcd t = numgcdh t (maxcoeff t)" |
|
654 |
|
655 recdef reducecoeffh "measure size" |
|
656 "reducecoeffh (C i) = (\<lambda> g. C (i div g))" |
|
657 "reducecoeffh (CN n c t) = (\<lambda> g. CN n (c div g) (reducecoeffh t g))" |
|
658 "reducecoeffh (CF c s t) = (\<lambda> g. CF (c div g) s (reducecoeffh t g))" |
|
659 "reducecoeffh t = (\<lambda>g. t)" |
|
660 |
|
661 definition |
|
662 reducecoeff :: "num \<Rightarrow> num" |
|
663 where |
|
664 reducecoeff_def: "reducecoeff t = |
|
665 (let g = numgcd t in |
|
666 if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)" |
|
667 |
|
668 recdef dvdnumcoeff "measure size" |
|
669 "dvdnumcoeff (C i) = (\<lambda> g. g dvd i)" |
|
670 "dvdnumcoeff (CN n c t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))" |
|
671 "dvdnumcoeff (CF c s t) = (\<lambda> g. g dvd c \<and> (dvdnumcoeff t g))" |
|
672 "dvdnumcoeff t = (\<lambda>g. False)" |
|
673 |
|
674 lemma dvdnumcoeff_trans: |
|
675 assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'" |
|
676 shows "dvdnumcoeff t g" |
|
677 using dgt' gdg |
|
678 by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg zdvd_trans[OF gdg]) |
|
679 |
|
680 declare zdvd_trans [trans add] |
|
681 |
|
682 lemma natabs0: "(nat (abs x) = 0) = (x = 0)" |
|
683 by arith |
|
684 |
|
685 lemma numgcd0: |
|
686 assumes g0: "numgcd t = 0" |
|
687 shows "Inum bs t = 0" |
|
688 proof- |
|
689 have "\<And>x. numgcdh t x= 0 \<Longrightarrow> Inum bs t = 0" |
|
690 by (induct t rule: numgcdh.induct, auto simp add: zgcd_def gcd_zero natabs0 max_def maxcoeff_pos) |
|
691 thus ?thesis using g0[simplified numgcd_def] by blast |
|
692 qed |
|
693 |
|
694 lemma numgcdh_pos: assumes gp: "g \<ge> 0" shows "numgcdh t g \<ge> 0" |
|
695 using gp |
|
696 by (induct t rule: numgcdh.induct, auto simp add: zgcd_def) |
|
697 |
|
698 lemma numgcd_pos: "numgcd t \<ge>0" |
|
699 by (simp add: numgcd_def numgcdh_pos maxcoeff_pos) |
|
700 |
|
701 lemma reducecoeffh: |
|
702 assumes gt: "dvdnumcoeff t g" and gp: "g > 0" |
|
703 shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t" |
|
704 using gt |
|
705 proof(induct t rule: reducecoeffh.induct) |
|
706 case (1 i) hence gd: "g dvd i" by simp |
|
707 from gp have gnz: "g \<noteq> 0" by simp |
|
708 from prems show ?case by (simp add: real_of_int_div[OF gnz gd]) |
|
709 next |
|
710 case (2 n c t) hence gd: "g dvd c" by simp |
|
711 from gp have gnz: "g \<noteq> 0" by simp |
|
712 from prems show ?case by (simp add: real_of_int_div[OF gnz gd] algebra_simps) |
|
713 next |
|
714 case (3 c s t) hence gd: "g dvd c" by simp |
|
715 from gp have gnz: "g \<noteq> 0" by simp |
|
716 from prems show ?case by (simp add: real_of_int_div[OF gnz gd] algebra_simps) |
|
717 qed (auto simp add: numgcd_def gp) |
|
718 consts ismaxcoeff:: "num \<Rightarrow> int \<Rightarrow> bool" |
|
719 recdef ismaxcoeff "measure size" |
|
720 "ismaxcoeff (C i) = (\<lambda> x. abs i \<le> x)" |
|
721 "ismaxcoeff (CN n c t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))" |
|
722 "ismaxcoeff (CF c s t) = (\<lambda>x. abs c \<le> x \<and> (ismaxcoeff t x))" |
|
723 "ismaxcoeff t = (\<lambda>x. True)" |
|
724 |
|
725 lemma ismaxcoeff_mono: "ismaxcoeff t c \<Longrightarrow> c \<le> c' \<Longrightarrow> ismaxcoeff t c'" |
|
726 by (induct t rule: ismaxcoeff.induct, auto) |
|
727 |
|
728 lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)" |
|
729 proof (induct t rule: maxcoeff.induct) |
|
730 case (2 n c t) |
|
731 hence H:"ismaxcoeff t (maxcoeff t)" . |
|
732 have thh: "maxcoeff t \<le> max (abs c) (maxcoeff t)" by (simp add: le_maxI2) |
|
733 from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1) |
|
734 next |
|
735 case (3 c t s) |
|
736 hence H1:"ismaxcoeff s (maxcoeff s)" by auto |
|
737 have thh1: "maxcoeff s \<le> max \<bar>c\<bar> (maxcoeff s)" by (simp add: max_def) |
|
738 from ismaxcoeff_mono[OF H1 thh1] show ?case by (simp add: le_maxI1) |
|
739 qed simp_all |
|
740 |
|
741 lemma zgcd_gt1: "zgcd i j > 1 \<Longrightarrow> ((abs i > 1 \<and> abs j > 1) \<or> (abs i = 0 \<and> abs j > 1) \<or> (abs i > 1 \<and> abs j = 0))" |
|
742 apply (unfold zgcd_def) |
|
743 apply (cases "i = 0", simp_all) |
|
744 apply (cases "j = 0", simp_all) |
|
745 apply (cases "abs i = 1", simp_all) |
|
746 apply (cases "abs j = 1", simp_all) |
|
747 apply auto |
|
748 done |
|
749 lemma numgcdh0:"numgcdh t m = 0 \<Longrightarrow> m =0" |
|
750 by (induct t rule: numgcdh.induct, auto simp add:zgcd0) |
|
751 |
|
752 lemma dvdnumcoeff_aux: |
|
753 assumes "ismaxcoeff t m" and mp:"m \<ge> 0" and "numgcdh t m > 1" |
|
754 shows "dvdnumcoeff t (numgcdh t m)" |
|
755 using prems |
|
756 proof(induct t rule: numgcdh.induct) |
|
757 case (2 n c t) |
|
758 let ?g = "numgcdh t m" |
|
759 from prems have th:"zgcd c ?g > 1" by simp |
|
760 from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] |
|
761 have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp |
|
762 moreover {assume "abs c > 1" and gp: "?g > 1" with prems |
|
763 have th: "dvdnumcoeff t ?g" by simp |
|
764 have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) |
|
765 from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)} |
|
766 moreover {assume "abs c = 0 \<and> ?g > 1" |
|
767 with prems have th: "dvdnumcoeff t ?g" by simp |
|
768 have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) |
|
769 from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1) |
|
770 hence ?case by simp } |
|
771 moreover {assume "abs c > 1" and g0:"?g = 0" |
|
772 from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } |
|
773 ultimately show ?case by blast |
|
774 next |
|
775 case (3 c s t) |
|
776 let ?g = "numgcdh t m" |
|
777 from prems have th:"zgcd c ?g > 1" by simp |
|
778 from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] |
|
779 have "(abs c > 1 \<and> ?g > 1) \<or> (abs c = 0 \<and> ?g > 1) \<or> (abs c > 1 \<and> ?g = 0)" by simp |
|
780 moreover {assume "abs c > 1" and gp: "?g > 1" with prems |
|
781 have th: "dvdnumcoeff t ?g" by simp |
|
782 have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) |
|
783 from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1)} |
|
784 moreover {assume "abs c = 0 \<and> ?g > 1" |
|
785 with prems have th: "dvdnumcoeff t ?g" by simp |
|
786 have th': "zgcd c ?g dvd ?g" by (simp add:zgcd_zdvd2) |
|
787 from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: zgcd_zdvd1) |
|
788 hence ?case by simp } |
|
789 moreover {assume "abs c > 1" and g0:"?g = 0" |
|
790 from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } |
|
791 ultimately show ?case by blast |
|
792 qed(auto simp add: zgcd_zdvd1) |
|
793 |
|
794 lemma dvdnumcoeff_aux2: |
|
795 assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \<and> numgcd t > 0" |
|
796 using prems |
|
797 proof (simp add: numgcd_def) |
|
798 let ?mc = "maxcoeff t" |
|
799 let ?g = "numgcdh t ?mc" |
|
800 have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff) |
|
801 have th2: "?mc \<ge> 0" by (rule maxcoeff_pos) |
|
802 assume H: "numgcdh t ?mc > 1" |
|
803 from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . |
|
804 qed |
|
805 |
|
806 lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" |
|
807 proof- |
|
808 let ?g = "numgcd t" |
|
809 have "?g \<ge> 0" by (simp add: numgcd_pos) |
|
810 hence "?g = 0 \<or> ?g = 1 \<or> ?g > 1" by auto |
|
811 moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} |
|
812 moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} |
|
813 moreover { assume g1:"?g > 1" |
|
814 from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+ |
|
815 from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis |
|
816 by (simp add: reducecoeff_def Let_def)} |
|
817 ultimately show ?thesis by blast |
|
818 qed |
|
819 |
|
820 lemma reducecoeffh_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeffh t g)" |
|
821 by (induct t rule: reducecoeffh.induct, auto) |
|
822 |
|
823 lemma reducecoeff_numbound0: "numbound0 t \<Longrightarrow> numbound0 (reducecoeff t)" |
|
824 using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def) |
|
825 |
|
826 consts |
|
827 simpnum:: "num \<Rightarrow> num" |
|
828 numadd:: "num \<times> num \<Rightarrow> num" |
|
829 nummul:: "num \<Rightarrow> int \<Rightarrow> num" |
|
830 |
|
831 recdef numadd "measure (\<lambda> (t,s). size t + size s)" |
|
832 "numadd (CN n1 c1 r1,CN n2 c2 r2) = |
|
833 (if n1=n2 then |
|
834 (let c = c1 + c2 |
|
835 in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2)))) |
|
836 else if n1 \<le> n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2)) |
|
837 else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))" |
|
838 "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))" |
|
839 "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" |
|
840 "numadd (CF c1 t1 r1,CF c2 t2 r2) = |
|
841 (if t1 = t2 then |
|
842 (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s)) |
|
843 else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2)) |
|
844 else CF c2 t2 (numadd(CF c1 t1 r1,r2)))" |
|
845 "numadd (CF c1 t1 r1,C c) = CF c1 t1 (numadd (r1, C c))" |
|
846 "numadd (C c,CF c1 t1 r1) = CF c1 t1 (numadd (r1, C c))" |
|
847 "numadd (C b1, C b2) = C (b1+b2)" |
|
848 "numadd (a,b) = Add a b" |
|
849 |
|
850 lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)" |
|
851 apply (induct t s rule: numadd.induct, simp_all add: Let_def) |
|
852 apply (case_tac "c1+c2 = 0",case_tac "n1 \<le> n2", simp_all) |
|
853 apply (case_tac "n1 = n2", simp_all add: algebra_simps) |
|
854 apply (simp only: left_distrib[symmetric]) |
|
855 apply simp |
|
856 apply (case_tac "lex_bnd t1 t2", simp_all) |
|
857 apply (case_tac "c1+c2 = 0") |
|
858 by (case_tac "t1 = t2", simp_all add: algebra_simps left_distrib[symmetric] real_of_int_mult[symmetric] real_of_int_add[symmetric]del: real_of_int_mult real_of_int_add left_distrib) |
|
859 |
|
860 lemma numadd_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numadd (t,s))" |
|
861 by (induct t s rule: numadd.induct, auto simp add: Let_def) |
|
862 |
|
863 recdef nummul "measure size" |
|
864 "nummul (C j) = (\<lambda> i. C (i*j))" |
|
865 "nummul (CN n c t) = (\<lambda> i. CN n (c*i) (nummul t i))" |
|
866 "nummul (CF c t s) = (\<lambda> i. CF (c*i) t (nummul s i))" |
|
867 "nummul (Mul c t) = (\<lambda> i. nummul t (i*c))" |
|
868 "nummul t = (\<lambda> i. Mul i t)" |
|
869 |
|
870 lemma nummul[simp]: "\<And> i. Inum bs (nummul t i) = Inum bs (Mul i t)" |
|
871 by (induct t rule: nummul.induct, auto simp add: algebra_simps) |
|
872 |
|
873 lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)" |
|
874 by (induct t rule: nummul.induct, auto) |
|
875 |
|
876 constdefs numneg :: "num \<Rightarrow> num" |
|
877 "numneg t \<equiv> nummul t (- 1)" |
|
878 |
|
879 constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num" |
|
880 "numsub s t \<equiv> (if s = t then C 0 else numadd (s,numneg t))" |
|
881 |
|
882 lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" |
|
883 using numneg_def nummul by simp |
|
884 |
|
885 lemma numneg_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numneg t)" |
|
886 using numneg_def by simp |
|
887 |
|
888 lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)" |
|
889 using numsub_def by simp |
|
890 |
|
891 lemma numsub_nb[simp]: "\<lbrakk> numbound0 t ; numbound0 s\<rbrakk> \<Longrightarrow> numbound0 (numsub t s)" |
|
892 using numsub_def by simp |
|
893 |
|
894 lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs" |
|
895 proof- |
|
896 have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor) |
|
897 |
|
898 have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def) |
|
899 also have "\<dots>" by (simp add: isint_add cti si) |
|
900 finally show ?thesis . |
|
901 qed |
|
902 |
|
903 consts split_int:: "num \<Rightarrow> num\<times>num" |
|
904 recdef split_int "measure num_size" |
|
905 "split_int (C c) = (C 0, C c)" |
|
906 "split_int (CN n c b) = |
|
907 (let (bv,bi) = split_int b |
|
908 in (CN n c bv, bi))" |
|
909 "split_int (CF c a b) = |
|
910 (let (bv,bi) = split_int b |
|
911 in (bv, CF c a bi))" |
|
912 "split_int a = (a,C 0)" |
|
913 |
|
914 lemma split_int:"\<And> tv ti. split_int t = (tv,ti) \<Longrightarrow> (Inum bs (Add tv ti) = Inum bs t) \<and> isint ti bs" |
|
915 proof (induct t rule: split_int.induct) |
|
916 case (2 c n b tv ti) |
|
917 let ?bv = "fst (split_int b)" |
|
918 let ?bi = "snd (split_int b)" |
|
919 have "split_int b = (?bv,?bi)" by simp |
|
920 with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+ |
|
921 from prems(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def) |
|
922 from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def) |
|
923 next |
|
924 case (3 c a b tv ti) |
|
925 let ?bv = "fst (split_int b)" |
|
926 let ?bi = "snd (split_int b)" |
|
927 have "split_int b = (?bv,?bi)" by simp |
|
928 with prems(1) have b:"Inum bs (Add ?bv ?bi) = Inum bs b" and bii: "isint ?bi bs" by blast+ |
|
929 from prems(2) have tibi: "ti = CF c a ?bi" by (simp add: Let_def split_def) |
|
930 from prems(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def isint_Floor isint_add isint_Mul isint_CF) |
|
931 qed (auto simp add: Let_def isint_iff isint_Floor isint_add isint_Mul split_def algebra_simps) |
|
932 |
|
933 lemma split_int_nb: "numbound0 t \<Longrightarrow> numbound0 (fst (split_int t)) \<and> numbound0 (snd (split_int t)) " |
|
934 by (induct t rule: split_int.induct, auto simp add: Let_def split_def) |
|
935 |
|
936 definition |
|
937 numfloor:: "num \<Rightarrow> num" |
|
938 where |
|
939 numfloor_def: "numfloor t = (let (tv,ti) = split_int t in |
|
940 (case tv of C i \<Rightarrow> numadd (tv,ti) |
|
941 | _ \<Rightarrow> numadd(CF 1 tv (C 0),ti)))" |
|
942 |
|
943 lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)") |
|
944 proof- |
|
945 let ?tv = "fst (split_int t)" |
|
946 let ?ti = "snd (split_int t)" |
|
947 have tvti:"split_int t = (?tv,?ti)" by simp |
|
948 {assume H: "\<forall> v. ?tv \<noteq> C v" |
|
949 hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" |
|
950 by (cases ?tv, auto simp add: numfloor_def Let_def split_def numadd) |
|
951 from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ |
|
952 hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp |
|
953 also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))" |
|
954 by (simp,subst tii[simplified isint_iff, symmetric]) simp |
|
955 also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) |
|
956 finally have ?thesis using th1 by simp} |
|
957 moreover {fix v assume H:"?tv = C v" |
|
958 from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ |
|
959 hence "?N (Floor t) = real (floor (?N (Add ?tv ?ti)))" by simp |
|
960 also have "\<dots> = real (floor (?N ?tv) + (floor (?N ?ti)))" |
|
961 by (simp,subst tii[simplified isint_iff, symmetric]) simp |
|
962 also have "\<dots> = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) |
|
963 finally have ?thesis by (simp add: H numfloor_def Let_def split_def numadd) } |
|
964 ultimately show ?thesis by auto |
|
965 qed |
|
966 |
|
967 lemma numfloor_nb[simp]: "numbound0 t \<Longrightarrow> numbound0 (numfloor t)" |
|
968 using split_int_nb[where t="t"] |
|
969 by (cases "fst(split_int t)" , auto simp add: numfloor_def Let_def split_def numadd_nb) |
|
970 |
|
971 recdef simpnum "measure num_size" |
|
972 "simpnum (C j) = C j" |
|
973 "simpnum (Bound n) = CN n 1 (C 0)" |
|
974 "simpnum (Neg t) = numneg (simpnum t)" |
|
975 "simpnum (Add t s) = numadd (simpnum t,simpnum s)" |
|
976 "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" |
|
977 "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)" |
|
978 "simpnum (Floor t) = numfloor (simpnum t)" |
|
979 "simpnum (CN n c t) = (if c=0 then simpnum t else CN n c (simpnum t))" |
|
980 "simpnum (CF c t s) = simpnum(Add (Mul c (Floor t)) s)" |
|
981 |
|
982 lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t" |
|
983 by (induct t rule: simpnum.induct, auto) |
|
984 |
|
985 lemma simpnum_numbound0[simp]: |
|
986 "numbound0 t \<Longrightarrow> numbound0 (simpnum t)" |
|
987 by (induct t rule: simpnum.induct, auto) |
|
988 |
|
989 consts nozerocoeff:: "num \<Rightarrow> bool" |
|
990 recdef nozerocoeff "measure size" |
|
991 "nozerocoeff (C c) = True" |
|
992 "nozerocoeff (CN n c t) = (c\<noteq>0 \<and> nozerocoeff t)" |
|
993 "nozerocoeff (CF c s t) = (c \<noteq> 0 \<and> nozerocoeff t)" |
|
994 "nozerocoeff (Mul c t) = (c\<noteq>0 \<and> nozerocoeff t)" |
|
995 "nozerocoeff t = True" |
|
996 |
|
997 lemma numadd_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numadd (a,b))" |
|
998 by (induct a b rule: numadd.induct,auto simp add: Let_def) |
|
999 |
|
1000 lemma nummul_nz : "\<And> i. i\<noteq>0 \<Longrightarrow> nozerocoeff a \<Longrightarrow> nozerocoeff (nummul a i)" |
|
1001 by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz) |
|
1002 |
|
1003 lemma numneg_nz : "nozerocoeff a \<Longrightarrow> nozerocoeff (numneg a)" |
|
1004 by (simp add: numneg_def nummul_nz) |
|
1005 |
|
1006 lemma numsub_nz: "nozerocoeff a \<Longrightarrow> nozerocoeff b \<Longrightarrow> nozerocoeff (numsub a b)" |
|
1007 by (simp add: numsub_def numneg_nz numadd_nz) |
|
1008 |
|
1009 lemma split_int_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (fst (split_int t)) \<and> nozerocoeff (snd (split_int t))" |
|
1010 by (induct t rule: split_int.induct,auto simp add: Let_def split_def) |
|
1011 |
|
1012 lemma numfloor_nz: "nozerocoeff t \<Longrightarrow> nozerocoeff (numfloor t)" |
|
1013 by (simp add: numfloor_def Let_def split_def) |
|
1014 (cases "fst (split_int t)", simp_all add: split_int_nz numadd_nz) |
|
1015 |
|
1016 lemma simpnum_nz: "nozerocoeff (simpnum t)" |
|
1017 by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz numfloor_nz) |
|
1018 |
|
1019 lemma maxcoeff_nz: "nozerocoeff t \<Longrightarrow> maxcoeff t = 0 \<Longrightarrow> t = C 0" |
|
1020 proof (induct t rule: maxcoeff.induct) |
|
1021 case (2 n c t) |
|
1022 hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ |
|
1023 have "max (abs c) (maxcoeff t) \<ge> abs c" by (simp add: le_maxI1) |
|
1024 with cnz have "max (abs c) (maxcoeff t) > 0" by arith |
|
1025 with prems show ?case by simp |
|
1026 next |
|
1027 case (3 c s t) |
|
1028 hence cnz: "c \<noteq>0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ |
|
1029 have "max (abs c) (maxcoeff t) \<ge> abs c" by (simp add: le_maxI1) |
|
1030 with cnz have "max (abs c) (maxcoeff t) > 0" by arith |
|
1031 with prems show ?case by simp |
|
1032 qed auto |
|
1033 |
|
1034 lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0" |
|
1035 proof- |
|
1036 from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def) |
|
1037 from numgcdh0[OF th] have th:"maxcoeff t = 0" . |
|
1038 from maxcoeff_nz[OF nz th] show ?thesis . |
|
1039 qed |
|
1040 |
|
1041 constdefs simp_num_pair:: "(num \<times> int) \<Rightarrow> num \<times> int" |
|
1042 "simp_num_pair \<equiv> (\<lambda> (t,n). (if n = 0 then (C 0, 0) else |
|
1043 (let t' = simpnum t ; g = numgcd t' in |
|
1044 if g > 1 then (let g' = zgcd n g in |
|
1045 if g' = 1 then (t',n) |
|
1046 else (reducecoeffh t' g', n div g')) |
|
1047 else (t',n))))" |
|
1048 |
|
1049 lemma simp_num_pair_ci: |
|
1050 shows "((\<lambda> (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((\<lambda> (t,n). Inum bs t / real n) (t,n))" |
|
1051 (is "?lhs = ?rhs") |
|
1052 proof- |
|
1053 let ?t' = "simpnum t" |
|
1054 let ?g = "numgcd ?t'" |
|
1055 let ?g' = "zgcd n ?g" |
|
1056 {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)} |
|
1057 moreover |
|
1058 { assume nnz: "n \<noteq> 0" |
|
1059 {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def)} |
|
1060 moreover |
|
1061 {assume g1:"?g>1" hence g0: "?g > 0" by simp |
|
1062 from zgcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp |
|
1063 hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith |
|
1064 hence "?g'= 1 \<or> ?g' > 1" by arith |
|
1065 moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def)} |
|
1066 moreover {assume g'1:"?g'>1" |
|
1067 from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" .. |
|
1068 let ?tt = "reducecoeffh ?t' ?g'" |
|
1069 let ?t = "Inum bs ?tt" |
|
1070 have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) |
|
1071 have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) |
|
1072 have gpdgp: "?g' dvd ?g'" by simp |
|
1073 from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] |
|
1074 have th2:"real ?g' * ?t = Inum bs ?t'" by simp |
|
1075 from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def) |
|
1076 also have "\<dots> = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp |
|
1077 also have "\<dots> = (Inum bs ?t' / real n)" |
|
1078 using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp |
|
1079 finally have "?lhs = Inum bs t / real n" by simp |
|
1080 then have ?thesis using prems by (simp add: simp_num_pair_def)} |
|
1081 ultimately have ?thesis by blast} |
|
1082 ultimately have ?thesis by blast} |
|
1083 ultimately show ?thesis by blast |
|
1084 qed |
|
1085 |
|
1086 lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')" |
|
1087 shows "numbound0 t' \<and> n' >0" |
|
1088 proof- |
|
1089 let ?t' = "simpnum t" |
|
1090 let ?g = "numgcd ?t'" |
|
1091 let ?g' = "zgcd n ?g" |
|
1092 {assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)} |
|
1093 moreover |
|
1094 { assume nnz: "n \<noteq> 0" |
|
1095 {assume "\<not> ?g > 1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def)} |
|
1096 moreover |
|
1097 {assume g1:"?g>1" hence g0: "?g > 0" by simp |
|
1098 from zgcd0 g1 nnz have gp0: "?g' \<noteq> 0" by simp |
|
1099 hence g'p: "?g' > 0" using zgcd_pos[where i="n" and j="numgcd ?t'"] by arith |
|
1100 hence "?g'= 1 \<or> ?g' > 1" by arith |
|
1101 moreover {assume "?g'=1" hence ?thesis using prems |
|
1102 by (auto simp add: Let_def simp_num_pair_def)} |
|
1103 moreover {assume g'1:"?g'>1" |
|
1104 have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) |
|
1105 have gpdd: "?g' dvd n" by (simp add: zgcd_zdvd1) |
|
1106 have gpdgp: "?g' dvd ?g'" by simp |
|
1107 from zdvd_imp_le[OF gpdd np] have g'n: "?g' \<le> n" . |
|
1108 from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]] |
|
1109 have "n div ?g' >0" by simp |
|
1110 hence ?thesis using prems |
|
1111 by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)} |
|
1112 ultimately have ?thesis by blast} |
|
1113 ultimately have ?thesis by blast} |
|
1114 ultimately show ?thesis by blast |
|
1115 qed |
|
1116 |
|
1117 consts not:: "fm \<Rightarrow> fm" |
|
1118 recdef not "measure size" |
|
1119 "not (NOT p) = p" |
|
1120 "not T = F" |
|
1121 "not F = T" |
|
1122 "not (Lt t) = Ge t" |
|
1123 "not (Le t) = Gt t" |
|
1124 "not (Gt t) = Le t" |
|
1125 "not (Ge t) = Lt t" |
|
1126 "not (Eq t) = NEq t" |
|
1127 "not (NEq t) = Eq t" |
|
1128 "not (Dvd i t) = NDvd i t" |
|
1129 "not (NDvd i t) = Dvd i t" |
|
1130 "not (And p q) = Or (not p) (not q)" |
|
1131 "not (Or p q) = And (not p) (not q)" |
|
1132 "not p = NOT p" |
|
1133 lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)" |
|
1134 by (induct p) auto |
|
1135 lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)" |
|
1136 by (induct p, auto) |
|
1137 lemma not_nb[simp]: "bound0 p \<Longrightarrow> bound0 (not p)" |
|
1138 by (induct p, auto) |
|
1139 |
|
1140 constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" |
|
1141 "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 |
|
1142 if p = q then p else And p q)" |
|
1143 lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" |
|
1144 by (cases "p=F \<or> q=F",simp_all add: conj_def) (cases p,simp_all) |
|
1145 |
|
1146 lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)" |
|
1147 using conj_def by auto |
|
1148 lemma conj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)" |
|
1149 using conj_def by auto |
|
1150 |
|
1151 constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" |
|
1152 "disj p q \<equiv> (if (p = T \<or> q=T) then T else if p=F then q else if q=F then p |
|
1153 else if p=q then p else Or p q)" |
|
1154 |
|
1155 lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)" |
|
1156 by (cases "p=T \<or> q=T",simp_all add: disj_def) (cases p,simp_all) |
|
1157 lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)" |
|
1158 using disj_def by auto |
|
1159 lemma disj_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)" |
|
1160 using disj_def by auto |
|
1161 |
|
1162 constdefs imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" |
|
1163 "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 |
|
1164 else Imp p q)" |
|
1165 lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" |
|
1166 by (cases "p=F \<or> q=T",simp_all add: imp_def) |
|
1167 lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)" |
|
1168 using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def) |
|
1169 lemma imp_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)" |
|
1170 using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def) |
|
1171 |
|
1172 constdefs iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" |
|
1173 "iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> not p = q) then F else |
|
1174 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 |
|
1175 Iff p q)" |
|
1176 lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)" |
|
1177 by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp add:not) |
|
1178 (cases "not p= q", auto simp add:not) |
|
1179 lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)" |
|
1180 by (unfold iff_def,cases "p=q", auto) |
|
1181 lemma iff_nb[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (iff p q)" |
|
1182 using iff_def by (unfold iff_def,cases "p=q", auto) |
|
1183 |
|
1184 consts check_int:: "num \<Rightarrow> bool" |
|
1185 recdef check_int "measure size" |
|
1186 "check_int (C i) = True" |
|
1187 "check_int (Floor t) = True" |
|
1188 "check_int (Mul i t) = check_int t" |
|
1189 "check_int (Add t s) = (check_int t \<and> check_int s)" |
|
1190 "check_int (Neg t) = check_int t" |
|
1191 "check_int (CF c t s) = check_int s" |
|
1192 "check_int t = False" |
|
1193 lemma check_int: "check_int t \<Longrightarrow> isint t bs" |
|
1194 by (induct t, auto simp add: isint_add isint_Floor isint_Mul isint_neg isint_c isint_CF) |
|
1195 |
|
1196 lemma rdvd_left1_int: "real \<lfloor>t\<rfloor> = t \<Longrightarrow> 1 rdvd t" |
|
1197 by (simp add: rdvd_def,rule_tac x="\<lfloor>t\<rfloor>" in exI) simp |
|
1198 |
|
1199 lemma rdvd_reduce: |
|
1200 assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0" |
|
1201 shows "real (d::int) rdvd real (c::int)*t = (real (d div g) rdvd real (c div g)*t)" |
|
1202 proof |
|
1203 assume d: "real d rdvd real c * t" |
|
1204 from d rdvd_def obtain k where k_def: "real c * t = real d* real (k::int)" by auto |
|
1205 from gd dvd_def obtain kd where kd_def: "d = g * kd" by auto |
|
1206 from gc dvd_def obtain kc where kc_def: "c = g * kc" by auto |
|
1207 from k_def kd_def kc_def have "real g * real kc * t = real g * real kd * real k" by simp |
|
1208 hence "real kc * t = real kd * real k" using gp by simp |
|
1209 hence th:"real kd rdvd real kc * t" using rdvd_def by blast |
|
1210 from kd_def gp have th':"kd = d div g" by simp |
|
1211 from kc_def gp have "kc = c div g" by simp |
|
1212 with th th' show "real (d div g) rdvd real (c div g) * t" by simp |
|
1213 next |
|
1214 assume d: "real (d div g) rdvd real (c div g) * t" |
|
1215 from gp have gnz: "g \<noteq> 0" by simp |
|
1216 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 gnz gd] real_of_int_div[OF gnz gc] by simp |
|
1217 qed |
|
1218 |
|
1219 constdefs simpdvd:: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)" |
|
1220 "simpdvd d t \<equiv> |
|
1221 (let g = numgcd t in |
|
1222 if g > 1 then (let g' = zgcd d g in |
|
1223 if g' = 1 then (d, t) |
|
1224 else (d div g',reducecoeffh t g')) |
|
1225 else (d, t))" |
|
1226 lemma simpdvd: |
|
1227 assumes tnz: "nozerocoeff t" and dnz: "d \<noteq> 0" |
|
1228 shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)" |
|
1229 proof- |
|
1230 let ?g = "numgcd t" |
|
1231 let ?g' = "zgcd d ?g" |
|
1232 {assume "\<not> ?g > 1" hence ?thesis by (simp add: Let_def simpdvd_def)} |
|
1233 moreover |
|
1234 {assume g1:"?g>1" hence g0: "?g > 0" by simp |
|
1235 from zgcd0 g1 dnz have gp0: "?g' \<noteq> 0" by simp |
|
1236 hence g'p: "?g' > 0" using zgcd_pos[where i="d" and j="numgcd t"] by arith |
|
1237 hence "?g'= 1 \<or> ?g' > 1" by arith |
|
1238 moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simpdvd_def)} |
|
1239 moreover {assume g'1:"?g'>1" |
|
1240 from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" .. |
|
1241 let ?tt = "reducecoeffh t ?g'" |
|
1242 let ?t = "Inum bs ?tt" |
|
1243 have gpdg: "?g' dvd ?g" by (simp add: zgcd_zdvd2) |
|
1244 have gpdd: "?g' dvd d" by (simp add: zgcd_zdvd1) |
|
1245 have gpdgp: "?g' dvd ?g'" by simp |
|
1246 from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] |
|
1247 have th2:"real ?g' * ?t = Inum bs t" by simp |
|
1248 from prems have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)" |
|
1249 by (simp add: simpdvd_def Let_def) |
|
1250 also have "\<dots> = (real d rdvd (Inum bs t))" |
|
1251 using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified zdiv_self[OF gp0]] |
|
1252 th2[symmetric] by simp |
|
1253 finally have ?thesis by simp } |
|
1254 ultimately have ?thesis by blast |
|
1255 } |
|
1256 ultimately show ?thesis by blast |
|
1257 qed |
|
1258 |
|
1259 consts simpfm :: "fm \<Rightarrow> fm" |
|
1260 recdef simpfm "measure fmsize" |
|
1261 "simpfm (And p q) = conj (simpfm p) (simpfm q)" |
|
1262 "simpfm (Or p q) = disj (simpfm p) (simpfm q)" |
|
1263 "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" |
|
1264 "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" |
|
1265 "simpfm (NOT p) = not (simpfm p)" |
|
1266 "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \<Rightarrow> if (v < 0) then T else F |
|
1267 | _ \<Rightarrow> Lt (reducecoeff a'))" |
|
1268 "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'))" |
|
1269 "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'))" |
|
1270 "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'))" |
|
1271 "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'))" |
|
1272 "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'))" |
|
1273 "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a) |
|
1274 else if (abs i = 1) \<and> check_int a then T |
|
1275 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))" |
|
1276 "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) |
|
1277 else if (abs i = 1) \<and> check_int a then F |
|
1278 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))" |
|
1279 "simpfm p = p" |
|
1280 |
|
1281 lemma simpfm[simp]: "Ifm bs (simpfm p) = Ifm bs p" |
|
1282 proof(induct p rule: simpfm.induct) |
|
1283 case (6 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp |
|
1284 {fix v assume "?sa = C v" hence ?case using sa by simp } |
|
1285 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" |
|
1286 let ?g = "numgcd ?sa" |
|
1287 let ?rsa = "reducecoeff ?sa" |
|
1288 let ?r = "Inum bs ?rsa" |
|
1289 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) |
|
1290 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} |
|
1291 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) |
|
1292 hence gp: "real ?g > 0" by simp |
|
1293 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) |
|
1294 with sa have "Inum bs a < 0 = (real ?g * ?r < real ?g * 0)" by simp |
|
1295 also have "\<dots> = (?r < 0)" using gp |
|
1296 by (simp only: mult_less_cancel_left) simp |
|
1297 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} |
|
1298 ultimately show ?case by blast |
|
1299 next |
|
1300 case (7 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp |
|
1301 {fix v assume "?sa = C v" hence ?case using sa by simp } |
|
1302 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" |
|
1303 let ?g = "numgcd ?sa" |
|
1304 let ?rsa = "reducecoeff ?sa" |
|
1305 let ?r = "Inum bs ?rsa" |
|
1306 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) |
|
1307 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} |
|
1308 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) |
|
1309 hence gp: "real ?g > 0" by simp |
|
1310 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) |
|
1311 with sa have "Inum bs a \<le> 0 = (real ?g * ?r \<le> real ?g * 0)" by simp |
|
1312 also have "\<dots> = (?r \<le> 0)" using gp |
|
1313 by (simp only: mult_le_cancel_left) simp |
|
1314 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} |
|
1315 ultimately show ?case by blast |
|
1316 next |
|
1317 case (8 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp |
|
1318 {fix v assume "?sa = C v" hence ?case using sa by simp } |
|
1319 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" |
|
1320 let ?g = "numgcd ?sa" |
|
1321 let ?rsa = "reducecoeff ?sa" |
|
1322 let ?r = "Inum bs ?rsa" |
|
1323 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) |
|
1324 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} |
|
1325 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) |
|
1326 hence gp: "real ?g > 0" by simp |
|
1327 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) |
|
1328 with sa have "Inum bs a > 0 = (real ?g * ?r > real ?g * 0)" by simp |
|
1329 also have "\<dots> = (?r > 0)" using gp |
|
1330 by (simp only: mult_less_cancel_left) simp |
|
1331 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} |
|
1332 ultimately show ?case by blast |
|
1333 next |
|
1334 case (9 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp |
|
1335 {fix v assume "?sa = C v" hence ?case using sa by simp } |
|
1336 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" |
|
1337 let ?g = "numgcd ?sa" |
|
1338 let ?rsa = "reducecoeff ?sa" |
|
1339 let ?r = "Inum bs ?rsa" |
|
1340 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) |
|
1341 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} |
|
1342 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) |
|
1343 hence gp: "real ?g > 0" by simp |
|
1344 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) |
|
1345 with sa have "Inum bs a \<ge> 0 = (real ?g * ?r \<ge> real ?g * 0)" by simp |
|
1346 also have "\<dots> = (?r \<ge> 0)" using gp |
|
1347 by (simp only: mult_le_cancel_left) simp |
|
1348 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} |
|
1349 ultimately show ?case by blast |
|
1350 next |
|
1351 case (10 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp |
|
1352 {fix v assume "?sa = C v" hence ?case using sa by simp } |
|
1353 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" |
|
1354 let ?g = "numgcd ?sa" |
|
1355 let ?rsa = "reducecoeff ?sa" |
|
1356 let ?r = "Inum bs ?rsa" |
|
1357 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) |
|
1358 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} |
|
1359 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) |
|
1360 hence gp: "real ?g > 0" by simp |
|
1361 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) |
|
1362 with sa have "Inum bs a = 0 = (real ?g * ?r = 0)" by simp |
|
1363 also have "\<dots> = (?r = 0)" using gp |
|
1364 by (simp add: mult_eq_0_iff) |
|
1365 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} |
|
1366 ultimately show ?case by blast |
|
1367 next |
|
1368 case (11 a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp |
|
1369 {fix v assume "?sa = C v" hence ?case using sa by simp } |
|
1370 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" |
|
1371 let ?g = "numgcd ?sa" |
|
1372 let ?rsa = "reducecoeff ?sa" |
|
1373 let ?r = "Inum bs ?rsa" |
|
1374 have sa_nz: "nozerocoeff ?sa" by (rule simpnum_nz) |
|
1375 {assume gz: "?g=0" from numgcd_nz[OF sa_nz gz] H have "False" by auto} |
|
1376 with numgcd_pos[where t="?sa"] have "?g > 0" by (cases "?g=0", auto) |
|
1377 hence gp: "real ?g > 0" by simp |
|
1378 have "Inum bs ?sa = real ?g* ?r" by (simp add: reducecoeff) |
|
1379 with sa have "Inum bs a \<noteq> 0 = (real ?g * ?r \<noteq> 0)" by simp |
|
1380 also have "\<dots> = (?r \<noteq> 0)" using gp |
|
1381 by (simp add: mult_eq_0_iff) |
|
1382 finally have ?case using H by (cases "?sa" , simp_all add: Let_def)} |
|
1383 ultimately show ?case by blast |
|
1384 next |
|
1385 case (12 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp |
|
1386 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 |
|
1387 {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)} |
|
1388 moreover |
|
1389 {assume ai1: "abs i = 1" and ai: "check_int a" |
|
1390 hence "i=1 \<or> i= - 1" by arith |
|
1391 moreover {assume i1: "i = 1" |
|
1392 from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] |
|
1393 have ?case using i1 ai by simp } |
|
1394 moreover {assume i1: "i = - 1" |
|
1395 from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] |
|
1396 rdvd_abs1[where d="- 1" and t="Inum bs a"] |
|
1397 have ?case using i1 ai by simp } |
|
1398 ultimately have ?case by blast} |
|
1399 moreover |
|
1400 {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)" |
|
1401 {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond |
|
1402 by (cases "abs i = 1", auto simp add: int_rdvd_iff) } |
|
1403 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" |
|
1404 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) |
|
1405 from simpnum_nz have nz:"nozerocoeff ?sa" by simp |
|
1406 from simpdvd [OF nz inz] th have ?case using sa by simp} |
|
1407 ultimately have ?case by blast} |
|
1408 ultimately show ?case by blast |
|
1409 next |
|
1410 case (13 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp |
|
1411 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 |
|
1412 {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)} |
|
1413 moreover |
|
1414 {assume ai1: "abs i = 1" and ai: "check_int a" |
|
1415 hence "i=1 \<or> i= - 1" by arith |
|
1416 moreover {assume i1: "i = 1" |
|
1417 from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] |
|
1418 have ?case using i1 ai by simp } |
|
1419 moreover {assume i1: "i = - 1" |
|
1420 from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] |
|
1421 rdvd_abs1[where d="- 1" and t="Inum bs a"] |
|
1422 have ?case using i1 ai by simp } |
|
1423 ultimately have ?case by blast} |
|
1424 moreover |
|
1425 {assume inz: "i\<noteq>0" and cond: "(abs i \<noteq> 1) \<or> (\<not> check_int a)" |
|
1426 {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond |
|
1427 by (cases "abs i = 1", auto simp add: int_rdvd_iff) } |
|
1428 moreover {assume H:"\<not> (\<exists> v. ?sa = C v)" |
|
1429 hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond |
|
1430 by (cases ?sa, auto simp add: Let_def split_def) |
|
1431 from simpnum_nz have nz:"nozerocoeff ?sa" by simp |
|
1432 from simpdvd [OF nz inz] th have ?case using sa by simp} |
|
1433 ultimately have ?case by blast} |
|
1434 ultimately show ?case by blast |
|
1435 qed (induct p rule: simpfm.induct, simp_all) |
|
1436 |
|
1437 lemma simpdvd_numbound0: "numbound0 t \<Longrightarrow> numbound0 (snd (simpdvd d t))" |
|
1438 by (simp add: simpdvd_def Let_def split_def reducecoeffh_numbound0) |
|
1439 |
|
1440 lemma simpfm_bound0[simp]: "bound0 p \<Longrightarrow> bound0 (simpfm p)" |
|
1441 proof(induct p rule: simpfm.induct) |
|
1442 case (6 a) hence nb: "numbound0 a" by simp |
|
1443 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) |
|
1444 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) |
|
1445 next |
|
1446 case (7 a) hence nb: "numbound0 a" by simp |
|
1447 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) |
|
1448 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) |
|
1449 next |
|
1450 case (8 a) hence nb: "numbound0 a" by simp |
|
1451 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) |
|
1452 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) |
|
1453 next |
|
1454 case (9 a) hence nb: "numbound0 a" by simp |
|
1455 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) |
|
1456 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) |
|
1457 next |
|
1458 case (10 a) hence nb: "numbound0 a" by simp |
|
1459 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) |
|
1460 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) |
|
1461 next |
|
1462 case (11 a) hence nb: "numbound0 a" by simp |
|
1463 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) |
|
1464 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0) |
|
1465 next |
|
1466 case (12 i a) hence nb: "numbound0 a" by simp |
|
1467 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) |
|
1468 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def) |
|
1469 next |
|
1470 case (13 i a) hence nb: "numbound0 a" by simp |
|
1471 hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) |
|
1472 thus ?case by (cases "simpnum a", auto simp add: Let_def reducecoeff_numbound0 simpdvd_numbound0 split_def) |
|
1473 qed(auto simp add: disj_def imp_def iff_def conj_def) |
|
1474 |
|
1475 lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)" |
|
1476 by (induct p rule: simpfm.induct, auto simp add: Let_def) |
|
1477 (case_tac "simpnum a",auto simp add: split_def Let_def)+ |
|
1478 |
|
1479 |
|
1480 (* Generic quantifier elimination *) |
|
1481 |
|
1482 constdefs list_conj :: "fm list \<Rightarrow> fm" |
|
1483 "list_conj ps \<equiv> foldr conj ps T" |
|
1484 lemma list_conj: "Ifm bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm bs p)" |
|
1485 by (induct ps, auto simp add: list_conj_def) |
|
1486 lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)" |
|
1487 by (induct ps, auto simp add: list_conj_def) |
|
1488 lemma list_conj_nb: " \<forall>p\<in> set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)" |
|
1489 by (induct ps, auto simp add: list_conj_def) |
|
1490 constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" |
|
1491 "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = List.partition bound0 cjs |
|
1492 in conj (decr (list_conj yes)) (f (list_conj no)))" |
|
1493 |
|
1494 lemma CJNB_qe: |
|
1495 assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))" |
|
1496 shows "\<forall> bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (Ifm bs ((CJNB qe p)) = Ifm bs (E p))" |
|
1497 proof(clarify) |
|
1498 fix bs p |
|
1499 assume qfp: "qfree p" |
|
1500 let ?cjs = "conjuncts p" |
|
1501 let ?yes = "fst (List.partition bound0 ?cjs)" |
|
1502 let ?no = "snd (List.partition bound0 ?cjs)" |
|
1503 let ?cno = "list_conj ?no" |
|
1504 let ?cyes = "list_conj ?yes" |
|
1505 have part: "List.partition bound0 ?cjs = (?yes,?no)" by simp |
|
1506 from partition_P[OF part] have "\<forall> q\<in> set ?yes. bound0 q" by blast |
|
1507 hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) |
|
1508 hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf) |
|
1509 from conjuncts_qf[OF qfp] partition_set[OF part] |
|
1510 have " \<forall>q\<in> set ?no. qfree q" by auto |
|
1511 hence no_qf: "qfree ?cno"by (simp add: list_conj_qf) |
|
1512 with qe have cno_qf:"qfree (qe ?cno )" |
|
1513 and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+ |
|
1514 from cno_qf yes_qf have qf: "qfree (CJNB qe p)" |
|
1515 by (simp add: CJNB_def Let_def conj_qf split_def) |
|
1516 {fix bs |
|
1517 from conjuncts have "Ifm bs p = (\<forall>q\<in> set ?cjs. Ifm bs q)" by blast |
|
1518 also have "\<dots> = ((\<forall>q\<in> set ?yes. Ifm bs q) \<and> (\<forall>q\<in> set ?no. Ifm bs q))" |
|
1519 using partition_set[OF part] by auto |
|
1520 finally have "Ifm bs p = ((Ifm bs ?cyes) \<and> (Ifm bs ?cno))" using list_conj by simp} |
|
1521 hence "Ifm bs (E p) = (\<exists>x. (Ifm (x#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))" by simp |
|
1522 also fix y have "\<dots> = (\<exists>x. (Ifm (y#bs) ?cyes) \<and> (Ifm (x#bs) ?cno))" |
|
1523 using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast |
|
1524 also have "\<dots> = (Ifm bs (decr ?cyes) \<and> Ifm bs (E ?cno))" |
|
1525 by (auto simp add: decr[OF yes_nb]) |
|
1526 also have "\<dots> = (Ifm bs (conj (decr ?cyes) (qe ?cno)))" |
|
1527 using qe[rule_format, OF no_qf] by auto |
|
1528 finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" |
|
1529 by (simp add: Let_def CJNB_def split_def) |
|
1530 with qf show "qfree (CJNB qe p) \<and> Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast |
|
1531 qed |
|
1532 |
|
1533 consts qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm" |
|
1534 recdef qelim "measure fmsize" |
|
1535 "qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))" |
|
1536 "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))" |
|
1537 "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))" |
|
1538 "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))" |
|
1539 "qelim (Or p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))" |
|
1540 "qelim (Imp p q) = (\<lambda> qe. disj (qelim (NOT p) qe) (qelim q qe))" |
|
1541 "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))" |
|
1542 "qelim p = (\<lambda> y. simpfm p)" |
|
1543 |
|
1544 lemma qelim_ci: |
|
1545 assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm bs (qe p) = Ifm bs (E p))" |
|
1546 shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm bs (qelim p qe) = Ifm bs p)" |
|
1547 using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] |
|
1548 by(induct p rule: qelim.induct) |
|
1549 (auto simp del: simpfm.simps) |
|
1550 |
|
1551 |
|
1552 text {* The @{text "\<int>"} Part *} |
|
1553 text{* Linearity for fm where Bound 0 ranges over @{text "\<int>"} *} |
|
1554 consts |
|
1555 zsplit0 :: "num \<Rightarrow> int \<times> num" (* splits the bounded from the unbounded part*) |
|
1556 recdef zsplit0 "measure num_size" |
|
1557 "zsplit0 (C c) = (0,C c)" |
|
1558 "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))" |
|
1559 "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)" |
|
1560 "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)" |
|
1561 "zsplit0 (Neg a) = (let (i',a') = zsplit0 a in (-i', Neg a'))" |
|
1562 "zsplit0 (Add a b) = (let (ia,a') = zsplit0 a ; |
|
1563 (ib,b') = zsplit0 b |
|
1564 in (ia+ib, Add a' b'))" |
|
1565 "zsplit0 (Sub a b) = (let (ia,a') = zsplit0 a ; |
|
1566 (ib,b') = zsplit0 b |
|
1567 in (ia-ib, Sub a' b'))" |
|
1568 "zsplit0 (Mul i a) = (let (i',a') = zsplit0 a in (i*i', Mul i a'))" |
|
1569 "zsplit0 (Floor a) = (let (i',a') = zsplit0 a in (i',Floor a'))" |
|
1570 (hints simp add: Let_def) |
|
1571 |
|
1572 lemma zsplit0_I: |
|
1573 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" |
|
1574 (is "\<And> n a. ?S t = (n,a) \<Longrightarrow> (?I x (CN 0 n a) = ?I x t) \<and> ?N a") |
|
1575 proof(induct t rule: zsplit0.induct) |
|
1576 case (1 c n a) thus ?case by auto |
|
1577 next |
|
1578 case (2 m n a) thus ?case by (cases "m=0") auto |
|
1579 next |
|
1580 case (3 n i a n a') thus ?case by auto |
|
1581 next |
|
1582 case (4 c a b n a') thus ?case by auto |
|
1583 next |
|
1584 case (5 t n a) |
|
1585 let ?nt = "fst (zsplit0 t)" |
|
1586 let ?at = "snd (zsplit0 t)" |
|
1587 have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \<and> n=-?nt" using prems |
|
1588 by (simp add: Let_def split_def) |
|
1589 from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast |
|
1590 from th2[simplified] th[simplified] show ?case by simp |
|
1591 next |
|
1592 case (6 s t n a) |
|
1593 let ?ns = "fst (zsplit0 s)" |
|
1594 let ?as = "snd (zsplit0 s)" |
|
1595 let ?nt = "fst (zsplit0 t)" |
|
1596 let ?at = "snd (zsplit0 t)" |
|
1597 have abjs: "zsplit0 s = (?ns,?as)" by simp |
|
1598 moreover have abjt: "zsplit0 t = (?nt,?at)" by simp |
|
1599 ultimately have th: "a=Add ?as ?at \<and> n=?ns + ?nt" using prems |
|
1600 by (simp add: Let_def split_def) |
|
1601 from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast |
|
1602 from prems 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 simp |
|
1603 with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast |
|
1604 from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast |
|
1605 from th3[simplified] th2[simplified] th[simplified] show ?case |
|
1606 by (simp add: left_distrib) |
|
1607 next |
|
1608 case (7 s t n a) |
|
1609 let ?ns = "fst (zsplit0 s)" |
|
1610 let ?as = "snd (zsplit0 s)" |
|
1611 let ?nt = "fst (zsplit0 t)" |
|
1612 let ?at = "snd (zsplit0 t)" |
|
1613 have abjs: "zsplit0 s = (?ns,?as)" by simp |
|
1614 moreover have abjt: "zsplit0 t = (?nt,?at)" by simp |
|
1615 ultimately have th: "a=Sub ?as ?at \<and> n=?ns - ?nt" using prems |
|
1616 by (simp add: Let_def split_def) |
|
1617 from abjs[symmetric] have bluddy: "\<exists> x y. (x,y) = zsplit0 s" by blast |
|
1618 from prems 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 simp |
|
1619 with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast |
|
1620 from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \<and> ?N ?as" by blast |
|
1621 from th3[simplified] th2[simplified] th[simplified] show ?case |
|
1622 by (simp add: left_diff_distrib) |
|
1623 next |
|
1624 case (8 i t n a) |
|
1625 let ?nt = "fst (zsplit0 t)" |
|
1626 let ?at = "snd (zsplit0 t)" |
|
1627 have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at \<and> n=i*?nt" using prems |
|
1628 by (simp add: Let_def split_def) |
|
1629 from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast |
|
1630 hence " ?I x (Mul i t) = (real i) * ?I x (CN 0 ?nt ?at)" by simp |
|
1631 also have "\<dots> = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: right_distrib) |
|
1632 finally show ?case using th th2 by simp |
|
1633 next |
|
1634 case (9 t n a) |
|
1635 let ?nt = "fst (zsplit0 t)" |
|
1636 let ?at = "snd (zsplit0 t)" |
|
1637 have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a= Floor ?at \<and> n=?nt" using prems |
|
1638 by (simp add: Let_def split_def) |
|
1639 from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \<and> ?N ?at" by blast |
|
1640 hence na: "?N a" using th by simp |
|
1641 have th': "(real ?nt)*(real x) = real (?nt * x)" by simp |
|
1642 have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp |
|
1643 also have "\<dots> = real (floor ((real ?nt)* real(x) + ?I x ?at))" by simp |
|
1644 also have "\<dots> = real (floor (?I x ?at + real (?nt* x)))" by (simp add: add_ac) |
|
1645 also have "\<dots> = real (floor (?I x ?at) + (?nt* x))" |
|
1646 using floor_add[where x="?I x ?at" and a="?nt* x"] by simp |
|
1647 also have "\<dots> = real (?nt)*(real x) + real (floor (?I x ?at))" by (simp add: add_ac) |
|
1648 finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp |
|
1649 with na show ?case by simp |
|
1650 qed |
|
1651 |
|
1652 consts |
|
1653 iszlfm :: "fm \<Rightarrow> real list \<Rightarrow> bool" (* Linearity test for fm *) |
|
1654 zlfm :: "fm \<Rightarrow> fm" (* Linearity transformation for fm *) |
|
1655 recdef iszlfm "measure size" |
|
1656 "iszlfm (And p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)" |
|
1657 "iszlfm (Or p q) = (\<lambda> bs. iszlfm p bs \<and> iszlfm q bs)" |
|
1658 "iszlfm (Eq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)" |
|
1659 "iszlfm (NEq (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)" |
|
1660 "iszlfm (Lt (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)" |
|
1661 "iszlfm (Le (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)" |
|
1662 "iszlfm (Gt (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)" |
|
1663 "iszlfm (Ge (CN 0 c e)) = (\<lambda> bs. c>0 \<and> numbound0 e \<and> isint e bs)" |
|
1664 "iszlfm (Dvd i (CN 0 c e)) = |
|
1665 (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)" |
|
1666 "iszlfm (NDvd i (CN 0 c e))= |
|
1667 (\<lambda> bs. c>0 \<and> i>0 \<and> numbound0 e \<and> isint e bs)" |
|
1668 "iszlfm p = (\<lambda> bs. isatom p \<and> (bound0 p))" |
|
1669 |
|
1670 lemma zlin_qfree: "iszlfm p bs \<Longrightarrow> qfree p" |
|
1671 by (induct p rule: iszlfm.induct) auto |
|
1672 |
|
1673 lemma iszlfm_gen: |
|
1674 assumes lp: "iszlfm p (x#bs)" |
|
1675 shows "\<forall> y. iszlfm p (y#bs)" |
|
1676 proof |
|
1677 fix y |
|
1678 show "iszlfm p (y#bs)" |
|
1679 using lp |
|
1680 by(induct p rule: iszlfm.induct, simp_all add: numbound0_gen[rule_format, where x="x" and y="y"]) |
|
1681 qed |
|
1682 |
|
1683 lemma conj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (conj p q) bs" |
|
1684 using conj_def by (cases p,auto) |
|
1685 lemma disj_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm q bs \<Longrightarrow> iszlfm (disj p q) bs" |
|
1686 using disj_def by (cases p,auto) |
|
1687 lemma not_zl[simp]: "iszlfm p bs \<Longrightarrow> iszlfm (not p) bs" |
|
1688 by (induct p rule:iszlfm.induct ,auto) |
|
1689 |
|
1690 recdef zlfm "measure fmsize" |
|
1691 "zlfm (And p q) = conj (zlfm p) (zlfm q)" |
|
1692 "zlfm (Or p q) = disj (zlfm p) (zlfm q)" |
|
1693 "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)" |
|
1694 "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))" |
|
1695 "zlfm (Lt a) = (let (c,r) = zsplit0 a in |
|
1696 if c=0 then Lt r else |
|
1697 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))) |
|
1698 else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" |
|
1699 "zlfm (Le a) = (let (c,r) = zsplit0 a in |
|
1700 if c=0 then Le r else |
|
1701 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))) |
|
1702 else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" |
|
1703 "zlfm (Gt a) = (let (c,r) = zsplit0 a in |
|
1704 if c=0 then Gt r else |
|
1705 if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) |
|
1706 else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" |
|
1707 "zlfm (Ge a) = (let (c,r) = zsplit0 a in |
|
1708 if c=0 then Ge r else |
|
1709 if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) |
|
1710 else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" |
|
1711 "zlfm (Eq a) = (let (c,r) = zsplit0 a in |
|
1712 if c=0 then Eq r else |
|
1713 if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r))) |
|
1714 else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))" |
|
1715 "zlfm (NEq a) = (let (c,r) = zsplit0 a in |
|
1716 if c=0 then NEq r else |
|
1717 if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r))) |
|
1718 else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))" |
|
1719 "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) |
|
1720 else (let (c,r) = zsplit0 a in |
|
1721 if c=0 then Dvd (abs i) r else |
|
1722 if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r))) |
|
1723 else And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 (-c) (Neg (Floor r))))))" |
|
1724 "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) |
|
1725 else (let (c,r) = zsplit0 a in |
|
1726 if c=0 then NDvd (abs i) r else |
|
1727 if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r))) |
|
1728 else Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 (-c) (Neg (Floor r))))))" |
|
1729 "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))" |
|
1730 "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))" |
|
1731 "zlfm (NOT (Imp p q)) = conj (zlfm p) (zlfm (NOT q))" |
|
1732 "zlfm (NOT (Iff p q)) = disj (conj(zlfm p) (zlfm(NOT q))) (conj (zlfm(NOT p)) (zlfm q))" |
|
1733 "zlfm (NOT (NOT p)) = zlfm p" |
|
1734 "zlfm (NOT T) = F" |
|
1735 "zlfm (NOT F) = T" |
|
1736 "zlfm (NOT (Lt a)) = zlfm (Ge a)" |
|
1737 "zlfm (NOT (Le a)) = zlfm (Gt a)" |
|
1738 "zlfm (NOT (Gt a)) = zlfm (Le a)" |
|
1739 "zlfm (NOT (Ge a)) = zlfm (Lt a)" |
|
1740 "zlfm (NOT (Eq a)) = zlfm (NEq a)" |
|
1741 "zlfm (NOT (NEq a)) = zlfm (Eq a)" |
|
1742 "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)" |
|
1743 "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)" |
|
1744 "zlfm p = p" (hints simp add: fmsize_pos) |
|
1745 |
|
1746 lemma split_int_less_real: |
|
1747 "(real (a::int) < b) = (a < floor b \<or> (a = floor b \<and> real (floor b) < b))" |
|
1748 proof( auto) |
|
1749 assume alb: "real a < b" and agb: "\<not> a < floor b" |
|
1750 from agb have "floor b \<le> a" by simp hence th: "b < real a + 1" by (simp only: floor_le_eq) |
|
1751 from floor_eq[OF alb th] show "a= floor b" by simp |
|
1752 next |
|
1753 assume alb: "a < floor b" |
|
1754 hence "real a < real (floor b)" by simp |
|
1755 moreover have "real (floor b) \<le> b" by simp ultimately show "real a < b" by arith |
|
1756 qed |
|
1757 |
|
1758 lemma split_int_less_real': |
|
1759 "(real (a::int) + b < 0) = (real a - real (floor(-b)) < 0 \<or> (real a - real (floor (-b)) = 0 \<and> real (floor (-b)) + b < 0))" |
|
1760 proof- |
|
1761 have "(real a + b <0) = (real a < -b)" by arith |
|
1762 with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith |
|
1763 qed |
|
1764 |
|
1765 lemma split_int_gt_real': |
|
1766 "(real (a::int) + b > 0) = (real a + real (floor b) > 0 \<or> (real a + real (floor b) = 0 \<and> real (floor b) - b < 0))" |
|
1767 proof- |
|
1768 have th: "(real a + b >0) = (real (-a) + (-b)< 0)" by arith |
|
1769 show ?thesis using myless[rule_format, where b="real (floor b)"] |
|
1770 by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) |
|
1771 (simp add: algebra_simps diff_def[symmetric],arith) |
|
1772 qed |
|
1773 |
|
1774 lemma split_int_le_real: |
|
1775 "(real (a::int) \<le> b) = (a \<le> floor b \<or> (a = floor b \<and> real (floor b) < b))" |
|
1776 proof( auto) |
|
1777 assume alb: "real a \<le> b" and agb: "\<not> a \<le> floor b" |
|
1778 from alb have "floor (real a) \<le> floor b " by (simp only: floor_mono2) |
|
1779 hence "a \<le> floor b" by simp with agb show "False" by simp |
|
1780 next |
|
1781 assume alb: "a \<le> floor b" |
|
1782 hence "real a \<le> real (floor b)" by (simp only: floor_mono2) |
|
1783 also have "\<dots>\<le> b" by simp finally show "real a \<le> b" . |
|
1784 qed |
|
1785 |
|
1786 lemma split_int_le_real': |
|
1787 "(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))" |
|
1788 proof- |
|
1789 have "(real a + b \<le>0) = (real a \<le> -b)" by arith |
|
1790 with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith |
|
1791 qed |
|
1792 |
|
1793 lemma split_int_ge_real': |
|
1794 "(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))" |
|
1795 proof- |
|
1796 have th: "(real a + b \<ge>0) = (real (-a) + (-b) \<le> 0)" by arith |
|
1797 show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"]) |
|
1798 (simp add: algebra_simps diff_def[symmetric],arith) |
|
1799 qed |
|
1800 |
|
1801 lemma split_int_eq_real: "(real (a::int) = b) = ( a = floor b \<and> b = real (floor b))" (is "?l = ?r") |
|
1802 by auto |
|
1803 |
|
1804 lemma split_int_eq_real': "(real (a::int) + b = 0) = ( a - floor (-b) = 0 \<and> real (floor (-b)) + b = 0)" (is "?l = ?r") |
|
1805 proof- |
|
1806 have "?l = (real a = -b)" by arith |
|
1807 with split_int_eq_real[where a="a" and b="-b"] show ?thesis by simp arith |
|
1808 qed |
|
1809 |
|
1810 lemma zlfm_I: |
|
1811 assumes qfp: "qfree p" |
|
1812 shows "(Ifm (real i #bs) (zlfm p) = Ifm (real i# bs) p) \<and> iszlfm (zlfm p) (real (i::int) #bs)" |
|
1813 (is "(?I (?l p) = ?I p) \<and> ?L (?l p)") |
|
1814 using qfp |
|
1815 proof(induct p rule: zlfm.induct) |
|
1816 case (5 a) |
|
1817 let ?c = "fst (zsplit0 a)" |
|
1818 let ?r = "snd (zsplit0 a)" |
|
1819 have spl: "zsplit0 a = (?c,?r)" by simp |
|
1820 from zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1821 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto |
|
1822 let ?N = "\<lambda> t. Inum (real i#bs) t" |
|
1823 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith |
|
1824 moreover |
|
1825 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1826 by (cases "?r", simp_all add: Let_def split_def,case_tac "nat", simp_all)} |
|
1827 moreover |
|
1828 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" |
|
1829 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1830 have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) |
|
1831 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_def) |
|
1832 finally have ?case using l by simp} |
|
1833 moreover |
|
1834 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Lt a))" |
|
1835 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1836 have "?I (Lt a) = (real (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) |
|
1837 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_def[symmetric] add_ac, arith) |
|
1838 finally have ?case using l by simp} |
|
1839 ultimately show ?case by blast |
|
1840 next |
|
1841 case (6 a) |
|
1842 let ?c = "fst (zsplit0 a)" |
|
1843 let ?r = "snd (zsplit0 a)" |
|
1844 have spl: "zsplit0 a = (?c,?r)" by simp |
|
1845 from zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1846 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto |
|
1847 let ?N = "\<lambda> t. Inum (real i#bs) t" |
|
1848 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith |
|
1849 moreover |
|
1850 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1851 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat",simp_all)} |
|
1852 moreover |
|
1853 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" |
|
1854 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1855 have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def) |
|
1856 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_def) |
|
1857 finally have ?case using l by simp} |
|
1858 moreover |
|
1859 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Le a))" |
|
1860 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1861 have "?I (Le a) = (real (?c * i) + (?N ?r) \<le> 0)" using Ia by (simp add: Let_def split_def) |
|
1862 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_def[symmetric] add_ac ,arith) |
|
1863 finally have ?case using l by simp} |
|
1864 ultimately show ?case by blast |
|
1865 next |
|
1866 case (7 a) |
|
1867 let ?c = "fst (zsplit0 a)" |
|
1868 let ?r = "snd (zsplit0 a)" |
|
1869 have spl: "zsplit0 a = (?c,?r)" by simp |
|
1870 from zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1871 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto |
|
1872 let ?N = "\<lambda> t. Inum (real i#bs) t" |
|
1873 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith |
|
1874 moreover |
|
1875 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1876 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} |
|
1877 moreover |
|
1878 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" |
|
1879 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1880 have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) |
|
1881 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_def) |
|
1882 finally have ?case using l by simp} |
|
1883 moreover |
|
1884 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Gt a))" |
|
1885 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1886 have "?I (Gt a) = (real (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) |
|
1887 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_def[symmetric] add_ac, arith) |
|
1888 finally have ?case using l by simp} |
|
1889 ultimately show ?case by blast |
|
1890 next |
|
1891 case (8 a) |
|
1892 let ?c = "fst (zsplit0 a)" |
|
1893 let ?r = "snd (zsplit0 a)" |
|
1894 have spl: "zsplit0 a = (?c,?r)" by simp |
|
1895 from zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1896 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto |
|
1897 let ?N = "\<lambda> t. Inum (real i#bs) t" |
|
1898 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith |
|
1899 moreover |
|
1900 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1901 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} |
|
1902 moreover |
|
1903 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" |
|
1904 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1905 have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def) |
|
1906 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_def) |
|
1907 finally have ?case using l by simp} |
|
1908 moreover |
|
1909 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Ge a))" |
|
1910 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1911 have "?I (Ge a) = (real (?c * i) + (?N ?r) \<ge> 0)" using Ia by (simp add: Let_def split_def) |
|
1912 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_def[symmetric] add_ac, arith) |
|
1913 finally have ?case using l by simp} |
|
1914 ultimately show ?case by blast |
|
1915 next |
|
1916 case (9 a) |
|
1917 let ?c = "fst (zsplit0 a)" |
|
1918 let ?r = "snd (zsplit0 a)" |
|
1919 have spl: "zsplit0 a = (?c,?r)" by simp |
|
1920 from zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1921 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto |
|
1922 let ?N = "\<lambda> t. Inum (real i#bs) t" |
|
1923 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith |
|
1924 moreover |
|
1925 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1926 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} |
|
1927 moreover |
|
1928 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))" |
|
1929 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1930 have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) |
|
1931 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) |
|
1932 finally have ?case using l by simp} |
|
1933 moreover |
|
1934 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (Eq a))" |
|
1935 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1936 have "?I (Eq a) = (real (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) |
|
1937 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) |
|
1938 finally have ?case using l by simp} |
|
1939 ultimately show ?case by blast |
|
1940 next |
|
1941 case (10 a) |
|
1942 let ?c = "fst (zsplit0 a)" |
|
1943 let ?r = "snd (zsplit0 a)" |
|
1944 have spl: "zsplit0 a = (?c,?r)" by simp |
|
1945 from zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1946 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto |
|
1947 let ?N = "\<lambda> t. Inum (real i#bs) t" |
|
1948 have "?c = 0 \<or> (?c >0 \<and> ?c\<noteq>0) \<or> (?c<0 \<and> ?c\<noteq>0)" by arith |
|
1949 moreover |
|
1950 {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1951 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} |
|
1952 moreover |
|
1953 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))" |
|
1954 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1955 have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def) |
|
1956 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) |
|
1957 finally have ?case using l by simp} |
|
1958 moreover |
|
1959 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" hence l: "?L (?l (NEq a))" |
|
1960 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1961 have "?I (NEq a) = (real (?c * i) + (?N ?r) \<noteq> 0)" using Ia by (simp add: Let_def split_def) |
|
1962 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) |
|
1963 finally have ?case using l by simp} |
|
1964 ultimately show ?case by blast |
|
1965 next |
|
1966 case (11 j a) |
|
1967 let ?c = "fst (zsplit0 a)" |
|
1968 let ?r = "snd (zsplit0 a)" |
|
1969 have spl: "zsplit0 a = (?c,?r)" by simp |
|
1970 from zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
1971 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto |
|
1972 let ?N = "\<lambda> t. Inum (real i#bs) t" |
|
1973 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 |
|
1974 moreover |
|
1975 {assume "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) |
|
1976 hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)} |
|
1977 moreover |
|
1978 {assume "?c=0" and "j\<noteq>0" hence ?case |
|
1979 using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] |
|
1980 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} |
|
1981 moreover |
|
1982 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" |
|
1983 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1984 have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" |
|
1985 using Ia by (simp add: Let_def split_def) |
|
1986 also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))" |
|
1987 by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp |
|
1988 also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> |
|
1989 (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" |
|
1990 by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) |
|
1991 also have "\<dots> = (?I (?l (Dvd j a)))" using cp cnz jnz |
|
1992 by (simp add: Let_def split_def int_rdvd_iff[symmetric] |
|
1993 del: real_of_int_mult) (auto simp add: add_ac) |
|
1994 finally have ?case using l jnz by simp } |
|
1995 moreover |
|
1996 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (Dvd j a))" |
|
1997 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
1998 have "?I (Dvd j a) = (real j rdvd (real (?c * i) + (?N ?r)))" |
|
1999 using Ia by (simp add: Let_def split_def) |
|
2000 also have "\<dots> = (real (abs j) rdvd real (?c*i) + (?N ?r))" |
|
2001 by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp |
|
2002 also have "\<dots> = ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> |
|
2003 (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r))))" |
|
2004 by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) |
|
2005 also have "\<dots> = (?I (?l (Dvd j a)))" using cn cnz jnz |
|
2006 using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric] |
|
2007 by (simp add: Let_def split_def int_rdvd_iff[symmetric] |
|
2008 del: real_of_int_mult) (auto simp add: add_ac) |
|
2009 finally have ?case using l jnz by blast } |
|
2010 ultimately show ?case by blast |
|
2011 next |
|
2012 case (12 j a) |
|
2013 let ?c = "fst (zsplit0 a)" |
|
2014 let ?r = "snd (zsplit0 a)" |
|
2015 have spl: "zsplit0 a = (?c,?r)" by simp |
|
2016 from zsplit0_I[OF spl, where x="i" and bs="bs"] |
|
2017 have Ia:"Inum (real i # bs) a = Inum (real i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto |
|
2018 let ?N = "\<lambda> t. Inum (real i#bs) t" |
|
2019 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 |
|
2020 moreover |
|
2021 {assume "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) |
|
2022 hence ?case using prems by (simp del: zlfm.simps add: rdvd_left_0_eq)} |
|
2023 moreover |
|
2024 {assume "?c=0" and "j\<noteq>0" hence ?case |
|
2025 using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] |
|
2026 by (cases "?r", simp_all add: Let_def split_def, case_tac "nat", simp_all)} |
|
2027 moreover |
|
2028 {assume cp: "?c > 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" |
|
2029 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
2030 have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))" |
|
2031 using Ia by (simp add: Let_def split_def) |
|
2032 also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))" |
|
2033 by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp |
|
2034 also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> |
|
2035 (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" |
|
2036 by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) |
|
2037 also have "\<dots> = (?I (?l (NDvd j a)))" using cp cnz jnz |
|
2038 by (simp add: Let_def split_def int_rdvd_iff[symmetric] |
|
2039 del: real_of_int_mult) (auto simp add: add_ac) |
|
2040 finally have ?case using l jnz by simp } |
|
2041 moreover |
|
2042 {assume cn: "?c < 0" and cnz: "?c\<noteq>0" and jnz: "j\<noteq>0" hence l: "?L (?l (NDvd j a))" |
|
2043 by (simp add: nb Let_def split_def isint_Floor isint_neg) |
|
2044 have "?I (NDvd j a) = (\<not> (real j rdvd (real (?c * i) + (?N ?r))))" |
|
2045 using Ia by (simp add: Let_def split_def) |
|
2046 also have "\<dots> = (\<not> (real (abs j) rdvd real (?c*i) + (?N ?r)))" |
|
2047 by (simp only: rdvd_abs1[where d="j" and t="real (?c*i) + ?N ?r", symmetric]) simp |
|
2048 also have "\<dots> = (\<not> ((abs j) dvd (floor ((?N ?r) + real (?c*i))) \<and> |
|
2049 (real (floor ((?N ?r) + real (?c*i))) = (real (?c*i) + (?N ?r)))))" |
|
2050 by(simp only: int_rdvd_real[where i="abs j" and x="real (?c*i) + (?N ?r)"]) (simp only: add_ac) |
|
2051 also have "\<dots> = (?I (?l (NDvd j a)))" using cn cnz jnz |
|
2052 using rdvd_minus [where d="abs j" and t="real (?c*i + floor (?N ?r))", simplified, symmetric] |
|
2053 by (simp add: Let_def split_def int_rdvd_iff[symmetric] |
|
2054 del: real_of_int_mult) (auto simp add: add_ac) |
|
2055 finally have ?case using l jnz by blast } |
|
2056 ultimately show ?case by blast |
|
2057 qed auto |
|
2058 |
|
2059 text{* plusinf : Virtual substitution of @{text "+\<infinity>"} |
|
2060 minusinf: Virtual substitution of @{text "-\<infinity>"} |
|
2061 @{text "\<delta>"} Compute lcm @{text "d| Dvd d c*x+t \<in> p"} |
|
2062 @{text "d\<delta>"} checks if a given l divides all the ds above*} |
|
2063 |
|
2064 consts |
|
2065 plusinf:: "fm \<Rightarrow> fm" |
|
2066 minusinf:: "fm \<Rightarrow> fm" |
|
2067 \<delta> :: "fm \<Rightarrow> int" |
|
2068 d\<delta> :: "fm \<Rightarrow> int \<Rightarrow> bool" |
|
2069 |
|
2070 recdef minusinf "measure size" |
|
2071 "minusinf (And p q) = conj (minusinf p) (minusinf q)" |
|
2072 "minusinf (Or p q) = disj (minusinf p) (minusinf q)" |
|
2073 "minusinf (Eq (CN 0 c e)) = F" |
|
2074 "minusinf (NEq (CN 0 c e)) = T" |
|
2075 "minusinf (Lt (CN 0 c e)) = T" |
|
2076 "minusinf (Le (CN 0 c e)) = T" |
|
2077 "minusinf (Gt (CN 0 c e)) = F" |
|
2078 "minusinf (Ge (CN 0 c e)) = F" |
|
2079 "minusinf p = p" |
|
2080 |
|
2081 lemma minusinf_qfree: "qfree p \<Longrightarrow> qfree (minusinf p)" |
|
2082 by (induct p rule: minusinf.induct, auto) |
|
2083 |
|
2084 recdef plusinf "measure size" |
|
2085 "plusinf (And p q) = conj (plusinf p) (plusinf q)" |
|
2086 "plusinf (Or p q) = disj (plusinf p) (plusinf q)" |
|
2087 "plusinf (Eq (CN 0 c e)) = F" |
|
2088 "plusinf (NEq (CN 0 c e)) = T" |
|
2089 "plusinf (Lt (CN 0 c e)) = F" |
|
2090 "plusinf (Le (CN 0 c e)) = F" |
|
2091 "plusinf (Gt (CN 0 c e)) = T" |
|
2092 "plusinf (Ge (CN 0 c e)) = T" |
|
2093 "plusinf p = p" |
|
2094 |
|
2095 recdef \<delta> "measure size" |
|
2096 "\<delta> (And p q) = zlcm (\<delta> p) (\<delta> q)" |
|
2097 "\<delta> (Or p q) = zlcm (\<delta> p) (\<delta> q)" |
|
2098 "\<delta> (Dvd i (CN 0 c e)) = i" |
|
2099 "\<delta> (NDvd i (CN 0 c e)) = i" |
|
2100 "\<delta> p = 1" |
|
2101 |
|
2102 recdef d\<delta> "measure size" |
|
2103 "d\<delta> (And p q) = (\<lambda> d. d\<delta> p d \<and> d\<delta> q d)" |
|
2104 "d\<delta> (Or p q) = (\<lambda> d. d\<delta> p d \<and> d\<delta> q d)" |
|
2105 "d\<delta> (Dvd i (CN 0 c e)) = (\<lambda> d. i dvd d)" |
|
2106 "d\<delta> (NDvd i (CN 0 c e)) = (\<lambda> d. i dvd d)" |
|
2107 "d\<delta> p = (\<lambda> d. True)" |
|
2108 |
|
2109 lemma delta_mono: |
|
2110 assumes lin: "iszlfm p bs" |
|
2111 and d: "d dvd d'" |
|
2112 and ad: "d\<delta> p d" |
|
2113 shows "d\<delta> p d'" |
|
2114 using lin ad d |
|
2115 proof(induct p rule: iszlfm.induct) |
|
2116 case (9 i c e) thus ?case using d |
|
2117 by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) |
|
2118 next |
|
2119 case (10 i c e) thus ?case using d |
|
2120 by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) |
|
2121 qed simp_all |
|
2122 |
|
2123 lemma \<delta> : assumes lin:"iszlfm p bs" |
|
2124 shows "d\<delta> p (\<delta> p) \<and> \<delta> p >0" |
|
2125 using lin |
|
2126 proof (induct p rule: iszlfm.induct) |
|
2127 case (1 p q) |
|
2128 let ?d = "\<delta> (And p q)" |
|
2129 from prems zlcm_pos have dp: "?d >0" by simp |
|
2130 have d1: "\<delta> p dvd \<delta> (And p q)" using prems by simp |
|
2131 hence th: "d\<delta> p ?d" |
|
2132 using delta_mono prems by (auto simp del: dvd_zlcm_self1) |
|
2133 have "\<delta> q dvd \<delta> (And p q)" using prems by simp |
|
2134 hence th': "d\<delta> q ?d" using delta_mono prems by (auto simp del: dvd_zlcm_self2) |
|
2135 from th th' dp show ?case by simp |
|
2136 next |
|
2137 case (2 p q) |
|
2138 let ?d = "\<delta> (And p q)" |
|
2139 from prems zlcm_pos have dp: "?d >0" by simp |
|
2140 have "\<delta> p dvd \<delta> (And p q)" using prems by simp hence th: "d\<delta> p ?d" using delta_mono prems |
|
2141 by (auto simp del: dvd_zlcm_self1) |
|
2142 have "\<delta> q dvd \<delta> (And p q)" using prems by simp hence th': "d\<delta> q ?d" using delta_mono prems by (auto simp del: dvd_zlcm_self2) |
|
2143 from th th' dp show ?case by simp |
|
2144 qed simp_all |
|
2145 |
|
2146 |
|
2147 lemma minusinf_inf: |
|
2148 assumes linp: "iszlfm p (a # bs)" |
|
2149 shows "\<exists> (z::int). \<forall> x < z. Ifm ((real x)#bs) (minusinf p) = Ifm ((real x)#bs) p" |
|
2150 (is "?P p" is "\<exists> (z::int). \<forall> x < z. ?I x (?M p) = ?I x p") |
|
2151 using linp |
|
2152 proof (induct p rule: minusinf.induct) |
|
2153 case (1 f g) |
|
2154 from prems have "?P f" by simp |
|
2155 then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast |
|
2156 from prems have "?P g" by simp |
|
2157 then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast |
|
2158 let ?z = "min z1 z2" |
|
2159 from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (And f g)) = ?I x (And f g)" by simp |
|
2160 thus ?case by blast |
|
2161 next |
|
2162 case (2 f g) from prems have "?P f" by simp |
|
2163 then obtain z1 where z1_def: "\<forall> x < z1. ?I x (?M f) = ?I x f" by blast |
|
2164 from prems have "?P g" by simp |
|
2165 then obtain z2 where z2_def: "\<forall> x < z2. ?I x (?M g) = ?I x g" by blast |
|
2166 let ?z = "min z1 z2" |
|
2167 from z1_def z2_def have "\<forall> x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp |
|
2168 thus ?case by blast |
|
2169 next |
|
2170 case (3 c e) |
|
2171 from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp |
|
2172 from prems have nbe: "numbound0 e" by simp |
|
2173 fix y |
|
2174 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))" |
|
2175 proof (simp add: less_floor_eq , rule allI, rule impI) |
|
2176 fix x |
|
2177 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)" |
|
2178 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp |
|
2179 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" |
|
2180 by (simp only: real_mult_less_mono2[OF rcpos th1]) |
|
2181 hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos by simp |
|
2182 thus "real c * real x + Inum (real x # bs) e \<noteq> 0" |
|
2183 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp |
|
2184 qed |
|
2185 thus ?case by blast |
|
2186 next |
|
2187 case (4 c e) |
|
2188 from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp |
|
2189 from prems have nbe: "numbound0 e" by simp |
|
2190 fix y |
|
2191 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))" |
|
2192 proof (simp add: less_floor_eq , rule allI, rule impI) |
|
2193 fix x |
|
2194 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)" |
|
2195 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp |
|
2196 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" |
|
2197 by (simp only: real_mult_less_mono2[OF rcpos th1]) |
|
2198 hence "real c * real x + Inum (y # bs) e \<noteq> 0"using rcpos by simp |
|
2199 thus "real c * real x + Inum (real x # bs) e \<noteq> 0" |
|
2200 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] by simp |
|
2201 qed |
|
2202 thus ?case by blast |
|
2203 next |
|
2204 case (5 c e) |
|
2205 from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp |
|
2206 from prems have nbe: "numbound0 e" by simp |
|
2207 fix y |
|
2208 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))" |
|
2209 proof (simp add: less_floor_eq , rule allI, rule impI) |
|
2210 fix x |
|
2211 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)" |
|
2212 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp |
|
2213 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" |
|
2214 by (simp only: real_mult_less_mono2[OF rcpos th1]) |
|
2215 thus "real c * real x + Inum (real x # bs) e < 0" |
|
2216 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp |
|
2217 qed |
|
2218 thus ?case by blast |
|
2219 next |
|
2220 case (6 c e) |
|
2221 from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp |
|
2222 from prems have nbe: "numbound0 e" by simp |
|
2223 fix y |
|
2224 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))" |
|
2225 proof (simp add: less_floor_eq , rule allI, rule impI) |
|
2226 fix x |
|
2227 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)" |
|
2228 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp |
|
2229 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" |
|
2230 by (simp only: real_mult_less_mono2[OF rcpos th1]) |
|
2231 thus "real c * real x + Inum (real x # bs) e \<le> 0" |
|
2232 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp |
|
2233 qed |
|
2234 thus ?case by blast |
|
2235 next |
|
2236 case (7 c e) |
|
2237 from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp |
|
2238 from prems have nbe: "numbound0 e" by simp |
|
2239 fix y |
|
2240 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))" |
|
2241 proof (simp add: less_floor_eq , rule allI, rule impI) |
|
2242 fix x |
|
2243 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)" |
|
2244 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp |
|
2245 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" |
|
2246 by (simp only: real_mult_less_mono2[OF rcpos th1]) |
|
2247 thus "\<not> (real c * real x + Inum (real x # bs) e>0)" |
|
2248 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp |
|
2249 qed |
|
2250 thus ?case by blast |
|
2251 next |
|
2252 case (8 c e) |
|
2253 from prems have "c > 0" by simp hence rcpos: "real c > 0" by simp |
|
2254 from prems have nbe: "numbound0 e" by simp |
|
2255 fix y |
|
2256 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))" |
|
2257 proof (simp add: less_floor_eq , rule allI, rule impI) |
|
2258 fix x |
|
2259 assume A: "real x + (1\<Colon>real) \<le> - (Inum (y # bs) e / real c)" |
|
2260 hence th1:"real x < - (Inum (y # bs) e / real c)" by simp |
|
2261 with rcpos have "(real c)*(real x) < (real c)*(- (Inum (y # bs) e / real c))" |
|
2262 by (simp only: real_mult_less_mono2[OF rcpos th1]) |
|
2263 thus "\<not> real c * real x + Inum (real x # bs) e \<ge> 0" |
|
2264 using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real x"] rcpos by simp |
|
2265 qed |
|
2266 thus ?case by blast |
|
2267 qed simp_all |
|
2268 |
|
2269 lemma minusinf_repeats: |
|
2270 assumes d: "d\<delta> p d" and linp: "iszlfm p (a # bs)" |
|
2271 shows "Ifm ((real(x - k*d))#bs) (minusinf p) = Ifm (real x #bs) (minusinf p)" |
|
2272 using linp d |
|
2273 proof(induct p rule: iszlfm.induct) |
|
2274 case (9 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ |
|
2275 hence "\<exists> k. d=i*k" by (simp add: dvd_def) |
|
2276 then obtain "di" where di_def: "d=i*di" by blast |
|
2277 show ?case |
|
2278 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) |
|
2279 assume |
|
2280 "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" |
|
2281 (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") |
|
2282 hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def) |
|
2283 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" |
|
2284 by (simp add: algebra_simps di_def) |
|
2285 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))" |
|
2286 by (simp add: algebra_simps) |
|
2287 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast |
|
2288 thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp |
|
2289 next |
|
2290 assume |
|
2291 "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e") |
|
2292 hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def) |
|
2293 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp |
|
2294 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) |
|
2295 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps) |
|
2296 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)" |
|
2297 by blast |
|
2298 thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp |
|
2299 qed |
|
2300 next |
|
2301 case (10 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ |
|
2302 hence "\<exists> k. d=i*k" by (simp add: dvd_def) |
|
2303 then obtain "di" where di_def: "d=i*di" by blast |
|
2304 show ?case |
|
2305 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) |
|
2306 assume |
|
2307 "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" |
|
2308 (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") |
|
2309 hence "\<exists> (l::int). ?rt = ?ri * (real l)" by (simp add: rdvd_def) |
|
2310 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real l)+?rc*(?rk * (real i) * (real di))" |
|
2311 by (simp add: algebra_simps di_def) |
|
2312 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri*(real (l + c*k*di))" |
|
2313 by (simp add: algebra_simps) |
|
2314 hence "\<exists> (l::int). ?rc*?rx+ ?I x e = ?ri* (real l)" by blast |
|
2315 thus "real i rdvd real c * real x + Inum (real x # bs) e" using rdvd_def by simp |
|
2316 next |
|
2317 assume |
|
2318 "real i rdvd real c * real x + Inum (real x # bs) e" (is "?ri rdvd ?rc*?rx+?e") |
|
2319 hence "\<exists> (l::int). ?rc*?rx+?e = ?ri * (real l)" by (simp add: rdvd_def) |
|
2320 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l) - real c * (real k * real d)" by simp |
|
2321 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) |
|
2322 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real (l - c*k*di))" by (simp add: algebra_simps) |
|
2323 hence "\<exists> (l::int). ?rc*?rx - real c * (real k * real d) +?e = ?ri * (real l)" |
|
2324 by blast |
|
2325 thus "real i rdvd real c * real x - real c * (real k * real d) + Inum (real x # bs) e" using rdvd_def by simp |
|
2326 qed |
|
2327 qed (auto simp add: nth_pos2 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) |
|
2328 |
|
2329 lemma minusinf_ex: |
|
2330 assumes lin: "iszlfm p (real (a::int) #bs)" |
|
2331 and exmi: "\<exists> (x::int). Ifm (real x#bs) (minusinf p)" (is "\<exists> x. ?P1 x") |
|
2332 shows "\<exists> (x::int). Ifm (real x#bs) p" (is "\<exists> x. ?P x") |
|
2333 proof- |
|
2334 let ?d = "\<delta> p" |
|
2335 from \<delta> [OF lin] have dpos: "?d >0" by simp |
|
2336 from \<delta> [OF lin] have alld: "d\<delta> p ?d" by simp |
|
2337 from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P1 x = ?P1 (x - (k * ?d))" by simp |
|
2338 from minusinf_inf[OF lin] have th2:"\<exists> z. \<forall> x. x<z \<longrightarrow> (?P x = ?P1 x)" by blast |
|
2339 from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast |
|
2340 qed |
|
2341 |
|
2342 lemma minusinf_bex: |
|
2343 assumes lin: "iszlfm p (real (a::int) #bs)" |
|
2344 shows "(\<exists> (x::int). Ifm (real x#bs) (minusinf p)) = |
|
2345 (\<exists> (x::int)\<in> {1..\<delta> p}. Ifm (real x#bs) (minusinf p))" |
|
2346 (is "(\<exists> x. ?P x) = _") |
|
2347 proof- |
|
2348 let ?d = "\<delta> p" |
|
2349 from \<delta> [OF lin] have dpos: "?d >0" by simp |
|
2350 from \<delta> [OF lin] have alld: "d\<delta> p ?d" by simp |
|
2351 from minusinf_repeats[OF alld lin] have th1:"\<forall> x k. ?P x = ?P (x - (k * ?d))" by simp |
|
2352 from periodic_finite_ex[OF dpos th1] show ?thesis by blast |
|
2353 qed |
|
2354 |
|
2355 lemma dvd1_eq1: "x >0 \<Longrightarrow> (x::int) dvd 1 = (x = 1)" by auto |
|
2356 |
|
2357 consts |
|
2358 a\<beta> :: "fm \<Rightarrow> int \<Rightarrow> fm" (* adjusts the coeffitients of a formula *) |
|
2359 d\<beta> :: "fm \<Rightarrow> int \<Rightarrow> bool" (* tests if all coeffs c of c divide a given l*) |
|
2360 \<zeta> :: "fm \<Rightarrow> int" (* computes the lcm of all coefficients of x*) |
|
2361 \<beta> :: "fm \<Rightarrow> num list" |
|
2362 \<alpha> :: "fm \<Rightarrow> num list" |
|
2363 |
|
2364 recdef a\<beta> "measure size" |
|
2365 "a\<beta> (And p q) = (\<lambda> k. And (a\<beta> p k) (a\<beta> q k))" |
|
2366 "a\<beta> (Or p q) = (\<lambda> k. Or (a\<beta> p k) (a\<beta> q k))" |
|
2367 "a\<beta> (Eq (CN 0 c e)) = (\<lambda> k. Eq (CN 0 1 (Mul (k div c) e)))" |
|
2368 "a\<beta> (NEq (CN 0 c e)) = (\<lambda> k. NEq (CN 0 1 (Mul (k div c) e)))" |
|
2369 "a\<beta> (Lt (CN 0 c e)) = (\<lambda> k. Lt (CN 0 1 (Mul (k div c) e)))" |
|
2370 "a\<beta> (Le (CN 0 c e)) = (\<lambda> k. Le (CN 0 1 (Mul (k div c) e)))" |
|
2371 "a\<beta> (Gt (CN 0 c e)) = (\<lambda> k. Gt (CN 0 1 (Mul (k div c) e)))" |
|
2372 "a\<beta> (Ge (CN 0 c e)) = (\<lambda> k. Ge (CN 0 1 (Mul (k div c) e)))" |
|
2373 "a\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" |
|
2374 "a\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" |
|
2375 "a\<beta> p = (\<lambda> k. p)" |
|
2376 |
|
2377 recdef d\<beta> "measure size" |
|
2378 "d\<beta> (And p q) = (\<lambda> k. (d\<beta> p k) \<and> (d\<beta> q k))" |
|
2379 "d\<beta> (Or p q) = (\<lambda> k. (d\<beta> p k) \<and> (d\<beta> q k))" |
|
2380 "d\<beta> (Eq (CN 0 c e)) = (\<lambda> k. c dvd k)" |
|
2381 "d\<beta> (NEq (CN 0 c e)) = (\<lambda> k. c dvd k)" |
|
2382 "d\<beta> (Lt (CN 0 c e)) = (\<lambda> k. c dvd k)" |
|
2383 "d\<beta> (Le (CN 0 c e)) = (\<lambda> k. c dvd k)" |
|
2384 "d\<beta> (Gt (CN 0 c e)) = (\<lambda> k. c dvd k)" |
|
2385 "d\<beta> (Ge (CN 0 c e)) = (\<lambda> k. c dvd k)" |
|
2386 "d\<beta> (Dvd i (CN 0 c e)) =(\<lambda> k. c dvd k)" |
|
2387 "d\<beta> (NDvd i (CN 0 c e))=(\<lambda> k. c dvd k)" |
|
2388 "d\<beta> p = (\<lambda> k. True)" |
|
2389 |
|
2390 recdef \<zeta> "measure size" |
|
2391 "\<zeta> (And p q) = zlcm (\<zeta> p) (\<zeta> q)" |
|
2392 "\<zeta> (Or p q) = zlcm (\<zeta> p) (\<zeta> q)" |
|
2393 "\<zeta> (Eq (CN 0 c e)) = c" |
|
2394 "\<zeta> (NEq (CN 0 c e)) = c" |
|
2395 "\<zeta> (Lt (CN 0 c e)) = c" |
|
2396 "\<zeta> (Le (CN 0 c e)) = c" |
|
2397 "\<zeta> (Gt (CN 0 c e)) = c" |
|
2398 "\<zeta> (Ge (CN 0 c e)) = c" |
|
2399 "\<zeta> (Dvd i (CN 0 c e)) = c" |
|
2400 "\<zeta> (NDvd i (CN 0 c e))= c" |
|
2401 "\<zeta> p = 1" |
|
2402 |
|
2403 recdef \<beta> "measure size" |
|
2404 "\<beta> (And p q) = (\<beta> p @ \<beta> q)" |
|
2405 "\<beta> (Or p q) = (\<beta> p @ \<beta> q)" |
|
2406 "\<beta> (Eq (CN 0 c e)) = [Sub (C -1) e]" |
|
2407 "\<beta> (NEq (CN 0 c e)) = [Neg e]" |
|
2408 "\<beta> (Lt (CN 0 c e)) = []" |
|
2409 "\<beta> (Le (CN 0 c e)) = []" |
|
2410 "\<beta> (Gt (CN 0 c e)) = [Neg e]" |
|
2411 "\<beta> (Ge (CN 0 c e)) = [Sub (C -1) e]" |
|
2412 "\<beta> p = []" |
|
2413 |
|
2414 recdef \<alpha> "measure size" |
|
2415 "\<alpha> (And p q) = (\<alpha> p @ \<alpha> q)" |
|
2416 "\<alpha> (Or p q) = (\<alpha> p @ \<alpha> q)" |
|
2417 "\<alpha> (Eq (CN 0 c e)) = [Add (C -1) e]" |
|
2418 "\<alpha> (NEq (CN 0 c e)) = [e]" |
|
2419 "\<alpha> (Lt (CN 0 c e)) = [e]" |
|
2420 "\<alpha> (Le (CN 0 c e)) = [Add (C -1) e]" |
|
2421 "\<alpha> (Gt (CN 0 c e)) = []" |
|
2422 "\<alpha> (Ge (CN 0 c e)) = []" |
|
2423 "\<alpha> p = []" |
|
2424 consts mirror :: "fm \<Rightarrow> fm" |
|
2425 recdef mirror "measure size" |
|
2426 "mirror (And p q) = And (mirror p) (mirror q)" |
|
2427 "mirror (Or p q) = Or (mirror p) (mirror q)" |
|
2428 "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))" |
|
2429 "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))" |
|
2430 "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))" |
|
2431 "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))" |
|
2432 "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))" |
|
2433 "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))" |
|
2434 "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))" |
|
2435 "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))" |
|
2436 "mirror p = p" |
|
2437 |
|
2438 lemma mirror\<alpha>\<beta>: |
|
2439 assumes lp: "iszlfm p (a#bs)" |
|
2440 shows "(Inum (real (i::int)#bs)) ` set (\<alpha> p) = (Inum (real i#bs)) ` set (\<beta> (mirror p))" |
|
2441 using lp |
|
2442 by (induct p rule: mirror.induct, auto) |
|
2443 |
|
2444 lemma mirror: |
|
2445 assumes lp: "iszlfm p (a#bs)" |
|
2446 shows "Ifm (real (x::int)#bs) (mirror p) = Ifm (real (- x)#bs) p" |
|
2447 using lp |
|
2448 proof(induct p rule: iszlfm.induct) |
|
2449 case (9 j c e) |
|
2450 have th: "(real j rdvd real c * real x - Inum (real x # bs) e) = |
|
2451 (real j rdvd - (real c * real x - Inum (real x # bs) e))" |
|
2452 by (simp only: rdvd_minus[symmetric]) |
|
2453 from prems th show ?case |
|
2454 by (simp add: algebra_simps |
|
2455 numbound0_I[where bs="bs" and b'="real x" and b="- real x"]) |
|
2456 next |
|
2457 case (10 j c e) |
|
2458 have th: "(real j rdvd real c * real x - Inum (real x # bs) e) = |
|
2459 (real j rdvd - (real c * real x - Inum (real x # bs) e))" |
|
2460 by (simp only: rdvd_minus[symmetric]) |
|
2461 from prems th show ?case |
|
2462 by (simp add: algebra_simps |
|
2463 numbound0_I[where bs="bs" and b'="real x" and b="- real x"]) |
|
2464 qed (auto simp add: numbound0_I[where bs="bs" and b="real x" and b'="- real x"] nth_pos2) |
|
2465 |
|
2466 lemma mirror_l: "iszlfm p (a#bs) \<Longrightarrow> iszlfm (mirror p) (a#bs)" |
|
2467 by (induct p rule: mirror.induct, auto simp add: isint_neg) |
|
2468 |
|
2469 lemma mirror_d\<beta>: "iszlfm p (a#bs) \<and> d\<beta> p 1 |
|
2470 \<Longrightarrow> iszlfm (mirror p) (a#bs) \<and> d\<beta> (mirror p) 1" |
|
2471 by (induct p rule: mirror.induct, auto simp add: isint_neg) |
|
2472 |
|
2473 lemma mirror_\<delta>: "iszlfm p (a#bs) \<Longrightarrow> \<delta> (mirror p) = \<delta> p" |
|
2474 by (induct p rule: mirror.induct,auto) |
|
2475 |
|
2476 |
|
2477 lemma mirror_ex: |
|
2478 assumes lp: "iszlfm p (real (i::int)#bs)" |
|
2479 shows "(\<exists> (x::int). Ifm (real x#bs) (mirror p)) = (\<exists> (x::int). Ifm (real x#bs) p)" |
|
2480 (is "(\<exists> x. ?I x ?mp) = (\<exists> x. ?I x p)") |
|
2481 proof(auto) |
|
2482 fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast |
|
2483 thus "\<exists> x. ?I x p" by blast |
|
2484 next |
|
2485 fix x assume "?I x p" hence "?I (- x) ?mp" |
|
2486 using mirror[OF lp, where x="- x", symmetric] by auto |
|
2487 thus "\<exists> x. ?I x ?mp" by blast |
|
2488 qed |
|
2489 |
|
2490 lemma \<beta>_numbound0: assumes lp: "iszlfm p bs" |
|
2491 shows "\<forall> b\<in> set (\<beta> p). numbound0 b" |
|
2492 using lp by (induct p rule: \<beta>.induct,auto) |
|
2493 |
|
2494 lemma d\<beta>_mono: |
|
2495 assumes linp: "iszlfm p (a #bs)" |
|
2496 and dr: "d\<beta> p l" |
|
2497 and d: "l dvd l'" |
|
2498 shows "d\<beta> p l'" |
|
2499 using dr linp zdvd_trans[where n="l" and k="l'", simplified d] |
|
2500 by (induct p rule: iszlfm.induct) simp_all |
|
2501 |
|
2502 lemma \<alpha>_l: assumes lp: "iszlfm p (a#bs)" |
|
2503 shows "\<forall> b\<in> set (\<alpha> p). numbound0 b \<and> isint b (a#bs)" |
|
2504 using lp |
|
2505 by(induct p rule: \<alpha>.induct, auto simp add: isint_add isint_c) |
|
2506 |
|
2507 lemma \<zeta>: |
|
2508 assumes linp: "iszlfm p (a #bs)" |
|
2509 shows "\<zeta> p > 0 \<and> d\<beta> p (\<zeta> p)" |
|
2510 using linp |
|
2511 proof(induct p rule: iszlfm.induct) |
|
2512 case (1 p q) |
|
2513 from prems have dl1: "\<zeta> p dvd zlcm (\<zeta> p) (\<zeta> q)" by simp |
|
2514 from prems have dl2: "\<zeta> q dvd zlcm (\<zeta> p) (\<zeta> q)" by simp |
|
2515 from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="zlcm (\<zeta> p) (\<zeta> q)"] |
|
2516 d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="zlcm (\<zeta> p) (\<zeta> q)"] |
|
2517 dl1 dl2 show ?case by (auto simp add: zlcm_pos) |
|
2518 next |
|
2519 case (2 p q) |
|
2520 from prems have dl1: "\<zeta> p dvd zlcm (\<zeta> p) (\<zeta> q)" by simp |
|
2521 from prems have dl2: "\<zeta> q dvd zlcm (\<zeta> p) (\<zeta> q)" by simp |
|
2522 from prems d\<beta>_mono[where p = "p" and l="\<zeta> p" and l'="zlcm (\<zeta> p) (\<zeta> q)"] |
|
2523 d\<beta>_mono[where p = "q" and l="\<zeta> q" and l'="zlcm (\<zeta> p) (\<zeta> q)"] |
|
2524 dl1 dl2 show ?case by (auto simp add: zlcm_pos) |
|
2525 qed (auto simp add: zlcm_pos) |
|
2526 |
|
2527 lemma a\<beta>: assumes linp: "iszlfm p (a #bs)" and d: "d\<beta> p l" and lp: "l > 0" |
|
2528 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)" |
|
2529 using linp d |
|
2530 proof (induct p rule: iszlfm.induct) |
|
2531 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+ |
|
2532 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) |
|
2533 from cp have cnz: "c \<noteq> 0" by simp |
|
2534 have "c div c\<le> l div c" |
|
2535 by (simp add: zdiv_mono1[OF clel cp]) |
|
2536 then have ldcp:"0 < l div c" |
|
2537 by (simp add: zdiv_self[OF cnz]) |
|
2538 have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp |
|
2539 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] |
|
2540 by simp |
|
2541 hence "(real l * real x + real (l div c) * Inum (real x # bs) e < (0\<Colon>real)) = |
|
2542 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e < 0)" |
|
2543 by simp |
|
2544 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) |
|
2545 also have "\<dots> = (real c * real x + Inum (real x # bs) e < 0)" |
|
2546 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 |
|
2547 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 |
|
2548 next |
|
2549 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+ |
|
2550 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) |
|
2551 from cp have cnz: "c \<noteq> 0" by simp |
|
2552 have "c div c\<le> l div c" |
|
2553 by (simp add: zdiv_mono1[OF clel cp]) |
|
2554 then have ldcp:"0 < l div c" |
|
2555 by (simp add: zdiv_self[OF cnz]) |
|
2556 have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp |
|
2557 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] |
|
2558 by simp |
|
2559 hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<le> (0\<Colon>real)) = |
|
2560 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<le> 0)" |
|
2561 by simp |
|
2562 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) |
|
2563 also have "\<dots> = (real c * real x + Inum (real x # bs) e \<le> 0)" |
|
2564 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 |
|
2565 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 |
|
2566 next |
|
2567 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+ |
|
2568 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) |
|
2569 from cp have cnz: "c \<noteq> 0" by simp |
|
2570 have "c div c\<le> l div c" |
|
2571 by (simp add: zdiv_mono1[OF clel cp]) |
|
2572 then have ldcp:"0 < l div c" |
|
2573 by (simp add: zdiv_self[OF cnz]) |
|
2574 have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp |
|
2575 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] |
|
2576 by simp |
|
2577 hence "(real l * real x + real (l div c) * Inum (real x # bs) e > (0\<Colon>real)) = |
|
2578 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e > 0)" |
|
2579 by simp |
|
2580 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) |
|
2581 also have "\<dots> = (real c * real x + Inum (real x # bs) e > 0)" |
|
2582 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 |
|
2583 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 |
|
2584 next |
|
2585 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+ |
|
2586 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) |
|
2587 from cp have cnz: "c \<noteq> 0" by simp |
|
2588 have "c div c\<le> l div c" |
|
2589 by (simp add: zdiv_mono1[OF clel cp]) |
|
2590 then have ldcp:"0 < l div c" |
|
2591 by (simp add: zdiv_self[OF cnz]) |
|
2592 have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp |
|
2593 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] |
|
2594 by simp |
|
2595 hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<ge> (0\<Colon>real)) = |
|
2596 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<ge> 0)" |
|
2597 by simp |
|
2598 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) |
|
2599 also have "\<dots> = (real c * real x + Inum (real x # bs) e \<ge> 0)" |
|
2600 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 |
|
2601 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 |
|
2602 next |
|
2603 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+ |
|
2604 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) |
|
2605 from cp have cnz: "c \<noteq> 0" by simp |
|
2606 have "c div c\<le> l div c" |
|
2607 by (simp add: zdiv_mono1[OF clel cp]) |
|
2608 then have ldcp:"0 < l div c" |
|
2609 by (simp add: zdiv_self[OF cnz]) |
|
2610 have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp |
|
2611 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] |
|
2612 by simp |
|
2613 hence "(real l * real x + real (l div c) * Inum (real x # bs) e = (0\<Colon>real)) = |
|
2614 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e = 0)" |
|
2615 by simp |
|
2616 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) |
|
2617 also have "\<dots> = (real c * real x + Inum (real x # bs) e = 0)" |
|
2618 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 |
|
2619 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 |
|
2620 next |
|
2621 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+ |
|
2622 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) |
|
2623 from cp have cnz: "c \<noteq> 0" by simp |
|
2624 have "c div c\<le> l div c" |
|
2625 by (simp add: zdiv_mono1[OF clel cp]) |
|
2626 then have ldcp:"0 < l div c" |
|
2627 by (simp add: zdiv_self[OF cnz]) |
|
2628 have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp |
|
2629 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] |
|
2630 by simp |
|
2631 hence "(real l * real x + real (l div c) * Inum (real x # bs) e \<noteq> (0\<Colon>real)) = |
|
2632 (real (c * (l div c)) * real x + real (l div c) * Inum (real x # bs) e \<noteq> 0)" |
|
2633 by simp |
|
2634 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) |
|
2635 also have "\<dots> = (real c * real x + Inum (real x # bs) e \<noteq> 0)" |
|
2636 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 |
|
2637 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 |
|
2638 next |
|
2639 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+ |
|
2640 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) |
|
2641 from cp have cnz: "c \<noteq> 0" by simp |
|
2642 have "c div c\<le> l div c" |
|
2643 by (simp add: zdiv_mono1[OF clel cp]) |
|
2644 then have ldcp:"0 < l div c" |
|
2645 by (simp add: zdiv_self[OF cnz]) |
|
2646 have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp |
|
2647 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] |
|
2648 by simp |
|
2649 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 |
|
2650 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) |
|
2651 also fix k have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)" |
|
2652 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 |
|
2653 also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp |
|
2654 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 |
|
2655 next |
|
2656 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+ |
|
2657 from lp cp have clel: "c\<le>l" by (simp add: zdvd_imp_le [OF d' lp]) |
|
2658 from cp have cnz: "c \<noteq> 0" by simp |
|
2659 have "c div c\<le> l div c" |
|
2660 by (simp add: zdiv_mono1[OF clel cp]) |
|
2661 then have ldcp:"0 < l div c" |
|
2662 by (simp add: zdiv_self[OF cnz]) |
|
2663 have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp |
|
2664 hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] |
|
2665 by simp |
|
2666 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 |
|
2667 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) |
|
2668 also fix k have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e - real j * real k = 0)" |
|
2669 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 |
|
2670 also have "\<dots> = (\<exists> (k::int). real c * real x + Inum (real x # bs) e = real j * real k)" by simp |
|
2671 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 |
|
2672 qed (simp_all add: nth_pos2 numbound0_I[where bs="bs" and b="real (l * x)" and b'="real x"] isint_Mul del: real_of_int_mult) |
|
2673 |
|
2674 lemma a\<beta>_ex: assumes linp: "iszlfm p (a#bs)" and d: "d\<beta> p l" and lp: "l>0" |
|
2675 shows "(\<exists> x. l dvd x \<and> Ifm (real x #bs) (a\<beta> p l)) = (\<exists> (x::int). Ifm (real x#bs) p)" |
|
2676 (is "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> x. ?P' x)") |
|
2677 proof- |
|
2678 have "(\<exists> x. l dvd x \<and> ?P x) = (\<exists> (x::int). ?P (l*x))" |
|
2679 using unity_coeff_ex[where l="l" and P="?P", simplified] by simp |
|
2680 also have "\<dots> = (\<exists> (x::int). ?P' x)" using a\<beta>[OF linp d lp] by simp |
|
2681 finally show ?thesis . |
|
2682 qed |
|
2683 |
|
2684 lemma \<beta>: |
|
2685 assumes lp: "iszlfm p (a#bs)" |
|
2686 and u: "d\<beta> p 1" |
|
2687 and d: "d\<delta> p d" |
|
2688 and dp: "d > 0" |
|
2689 and nob: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)" |
|
2690 and p: "Ifm (real x#bs) p" (is "?P x") |
|
2691 shows "?P (x - d)" |
|
2692 using lp u d dp nob p |
|
2693 proof(induct p rule: iszlfm.induct) |
|
2694 case (5 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ |
|
2695 with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems |
|
2696 show ?case by (simp del: real_of_int_minus) |
|
2697 next |
|
2698 case (6 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ |
|
2699 with dp p c1 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] prems |
|
2700 show ?case by (simp del: real_of_int_minus) |
|
2701 next |
|
2702 case (7 c e) hence p: "Ifm (real x #bs) (Gt (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" and ie1:"isint e (a#bs)" using dvd1_eq1[where x="c"] by simp+ |
|
2703 let ?e = "Inum (real x # bs) e" |
|
2704 from ie1 have ie: "real (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"] |
|
2705 numbound0_I[OF bn,where b="a" and b'="real x" and bs="bs"] |
|
2706 by (simp add: isint_iff) |
|
2707 {assume "real (x-d) +?e > 0" hence ?case using c1 |
|
2708 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] |
|
2709 by (simp del: real_of_int_minus)} |
|
2710 moreover |
|
2711 {assume H: "\<not> real (x-d) + ?e > 0" |
|
2712 let ?v="Neg e" |
|
2713 have vb: "?v \<in> set (\<beta> (Gt (CN 0 c e)))" by simp |
|
2714 from prems(11)[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"]] |
|
2715 have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x = - ?e + real j)" by auto |
|
2716 from H p have "real x + ?e > 0 \<and> real x + ?e \<le> real d" by (simp add: c1) |
|
2717 hence "real (x + floor ?e) > real (0::int) \<and> real (x + floor ?e) \<le> real d" |
|
2718 using ie by simp |
|
2719 hence "x + floor ?e \<ge> 1 \<and> x + floor ?e \<le> d" by simp |
|
2720 hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e" by simp |
|
2721 hence "\<exists> (j::int) \<in> {1 .. d}. real x = real (- floor ?e + j)" |
|
2722 by (simp only: real_of_int_inject) (simp add: algebra_simps) |
|
2723 hence "\<exists> (j::int) \<in> {1 .. d}. real x = - ?e + real j" |
|
2724 by (simp add: ie[simplified isint_iff]) |
|
2725 with nob have ?case by auto} |
|
2726 ultimately show ?case by blast |
|
2727 next |
|
2728 case (8 c e) hence p: "Ifm (real x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" |
|
2729 and ie1:"isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ |
|
2730 let ?e = "Inum (real x # bs) e" |
|
2731 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"] |
|
2732 by (simp add: isint_iff) |
|
2733 {assume "real (x-d) +?e \<ge> 0" hence ?case using c1 |
|
2734 numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] |
|
2735 by (simp del: real_of_int_minus)} |
|
2736 moreover |
|
2737 {assume H: "\<not> real (x-d) + ?e \<ge> 0" |
|
2738 let ?v="Sub (C -1) e" |
|
2739 have vb: "?v \<in> set (\<beta> (Ge (CN 0 c e)))" by simp |
|
2740 from prems(11)[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"]] |
|
2741 have nob: "\<not> (\<exists> j\<in> {1 ..d}. real x = - ?e - 1 + real j)" by auto |
|
2742 from H p have "real x + ?e \<ge> 0 \<and> real x + ?e < real d" by (simp add: c1) |
|
2743 hence "real (x + floor ?e) \<ge> real (0::int) \<and> real (x + floor ?e) < real d" |
|
2744 using ie by simp |
|
2745 hence "x + floor ?e +1 \<ge> 1 \<and> x + floor ?e + 1 \<le> d" by simp |
|
2746 hence "\<exists> (j::int) \<in> {1 .. d}. j = x + floor ?e + 1" by simp |
|
2747 hence "\<exists> (j::int) \<in> {1 .. d}. x= - floor ?e - 1 + j" by (simp add: algebra_simps) |
|
2748 hence "\<exists> (j::int) \<in> {1 .. d}. real x= real (- floor ?e - 1 + j)" |
|
2749 by (simp only: real_of_int_inject) |
|
2750 hence "\<exists> (j::int) \<in> {1 .. d}. real x= - ?e - 1 + real j" |
|
2751 by (simp add: ie[simplified isint_iff]) |
|
2752 with nob have ?case by simp } |
|
2753 ultimately show ?case by blast |
|
2754 next |
|
2755 case (3 c e) hence p: "Ifm (real x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" |
|
2756 and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ |
|
2757 let ?e = "Inum (real x # bs) e" |
|
2758 let ?v="(Sub (C -1) e)" |
|
2759 have vb: "?v \<in> set (\<beta> (Eq (CN 0 c e)))" by simp |
|
2760 from p have "real x= - ?e" by (simp add: c1) with prems(11) show ?case using dp |
|
2761 by simp (erule ballE[where x="1"], |
|
2762 simp_all add:algebra_simps numbound0_I[OF bn,where b="real x"and b'="a"and bs="bs"]) |
|
2763 next |
|
2764 case (4 c e)hence p: "Ifm (real x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" |
|
2765 and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ |
|
2766 let ?e = "Inum (real x # bs) e" |
|
2767 let ?v="Neg e" |
|
2768 have vb: "?v \<in> set (\<beta> (NEq (CN 0 c e)))" by simp |
|
2769 {assume "real x - real d + Inum ((real (x -d)) # bs) e \<noteq> 0" |
|
2770 hence ?case by (simp add: c1)} |
|
2771 moreover |
|
2772 {assume H: "real x - real d + Inum ((real (x -d)) # bs) e = 0" |
|
2773 hence "real x = - Inum ((real (x -d)) # bs) e + real d" by simp |
|
2774 hence "real x = - Inum (a # bs) e + real d" |
|
2775 by (simp add: numbound0_I[OF bn,where b="real x - real d"and b'="a"and bs="bs"]) |
|
2776 with prems(11) have ?case using dp by simp} |
|
2777 ultimately show ?case by blast |
|
2778 next |
|
2779 case (9 j c e) hence p: "Ifm (real x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" |
|
2780 and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ |
|
2781 let ?e = "Inum (real x # bs) e" |
|
2782 from prems have "isint e (a #bs)" by simp |
|
2783 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"] |
|
2784 by (simp add: isint_iff) |
|
2785 from prems have id: "j dvd d" by simp |
|
2786 from c1 ie[symmetric] have "?p x = (real j rdvd real (x+ floor ?e))" by simp |
|
2787 also have "\<dots> = (j dvd x + floor ?e)" |
|
2788 using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp |
|
2789 also have "\<dots> = (j dvd x - d + floor ?e)" |
|
2790 using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp |
|
2791 also have "\<dots> = (real j rdvd real (x - d + floor ?e))" |
|
2792 using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified] |
|
2793 ie by simp |
|
2794 also have "\<dots> = (real j rdvd real x - real d + ?e)" |
|
2795 using ie by simp |
|
2796 finally show ?case |
|
2797 using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp |
|
2798 next |
|
2799 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+ |
|
2800 let ?e = "Inum (real x # bs) e" |
|
2801 from prems have "isint e (a#bs)" by simp |
|
2802 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"] |
|
2803 by (simp add: isint_iff) |
|
2804 from prems have id: "j dvd d" by simp |
|
2805 from c1 ie[symmetric] have "?p x = (\<not> real j rdvd real (x+ floor ?e))" by simp |
|
2806 also have "\<dots> = (\<not> j dvd x + floor ?e)" |
|
2807 using int_rdvd_real[where i="j" and x="real (x+ floor ?e)"] by simp |
|
2808 also have "\<dots> = (\<not> j dvd x - d + floor ?e)" |
|
2809 using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp |
|
2810 also have "\<dots> = (\<not> real j rdvd real (x - d + floor ?e))" |
|
2811 using int_rdvd_real[where i="j" and x="real (x-d + floor ?e)",symmetric, simplified] |
|
2812 ie by simp |
|
2813 also have "\<dots> = (\<not> real j rdvd real x - real d + ?e)" |
|
2814 using ie by simp |
|
2815 finally show ?case using numbound0_I[OF bn,where b="real (x-d)" and b'="real x" and bs="bs"] c1 p by simp |
|
2816 qed (auto simp add: numbound0_I[where bs="bs" and b="real (x - d)" and b'="real x"] nth_pos2 simp del: real_of_int_diff) |
|
2817 |
|
2818 lemma \<beta>': |
|
2819 assumes lp: "iszlfm p (a #bs)" |
|
2820 and u: "d\<beta> p 1" |
|
2821 and d: "d\<delta> p d" |
|
2822 and dp: "d > 0" |
|
2823 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)") |
|
2824 proof(clarify) |
|
2825 fix x |
|
2826 assume nb:"?b" and px: "?P x" |
|
2827 hence nb2: "\<not>(\<exists>(j::int) \<in> {1 .. d}. \<exists> b\<in> (Inum (a#bs)) ` set(\<beta> p). real x = b + real j)" |
|
2828 by auto |
|
2829 from \<beta>[OF lp u d dp nb2 px] show "?P (x -d )" . |
|
2830 qed |
|
2831 |
|
2832 lemma \<beta>_int: assumes lp: "iszlfm p bs" |
|
2833 shows "\<forall> b\<in> set (\<beta> p). isint b bs" |
|
2834 using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub) |
|
2835 |
|
2836 lemma cpmi_eq: "0 < D \<Longrightarrow> (EX z::int. ALL x. x < z --> (P x = P1 x)) |
|
2837 ==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) |
|
2838 ==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D)))) |
|
2839 ==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))" |
|
2840 apply(rule iffI) |
|
2841 prefer 2 |
|
2842 apply(drule minusinfinity) |
|
2843 apply assumption+ |
|
2844 apply(fastsimp) |
|
2845 apply clarsimp |
|
2846 apply(subgoal_tac "!!k. 0<=k \<Longrightarrow> !x. P x \<longrightarrow> P (x - k*D)") |
|
2847 apply(frule_tac x = x and z=z in decr_lemma) |
|
2848 apply(subgoal_tac "P1(x - (\<bar>x - z\<bar> + 1) * D)") |
|
2849 prefer 2 |
|
2850 apply(subgoal_tac "0 <= (\<bar>x - z\<bar> + 1)") |
|
2851 prefer 2 apply arith |
|
2852 apply fastsimp |
|
2853 apply(drule (1) periodic_finite_ex) |
|
2854 apply blast |
|
2855 apply(blast dest:decr_mult_lemma) |
|
2856 done |
|
2857 |
|
2858 |
|
2859 theorem cp_thm: |
|
2860 assumes lp: "iszlfm p (a #bs)" |
|
2861 and u: "d\<beta> p 1" |
|
2862 and d: "d\<delta> p d" |
|
2863 and dp: "d > 0" |
|
2864 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))" |
|
2865 (is "(\<exists> (x::int). ?P (real x)) = (\<exists> j\<in> ?D. ?M j \<or> (\<exists> b\<in> ?B. ?P (?I b + real j)))") |
|
2866 proof- |
|
2867 from minusinf_inf[OF lp] |
|
2868 have th: "\<exists>(z::int). \<forall>x<z. ?P (real x) = ?M x" by blast |
|
2869 let ?B' = "{floor (?I b) | b. b\<in> ?B}" |
|
2870 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 |
|
2871 from B[rule_format] |
|
2872 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))" |
|
2873 by simp |
|
2874 also have "\<dots> = (\<exists>j\<in>?D. \<exists>b\<in> ?B. ?P (real (floor (?I b) + j)))" by simp |
|
2875 also have"\<dots> = (\<exists> j \<in> ?D. \<exists> b \<in> ?B'. ?P (real (b + j)))" by blast |
|
2876 finally have BB': |
|
2877 "(\<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)))" |
|
2878 by blast |
|
2879 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 |
|
2880 from minusinf_repeats[OF d lp] |
|
2881 have th3: "\<forall> x k. ?M x = ?M (x-k*d)" by simp |
|
2882 from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast |
|
2883 qed |
|
2884 |
|
2885 (* Reddy and Loveland *) |
|
2886 |
|
2887 |
|
2888 consts |
|
2889 \<rho> :: "fm \<Rightarrow> (num \<times> int) list" (* Compute the Reddy and Loveland Bset*) |
|
2890 \<sigma>\<rho>:: "fm \<Rightarrow> num \<times> int \<Rightarrow> fm" (* Performs the modified substitution of Reddy and Loveland*) |
|
2891 \<alpha>\<rho> :: "fm \<Rightarrow> (num\<times>int) list" |
|
2892 a\<rho> :: "fm \<Rightarrow> int \<Rightarrow> fm" |
|
2893 recdef \<rho> "measure size" |
|
2894 "\<rho> (And p q) = (\<rho> p @ \<rho> q)" |
|
2895 "\<rho> (Or p q) = (\<rho> p @ \<rho> q)" |
|
2896 "\<rho> (Eq (CN 0 c e)) = [(Sub (C -1) e,c)]" |
|
2897 "\<rho> (NEq (CN 0 c e)) = [(Neg e,c)]" |
|
2898 "\<rho> (Lt (CN 0 c e)) = []" |
|
2899 "\<rho> (Le (CN 0 c e)) = []" |
|
2900 "\<rho> (Gt (CN 0 c e)) = [(Neg e, c)]" |
|
2901 "\<rho> (Ge (CN 0 c e)) = [(Sub (C (-1)) e, c)]" |
|
2902 "\<rho> p = []" |
|
2903 |
|
2904 recdef \<sigma>\<rho> "measure size" |
|
2905 "\<sigma>\<rho> (And p q) = (\<lambda> (t,k). And (\<sigma>\<rho> p (t,k)) (\<sigma>\<rho> q (t,k)))" |
|
2906 "\<sigma>\<rho> (Or p q) = (\<lambda> (t,k). Or (\<sigma>\<rho> p (t,k)) (\<sigma>\<rho> q (t,k)))" |
|
2907 "\<sigma>\<rho> (Eq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) |
|
2908 else (Eq (Add (Mul c t) (Mul k e))))" |
|
2909 "\<sigma>\<rho> (NEq (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) |
|
2910 else (NEq (Add (Mul c t) (Mul k e))))" |
|
2911 "\<sigma>\<rho> (Lt (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) |
|
2912 else (Lt (Add (Mul c t) (Mul k e))))" |
|
2913 "\<sigma>\<rho> (Le (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) |
|
2914 else (Le (Add (Mul c t) (Mul k e))))" |
|
2915 "\<sigma>\<rho> (Gt (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) |
|
2916 else (Gt (Add (Mul c t) (Mul k e))))" |
|
2917 "\<sigma>\<rho> (Ge (CN 0 c e)) = (\<lambda> (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) |
|
2918 else (Ge (Add (Mul c t) (Mul k e))))" |
|
2919 "\<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)) |
|
2920 else (Dvd (i*k) (Add (Mul c t) (Mul k e))))" |
|
2921 "\<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)) |
|
2922 else (NDvd (i*k) (Add (Mul c t) (Mul k e))))" |
|
2923 "\<sigma>\<rho> p = (\<lambda> (t,k). p)" |
|
2924 |
|
2925 recdef \<alpha>\<rho> "measure size" |
|
2926 "\<alpha>\<rho> (And p q) = (\<alpha>\<rho> p @ \<alpha>\<rho> q)" |
|
2927 "\<alpha>\<rho> (Or p q) = (\<alpha>\<rho> p @ \<alpha>\<rho> q)" |
|
2928 "\<alpha>\<rho> (Eq (CN 0 c e)) = [(Add (C -1) e,c)]" |
|
2929 "\<alpha>\<rho> (NEq (CN 0 c e)) = [(e,c)]" |
|
2930 "\<alpha>\<rho> (Lt (CN 0 c e)) = [(e,c)]" |
|
2931 "\<alpha>\<rho> (Le (CN 0 c e)) = [(Add (C -1) e,c)]" |
|
2932 "\<alpha>\<rho> p = []" |
|
2933 |
|
2934 (* Simulates normal substituion by modifying the formula see correctness theorem *) |
|
2935 |
|
2936 recdef a\<rho> "measure size" |
|
2937 "a\<rho> (And p q) = (\<lambda> k. And (a\<rho> p k) (a\<rho> q k))" |
|
2938 "a\<rho> (Or p q) = (\<lambda> k. Or (a\<rho> p k) (a\<rho> q k))" |
|
2939 "a\<rho> (Eq (CN 0 c e)) = (\<lambda> k. if k dvd c then (Eq (CN 0 (c div k) e)) |
|
2940 else (Eq (CN 0 c (Mul k e))))" |
|
2941 "a\<rho> (NEq (CN 0 c e)) = (\<lambda> k. if k dvd c then (NEq (CN 0 (c div k) e)) |
|
2942 else (NEq (CN 0 c (Mul k e))))" |
|
2943 "a\<rho> (Lt (CN 0 c e)) = (\<lambda> k. if k dvd c then (Lt (CN 0 (c div k) e)) |
|
2944 else (Lt (CN 0 c (Mul k e))))" |
|
2945 "a\<rho> (Le (CN 0 c e)) = (\<lambda> k. if k dvd c then (Le (CN 0 (c div k) e)) |
|
2946 else (Le (CN 0 c (Mul k e))))" |
|
2947 "a\<rho> (Gt (CN 0 c e)) = (\<lambda> k. if k dvd c then (Gt (CN 0 (c div k) e)) |
|
2948 else (Gt (CN 0 c (Mul k e))))" |
|
2949 "a\<rho> (Ge (CN 0 c e)) = (\<lambda> k. if k dvd c then (Ge (CN 0 (c div k) e)) |
|
2950 else (Ge (CN 0 c (Mul k e))))" |
|
2951 "a\<rho> (Dvd i (CN 0 c e)) = (\<lambda> k. if k dvd c then (Dvd i (CN 0 (c div k) e)) |
|
2952 else (Dvd (i*k) (CN 0 c (Mul k e))))" |
|
2953 "a\<rho> (NDvd i (CN 0 c e)) = (\<lambda> k. if k dvd c then (NDvd i (CN 0 (c div k) e)) |
|
2954 else (NDvd (i*k) (CN 0 c (Mul k e))))" |
|
2955 "a\<rho> p = (\<lambda> k. p)" |
|
2956 |
|
2957 constdefs \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" |
|
2958 "\<sigma> p k t \<equiv> And (Dvd k t) (\<sigma>\<rho> p (t,k))" |
|
2959 |
|
2960 lemma \<sigma>\<rho>: |
|
2961 assumes linp: "iszlfm p (real (x::int)#bs)" |
|
2962 and kpos: "real k > 0" |
|
2963 and tnb: "numbound0 t" |
|
2964 and tint: "isint t (real x#bs)" |
|
2965 and kdt: "k dvd floor (Inum (b'#bs) t)" |
|
2966 shows "Ifm (real x#bs) (\<sigma>\<rho> p (t,k)) = |
|
2967 (Ifm ((real ((floor (Inum (b'#bs) t)) div k))#bs) p)" |
|
2968 (is "?I (real x) (?s p) = (?I (real ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)") |
|
2969 using linp kpos tnb |
|
2970 proof(induct p rule: \<sigma>\<rho>.induct) |
|
2971 case (3 c e) |
|
2972 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
2973 {assume kdc: "k dvd c" |
|
2974 from kpos have knz: "k\<noteq>0" by simp |
|
2975 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
2976 from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] |
|
2977 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
2978 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } |
|
2979 moreover |
|
2980 {assume "\<not> k dvd c" |
|
2981 from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
2982 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
2983 from prems 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)" |
|
2984 using real_of_int_div[OF knz kdt] |
|
2985 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
2986 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps) |
|
2987 also have "\<dots> = (?I ?tk (Eq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
2988 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] |
|
2989 by (simp add: ti) |
|
2990 finally have ?case . } |
|
2991 ultimately show ?case by blast |
|
2992 next |
|
2993 case (4 c e) |
|
2994 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
2995 {assume kdc: "k dvd c" |
|
2996 from kpos have knz: "k\<noteq>0" by simp |
|
2997 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
2998 from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] |
|
2999 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3000 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } |
|
3001 moreover |
|
3002 {assume "\<not> k dvd c" |
|
3003 from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3004 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3005 from prems 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)" |
|
3006 using real_of_int_div[OF knz kdt] |
|
3007 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3008 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps) |
|
3009 also have "\<dots> = (?I ?tk (NEq (CN 0 c e)))" using nonzero_eq_divide_eq[OF knz', where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3010 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] |
|
3011 by (simp add: ti) |
|
3012 finally have ?case . } |
|
3013 ultimately show ?case by blast |
|
3014 next |
|
3015 case (5 c e) |
|
3016 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3017 {assume kdc: "k dvd c" |
|
3018 from kpos have knz: "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 prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] |
|
3021 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3022 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } |
|
3023 moreover |
|
3024 {assume "\<not> k dvd c" |
|
3025 from kpos have knz: "k\<noteq>0" by simp |
|
3026 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3027 from prems 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)" |
|
3028 using real_of_int_div[OF knz kdt] |
|
3029 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3030 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps) |
|
3031 also have "\<dots> = (?I ?tk (Lt (CN 0 c e)))" using pos_less_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3032 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] |
|
3033 by (simp add: ti) |
|
3034 finally have ?case . } |
|
3035 ultimately show ?case by blast |
|
3036 next |
|
3037 case (6 c e) |
|
3038 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3039 {assume kdc: "k dvd c" |
|
3040 from kpos have knz: "k\<noteq>0" by simp |
|
3041 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3042 from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] |
|
3043 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3044 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } |
|
3045 moreover |
|
3046 {assume "\<not> k dvd c" |
|
3047 from kpos have knz: "k\<noteq>0" by simp |
|
3048 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3049 from prems 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)" |
|
3050 using real_of_int_div[OF knz kdt] |
|
3051 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3052 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps) |
|
3053 also have "\<dots> = (?I ?tk (Le (CN 0 c e)))" using pos_le_divide_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3054 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] |
|
3055 by (simp add: ti) |
|
3056 finally have ?case . } |
|
3057 ultimately show ?case by blast |
|
3058 next |
|
3059 case (7 c e) |
|
3060 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3061 {assume kdc: "k dvd c" |
|
3062 from kpos have knz: "k\<noteq>0" by simp |
|
3063 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3064 from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] |
|
3065 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3066 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } |
|
3067 moreover |
|
3068 {assume "\<not> k dvd c" |
|
3069 from kpos have knz: "k\<noteq>0" by simp |
|
3070 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3071 from prems 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)" |
|
3072 using real_of_int_div[OF knz kdt] |
|
3073 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3074 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps) |
|
3075 also have "\<dots> = (?I ?tk (Gt (CN 0 c e)))" using pos_divide_less_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3076 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] |
|
3077 by (simp add: ti) |
|
3078 finally have ?case . } |
|
3079 ultimately show ?case by blast |
|
3080 next |
|
3081 case (8 c e) |
|
3082 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3083 {assume kdc: "k dvd c" |
|
3084 from kpos have knz: "k\<noteq>0" by simp |
|
3085 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3086 from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] |
|
3087 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3088 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } |
|
3089 moreover |
|
3090 {assume "\<not> k dvd c" |
|
3091 from kpos have knz: "k\<noteq>0" by simp |
|
3092 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3093 from prems 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)" |
|
3094 using real_of_int_div[OF knz kdt] |
|
3095 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3096 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps) |
|
3097 also have "\<dots> = (?I ?tk (Ge (CN 0 c e)))" using pos_divide_le_eq[OF kpos, where a="real c * (?N (real x) t / real k) + ?N (real x) e" and b="0", symmetric] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3098 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] |
|
3099 by (simp add: ti) |
|
3100 finally have ?case . } |
|
3101 ultimately show ?case by blast |
|
3102 next |
|
3103 case (9 i c e) from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3104 {assume kdc: "k dvd c" |
|
3105 from kpos have knz: "k\<noteq>0" by simp |
|
3106 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3107 from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] |
|
3108 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3109 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } |
|
3110 moreover |
|
3111 {assume "\<not> k dvd c" |
|
3112 from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3113 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3114 from prems 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)" |
|
3115 using real_of_int_div[OF knz kdt] |
|
3116 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3117 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps) |
|
3118 also have "\<dots> = (?I ?tk (Dvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3119 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] |
|
3120 by (simp add: ti) |
|
3121 finally have ?case . } |
|
3122 ultimately show ?case by blast |
|
3123 next |
|
3124 case (10 i c e) from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3125 {assume kdc: "k dvd c" |
|
3126 from kpos have knz: "k\<noteq>0" by simp |
|
3127 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3128 from prems have ?case using real_of_int_div[OF knz kdc] real_of_int_div[OF knz kdt] |
|
3129 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3130 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti) } |
|
3131 moreover |
|
3132 {assume "\<not> k dvd c" |
|
3133 from kpos have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3134 from tint have ti: "real (floor (?N (real x) t)) = ?N (real x) t" using isint_def by simp |
|
3135 from prems 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))" |
|
3136 using real_of_int_div[OF knz kdt] |
|
3137 numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3138 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] by (simp add: ti algebra_simps) |
|
3139 also have "\<dots> = (?I ?tk (NDvd i (CN 0 c e)))" using rdvd_mult[OF knz, where n="i"] real_of_int_div[OF knz kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real x"] |
|
3140 numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real x"] |
|
3141 by (simp add: ti) |
|
3142 finally have ?case . } |
|
3143 ultimately show ?case by blast |
|
3144 qed (simp_all add: nth_pos2 bound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"] numbound0_I[where bs="bs" and b="real ((floor (?N b' t)) div k)" and b'="real x"]) |
|
3145 |
|
3146 |
|
3147 lemma a\<rho>: |
|
3148 assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "real k > 0" |
|
3149 shows "Ifm (real (x*k)#bs) (a\<rho> p k) = Ifm (real x#bs) p" (is "?I (x*k) (?f p k) = ?I x p") |
|
3150 using lp bound0_I[where bs="bs" and b="real (x*k)" and b'="real x"] numbound0_I[where bs="bs" and b="real (x*k)" and b'="real x"] |
|
3151 proof(induct p rule: a\<rho>.induct) |
|
3152 case (3 c e) |
|
3153 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3154 from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3155 {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } |
|
3156 moreover |
|
3157 {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)} |
|
3158 ultimately show ?case by blast |
|
3159 next |
|
3160 case (4 c e) |
|
3161 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3162 from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3163 {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } |
|
3164 moreover |
|
3165 {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] nonzero_eq_divide_eq[OF knz', where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)} |
|
3166 ultimately show ?case by blast |
|
3167 next |
|
3168 case (5 c e) |
|
3169 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3170 from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3171 {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } |
|
3172 moreover |
|
3173 {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_less_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)} |
|
3174 ultimately show ?case by blast |
|
3175 next |
|
3176 case (6 c e) |
|
3177 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3178 from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3179 {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } |
|
3180 moreover |
|
3181 {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_le_divide_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)} |
|
3182 ultimately show ?case by blast |
|
3183 next |
|
3184 case (7 c e) |
|
3185 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3186 from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3187 {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } |
|
3188 moreover |
|
3189 {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_less_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)} |
|
3190 ultimately show ?case by blast |
|
3191 next |
|
3192 case (8 c e) |
|
3193 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3194 from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3195 {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } |
|
3196 moreover |
|
3197 {assume nkdc: "\<not> k dvd c" hence ?case using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] pos_divide_le_eq[OF kp, where b="0" and a="real c * real x + Inum (real x # bs) e", symmetric] by (simp add: algebra_simps)} |
|
3198 ultimately show ?case by blast |
|
3199 next |
|
3200 case (9 i c e) |
|
3201 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3202 from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3203 {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } |
|
3204 moreover |
|
3205 {assume "\<not> k dvd c" |
|
3206 hence "Ifm (real (x*k)#bs) (a\<rho> (Dvd i (CN 0 c e)) k) = |
|
3207 (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k)" |
|
3208 using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] |
|
3209 by (simp add: algebra_simps) |
|
3210 also have "\<dots> = (Ifm (real x#bs) (Dvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"]) |
|
3211 finally have ?case . } |
|
3212 ultimately show ?case by blast |
|
3213 next |
|
3214 case (10 i c e) |
|
3215 from prems have cp: "c > 0" and nb: "numbound0 e" by auto |
|
3216 from kp have knz: "k\<noteq>0" by simp hence knz': "real k \<noteq> 0" by simp |
|
3217 {assume kdc: "k dvd c" from prems have ?case using real_of_int_div[OF knz kdc] by simp } |
|
3218 moreover |
|
3219 {assume "\<not> k dvd c" |
|
3220 hence "Ifm (real (x*k)#bs) (a\<rho> (NDvd i (CN 0 c e)) k) = |
|
3221 (\<not> (real i * real k rdvd (real c * real x + Inum (real x#bs) e) * real k))" |
|
3222 using numbound0_I[OF nb, where bs="bs" and b="real (x*k)" and b'="real x"] |
|
3223 by (simp add: algebra_simps) |
|
3224 also have "\<dots> = (Ifm (real x#bs) (NDvd i (CN 0 c e)))" by (simp add: rdvd_mult[OF knz, where n="i"]) |
|
3225 finally have ?case . } |
|
3226 ultimately show ?case by blast |
|
3227 qed (simp_all add: nth_pos2) |
|
3228 |
|
3229 lemma a\<rho>_ex: |
|
3230 assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0" |
|
3231 shows "(\<exists> (x::int). real k rdvd real x \<and> Ifm (real x#bs) (a\<rho> p k)) = |
|
3232 (\<exists> (x::int). Ifm (real x#bs) p)" (is "(\<exists> x. ?D x \<and> ?P' x) = (\<exists> x. ?P x)") |
|
3233 proof- |
|
3234 have "(\<exists> x. ?D x \<and> ?P' x) = (\<exists> x. k dvd x \<and> ?P' x)" using int_rdvd_iff by simp |
|
3235 also have "\<dots> = (\<exists>x. ?P' (x*k))" using unity_coeff_ex[where P="?P'" and l="k", simplified] |
|
3236 by (simp add: algebra_simps) |
|
3237 also have "\<dots> = (\<exists> x. ?P x)" using a\<rho> iszlfm_gen[OF lp] kp by auto |
|
3238 finally show ?thesis . |
|
3239 qed |
|
3240 |
|
3241 lemma \<sigma>\<rho>': assumes lp: "iszlfm p (real (x::int)#bs)" and kp: "k > 0" and nb: "numbound0 t" |
|
3242 shows "Ifm (real x#bs) (\<sigma>\<rho> p (t,k)) = Ifm ((Inum (real x#bs) t)#bs) (a\<rho> p k)" |
|
3243 using lp |
|
3244 by(induct p rule: \<sigma>\<rho>.induct, simp_all add: |
|
3245 numbound0_I[OF nb, where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] |
|
3246 numbound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] |
|
3247 bound0_I[where bs="bs" and b="Inum (real x#bs) t" and b'="real x"] nth_pos2 cong: imp_cong) |
|
3248 |
|
3249 lemma \<sigma>\<rho>_nb: assumes lp:"iszlfm p (a#bs)" and nb: "numbound0 t" |
|
3250 shows "bound0 (\<sigma>\<rho> p (t,k))" |
|
3251 using lp |
|
3252 by (induct p rule: iszlfm.induct, auto simp add: nb) |
|
3253 |
|
3254 lemma \<rho>_l: |
|
3255 assumes lp: "iszlfm p (real (i::int)#bs)" |
|
3256 shows "\<forall> (b,k) \<in> set (\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)" |
|
3257 using lp by (induct p rule: \<rho>.induct, auto simp add: isint_sub isint_neg) |
|
3258 |
|
3259 lemma \<alpha>\<rho>_l: |
|
3260 assumes lp: "iszlfm p (real (i::int)#bs)" |
|
3261 shows "\<forall> (b,k) \<in> set (\<alpha>\<rho> p). k >0 \<and> numbound0 b \<and> isint b (real i#bs)" |
|
3262 using lp isint_add [OF isint_c[where j="- 1"],where bs="real i#bs"] |
|
3263 by (induct p rule: \<alpha>\<rho>.induct, auto) |
|
3264 |
|
3265 lemma zminusinf_\<rho>: |
|
3266 assumes lp: "iszlfm p (real (i::int)#bs)" |
|
3267 and nmi: "\<not> (Ifm (real i#bs) (minusinf p))" (is "\<not> (Ifm (real i#bs) (?M p))") |
|
3268 and ex: "Ifm (real i#bs) p" (is "?I i p") |
|
3269 shows "\<exists> (e,c) \<in> set (\<rho> p). real (c*i) > Inum (real i#bs) e" (is "\<exists> (e,c) \<in> ?R p. real (c*i) > ?N i e") |
|
3270 using lp nmi ex |
|
3271 by (induct p rule: minusinf.induct, auto) |
|
3272 |
|
3273 |
|
3274 lemma \<sigma>_And: "Ifm bs (\<sigma> (And p q) k t) = Ifm bs (And (\<sigma> p k t) (\<sigma> q k t))" |
|
3275 using \<sigma>_def by auto |
|
3276 lemma \<sigma>_Or: "Ifm bs (\<sigma> (Or p q) k t) = Ifm bs (Or (\<sigma> p k t) (\<sigma> q k t))" |
|
3277 using \<sigma>_def by auto |
|
3278 |
|
3279 lemma \<rho>: assumes lp: "iszlfm p (real (i::int) #bs)" |
|
3280 and pi: "Ifm (real i#bs) p" |
|
3281 and d: "d\<delta> p d" |
|
3282 and dp: "d > 0" |
|
3283 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" |
|
3284 (is "\<forall>(e,c) \<in> set (\<rho> p). \<forall> j\<in> {1 .. c*d}. _ \<noteq> ?N i e + _") |
|
3285 shows "Ifm (real(i - d)#bs) p" |
|
3286 using lp pi d nob |
|
3287 proof(induct p rule: iszlfm.induct) |
|
3288 case (3 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" |
|
3289 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" |
|
3290 by simp+ |
|
3291 from mult_strict_left_mono[OF dp cp] have one:"1 \<in> {1 .. c*d}" by auto |
|
3292 from nob[rule_format, where j="1", OF one] pi show ?case by simp |
|
3293 next |
|
3294 case (4 c e) |
|
3295 hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" |
|
3296 and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j" |
|
3297 by simp+ |
|
3298 {assume "real (c*i) \<noteq> - ?N i e + real (c*d)" |
|
3299 with numbound0_I[OF nb, where bs="bs" and b="real i - real d" and b'="real i"] |
|
3300 have ?case by (simp add: algebra_simps)} |
|
3301 moreover |
|
3302 {assume pi: "real (c*i) = - ?N i e + real (c*d)" |
|
3303 from mult_strict_left_mono[OF dp cp] have d: "(c*d) \<in> {1 .. c*d}" by simp |
|
3304 from nob[rule_format, where j="c*d", OF d] pi have ?case by simp } |
|
3305 ultimately show ?case by blast |
|
3306 next |
|
3307 case (5 c e) hence cp: "c > 0" by simp |
|
3308 from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] |
|
3309 real_of_int_mult] |
|
3310 show ?case using prems dp |
|
3311 by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] |
|
3312 algebra_simps) |
|
3313 next |
|
3314 case (6 c e) hence cp: "c > 0" by simp |
|
3315 from prems mult_strict_left_mono[OF dp cp, simplified real_of_int_less_iff[symmetric] |
|
3316 real_of_int_mult] |
|
3317 show ?case using prems dp |
|
3318 by (simp add: add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] |
|
3319 algebra_simps) |
|
3320 next |
|
3321 case (7 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" |
|
3322 and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - ?N i e + real j" |
|
3323 and pi: "real (c*i) + ?N i e > 0" and cp': "real c >0" |
|
3324 by simp+ |
|
3325 let ?fe = "floor (?N i e)" |
|
3326 from pi cp have th:"(real i +?N i e / real c)*real c > 0" by (simp add: algebra_simps) |
|
3327 from pi ei[simplified isint_iff] have "real (c*i + ?fe) > real (0::int)" by simp |
|
3328 hence pi': "c*i + ?fe > 0" by (simp only: real_of_int_less_iff[symmetric]) |
|
3329 have "real (c*i) + ?N i e > real (c*d) \<or> real (c*i) + ?N i e \<le> real (c*d)" by auto |
|
3330 moreover |
|
3331 {assume "real (c*i) + ?N i e > real (c*d)" hence ?case |
|
3332 by (simp add: algebra_simps |
|
3333 numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} |
|
3334 moreover |
|
3335 {assume H:"real (c*i) + ?N i e \<le> real (c*d)" |
|
3336 with ei[simplified isint_iff] have "real (c*i + ?fe) \<le> real (c*d)" by simp |
|
3337 hence pid: "c*i + ?fe \<le> c*d" by (simp only: real_of_int_le_iff) |
|
3338 with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + ?fe = j1" by auto |
|
3339 hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - ?N i e + real j1" |
|
3340 by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps) |
|
3341 with nob have ?case by blast } |
|
3342 ultimately show ?case by blast |
|
3343 next |
|
3344 case (8 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real i#bs)" |
|
3345 and nob: "\<forall> j\<in> {1 .. c*d}. real (c*i) \<noteq> - 1 - ?N i e + real j" |
|
3346 and pi: "real (c*i) + ?N i e \<ge> 0" and cp': "real c >0" |
|
3347 by simp+ |
|
3348 let ?fe = "floor (?N i e)" |
|
3349 from pi cp have th:"(real i +?N i e / real c)*real c \<ge> 0" by (simp add: algebra_simps) |
|
3350 from pi ei[simplified isint_iff] have "real (c*i + ?fe) \<ge> real (0::int)" by simp |
|
3351 hence pi': "c*i + 1 + ?fe \<ge> 1" by (simp only: real_of_int_le_iff[symmetric]) |
|
3352 have "real (c*i) + ?N i e \<ge> real (c*d) \<or> real (c*i) + ?N i e < real (c*d)" by auto |
|
3353 moreover |
|
3354 {assume "real (c*i) + ?N i e \<ge> real (c*d)" hence ?case |
|
3355 by (simp add: algebra_simps |
|
3356 numbound0_I[OF nb,where bs="bs" and b="real i - real d" and b'="real i"])} |
|
3357 moreover |
|
3358 {assume H:"real (c*i) + ?N i e < real (c*d)" |
|
3359 with ei[simplified isint_iff] have "real (c*i + ?fe) < real (c*d)" by simp |
|
3360 hence pid: "c*i + 1 + ?fe \<le> c*d" by (simp only: real_of_int_le_iff) |
|
3361 with pi' have "\<exists> j1\<in> {1 .. c*d}. c*i + 1+ ?fe = j1" by auto |
|
3362 hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) + 1= - ?N i e + real j1" |
|
3363 by (simp only: diff_def[symmetric] real_of_int_mult real_of_int_add real_of_int_inject[symmetric] ei[simplified isint_iff] algebra_simps real_of_one) |
|
3364 hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = (- ?N i e + real j1) - 1" |
|
3365 by (simp only: algebra_simps diff_def[symmetric]) |
|
3366 hence "\<exists> j1\<in> {1 .. c*d}. real (c*i) = - 1 - ?N i e + real j1" |
|
3367 by (simp only: add_ac diff_def) |
|
3368 with nob have ?case by blast } |
|
3369 ultimately show ?case by blast |
|
3370 next |
|
3371 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+ |
|
3372 let ?e = "Inum (real i # bs) e" |
|
3373 from prems have "isint e (real i #bs)" by simp |
|
3374 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"] |
|
3375 by (simp add: isint_iff) |
|
3376 from prems have id: "j dvd d" by simp |
|
3377 from ie[symmetric] have "?p i = (real j rdvd real (c*i+ floor ?e))" by simp |
|
3378 also have "\<dots> = (j dvd c*i + floor ?e)" |
|
3379 using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp |
|
3380 also have "\<dots> = (j dvd c*i - c*d + floor ?e)" |
|
3381 using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp |
|
3382 also have "\<dots> = (real j rdvd real (c*i - c*d + floor ?e))" |
|
3383 using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] |
|
3384 ie by simp |
|
3385 also have "\<dots> = (real j rdvd real (c*(i - d)) + ?e)" |
|
3386 using ie by (simp add:algebra_simps) |
|
3387 finally show ?case |
|
3388 using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p |
|
3389 by (simp add: algebra_simps) |
|
3390 next |
|
3391 case (10 j c e) hence p: "\<not> (real j rdvd real (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ |
|
3392 let ?e = "Inum (real i # bs) e" |
|
3393 from prems have "isint e (real i #bs)" by simp |
|
3394 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"] |
|
3395 by (simp add: isint_iff) |
|
3396 from prems have id: "j dvd d" by simp |
|
3397 from ie[symmetric] have "?p i = (\<not> (real j rdvd real (c*i+ floor ?e)))" by simp |
|
3398 also have "\<dots> = Not (j dvd c*i + floor ?e)" |
|
3399 using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp |
|
3400 also have "\<dots> = Not (j dvd c*i - c*d + floor ?e)" |
|
3401 using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp |
|
3402 also have "\<dots> = Not (real j rdvd real (c*i - c*d + floor ?e))" |
|
3403 using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] |
|
3404 ie by simp |
|
3405 also have "\<dots> = Not (real j rdvd real (c*(i - d)) + ?e)" |
|
3406 using ie by (simp add:algebra_simps) |
|
3407 finally show ?case |
|
3408 using numbound0_I[OF bn,where b="real i - real d" and b'="real i" and bs="bs"] p |
|
3409 by (simp add: algebra_simps) |
|
3410 qed(auto simp add: numbound0_I[where bs="bs" and b="real i - real d" and b'="real i"] nth_pos2) |
|
3411 |
|
3412 lemma \<sigma>_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t" |
|
3413 shows "bound0 (\<sigma> p k t)" |
|
3414 using \<sigma>\<rho>_nb[OF lp nb] nb by (simp add: \<sigma>_def) |
|
3415 |
|
3416 lemma \<rho>': assumes lp: "iszlfm p (a #bs)" |
|
3417 and d: "d\<delta> p d" |
|
3418 and dp: "d > 0" |
|
3419 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)") |
|
3420 proof(clarify) |
|
3421 fix x |
|
3422 assume nob1:"?b x" and px: "?P x" |
|
3423 from iszlfm_gen[OF lp, rule_format, where y="real x"] have lp': "iszlfm p (real x#bs)". |
|
3424 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" |
|
3425 proof(clarify) |
|
3426 fix e c j assume ecR: "(e,c) \<in> set (\<rho> p)" and jD: "j\<in> {1 .. c*d}" |
|
3427 and cx: "real (c*x) = Inum (real x#bs) e + real j" |
|
3428 let ?e = "Inum (real x#bs) e" |
|
3429 let ?fe = "floor ?e" |
|
3430 from \<rho>_l[OF lp'] ecR have ei:"isint e (real x#bs)" and cp:"c>0" and nb:"numbound0 e" |
|
3431 by auto |
|
3432 from numbound0_gen [OF nb ei, rule_format,where y="a"] have "isint e (a#bs)" . |
|
3433 from cx ei[simplified isint_iff] have "real (c*x) = real (?fe + j)" by simp |
|
3434 hence cx: "c*x = ?fe + j" by (simp only: real_of_int_inject) |
|
3435 hence cdej:"c dvd ?fe + j" by (simp add: dvd_def) (rule_tac x="x" in exI, simp) |
|
3436 hence "real c rdvd real (?fe + j)" by (simp only: int_rdvd_iff) |
|
3437 hence rcdej: "real c rdvd ?e + real j" by (simp add: ei[simplified isint_iff]) |
|
3438 from cx have "(c*x) div c = (?fe + j) div c" by simp |
|
3439 with cp have "x = (?fe + j) div c" by simp |
|
3440 with px have th: "?P ((?fe + j) div c)" by auto |
|
3441 from cp have cp': "real c > 0" by simp |
|
3442 from cdej have cdej': "c dvd floor (Inum (real x#bs) (Add e (C j)))" by simp |
|
3443 from nb have nb': "numbound0 (Add e (C j))" by simp |
|
3444 have ji: "isint (C j) (real x#bs)" by (simp add: isint_def) |
|
3445 from isint_add[OF ei ji] have ei':"isint (Add e (C j)) (real x#bs)" . |
|
3446 from th \<sigma>\<rho>[where b'="real x", OF lp' cp' nb' ei' cdej',symmetric] |
|
3447 have "Ifm (real x#bs) (\<sigma>\<rho> p (Add e (C j), c))" by simp |
|
3448 with rcdej have th: "Ifm (real x#bs) (\<sigma> p c (Add e (C j)))" by (simp add: \<sigma>_def) |
|
3449 from th bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"],where bs="bs" and b="real x" and b'="a"] |
|
3450 have "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" by blast |
|
3451 with ecR jD nob1 show "False" by blast |
|
3452 qed |
|
3453 from \<rho>[OF lp' px d dp nob] show "?P (x -d )" . |
|
3454 qed |
|
3455 |
|
3456 |
|
3457 lemma rl_thm: |
|
3458 assumes lp: "iszlfm p (real (i::int)#bs)" |
|
3459 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)))))" |
|
3460 (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))" |
|
3461 is "?lhs = (?MD \<or> ?RD)" is "?lhs = ?rhs") |
|
3462 proof- |
|
3463 let ?d= "\<delta> p" |
|
3464 from \<delta>[OF lp] have d:"d\<delta> p ?d" and dp: "?d > 0" by auto |
|
3465 { assume H:"?MD" hence th:"\<exists> (x::int). ?MP x" by blast |
|
3466 from H minusinf_ex[OF lp th] have ?thesis by blast} |
|
3467 moreover |
|
3468 { fix e c j assume exR:"(e,c) \<in> ?R" and jD:"j\<in> {1 .. c*?d}" and spx:"?SP c e j" |
|
3469 from exR \<rho>_l[OF lp] have nb: "numbound0 e" and ei:"isint e (real i#bs)" and cp: "c > 0" |
|
3470 by auto |
|
3471 have "isint (C j) (real i#bs)" by (simp add: isint_iff) |
|
3472 with isint_add[OF numbound0_gen[OF nb ei,rule_format, where y="real i"]] |
|
3473 have eji:"isint (Add e (C j)) (real i#bs)" by simp |
|
3474 from nb have nb': "numbound0 (Add e (C j))" by simp |
|
3475 from spx bound0_I[OF \<sigma>_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real i"] |
|
3476 have spx': "Ifm (real i # bs) (\<sigma> p c (Add e (C j)))" by blast |
|
3477 from spx' have rcdej:"real c rdvd (Inum (real i#bs) (Add e (C j)))" |
|
3478 and sr:"Ifm (real i#bs) (\<sigma>\<rho> p (Add e (C j),c))" by (simp add: \<sigma>_def)+ |
|
3479 from rcdej eji[simplified isint_iff] |
|
3480 have "real c rdvd real (floor (Inum (real i#bs) (Add e (C j))))" by simp |
|
3481 hence cdej:"c dvd floor (Inum (real i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff) |
|
3482 from cp have cp': "real c > 0" by simp |
|
3483 from \<sigma>\<rho>[OF lp cp' nb' eji cdej] spx' have "?P (\<lfloor>Inum (real i # bs) (Add e (C j))\<rfloor> div c)" |
|
3484 by (simp add: \<sigma>_def) |
|
3485 hence ?lhs by blast |
|
3486 with exR jD spx have ?thesis by blast} |
|
3487 moreover |
|
3488 { fix x assume px: "?P x" and nob: "\<not> ?RD" |
|
3489 from iszlfm_gen [OF lp,rule_format, where y="a"] have lp':"iszlfm p (a#bs)" . |
|
3490 from \<rho>'[OF lp' d dp, rule_format, OF nob] have th:"\<forall> x. ?P x \<longrightarrow> ?P (x - ?d)" by blast |
|
3491 from minusinf_inf[OF lp] obtain z where z:"\<forall> x<z. ?MP x = ?P x" by blast |
|
3492 have zp: "abs (x - z) + 1 \<ge> 0" by arith |
|
3493 from decr_lemma[OF dp,where x="x" and z="z"] |
|
3494 decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\<exists> x. ?MP x" by auto |
|
3495 with minusinf_bex[OF lp] px nob have ?thesis by blast} |
|
3496 ultimately show ?thesis by blast |
|
3497 qed |
|
3498 |
|
3499 lemma mirror_\<alpha>\<rho>: assumes lp: "iszlfm p (a#bs)" |
|
3500 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))" |
|
3501 using lp |
|
3502 by (induct p rule: mirror.induct, simp_all add: split_def image_Un ) |
|
3503 |
|
3504 text {* The @{text "\<real>"} part*} |
|
3505 |
|
3506 text{* Linearity for fm where Bound 0 ranges over @{text "\<real>"}*} |
|
3507 consts |
|
3508 isrlfm :: "fm \<Rightarrow> bool" (* Linearity test for fm *) |
|
3509 recdef isrlfm "measure size" |
|
3510 "isrlfm (And p q) = (isrlfm p \<and> isrlfm q)" |
|
3511 "isrlfm (Or p q) = (isrlfm p \<and> isrlfm q)" |
|
3512 "isrlfm (Eq (CN 0 c e)) = (c>0 \<and> numbound0 e)" |
|
3513 "isrlfm (NEq (CN 0 c e)) = (c>0 \<and> numbound0 e)" |
|
3514 "isrlfm (Lt (CN 0 c e)) = (c>0 \<and> numbound0 e)" |
|
3515 "isrlfm (Le (CN 0 c e)) = (c>0 \<and> numbound0 e)" |
|
3516 "isrlfm (Gt (CN 0 c e)) = (c>0 \<and> numbound0 e)" |
|
3517 "isrlfm (Ge (CN 0 c e)) = (c>0 \<and> numbound0 e)" |
|
3518 "isrlfm p = (isatom p \<and> (bound0 p))" |
|
3519 |
|
3520 constdefs fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm" |
|
3521 "fp p n s j \<equiv> (if n > 0 then |
|
3522 (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j))))) |
|
3523 (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1)))))))) |
|
3524 else |
|
3525 (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) |
|
3526 (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))" |
|
3527 |
|
3528 (* splits the bounded from the unbounded part*) |
|
3529 consts rsplit0 :: "num \<Rightarrow> (fm \<times> int \<times> num) list" |
|
3530 recdef rsplit0 "measure num_size" |
|
3531 "rsplit0 (Bound 0) = [(T,1,C 0)]" |
|
3532 "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b |
|
3533 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])" |
|
3534 "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))" |
|
3535 "rsplit0 (Neg a) = map (\<lambda> (p,n,s). (p,-n,Neg s)) (rsplit0 a)" |
|
3536 "rsplit0 (Floor a) = foldl (op @) [] (map |
|
3537 (\<lambda> (p,n,s). if n=0 then [(p,0,Floor s)] |
|
3538 else (map (\<lambda> j. (fp p n s j, 0, Add (Floor s) (C j))) (if n > 0 then iupt (0,n) else iupt(n,0)))) |
|
3539 (rsplit0 a))" |
|
3540 "rsplit0 (CN 0 c a) = map (\<lambda> (p,n,s). (p,n+c,s)) (rsplit0 a)" |
|
3541 "rsplit0 (CN m c a) = map (\<lambda> (p,n,s). (p,n,CN m c s)) (rsplit0 a)" |
|
3542 "rsplit0 (CF c t s) = rsplit0 (Add (Mul c (Floor t)) s)" |
|
3543 "rsplit0 (Mul c a) = map (\<lambda> (p,n,s). (p,c*n,Mul c s)) (rsplit0 a)" |
|
3544 "rsplit0 t = [(T,0,t)]" |
|
3545 |
|
3546 lemma not_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm (not p)" |
|
3547 by (induct p rule: isrlfm.induct, auto) |
|
3548 lemma conj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (conj p q)" |
|
3549 using conj_def by (cases p, auto) |
|
3550 lemma disj_rl[simp]: "isrlfm p \<Longrightarrow> isrlfm q \<Longrightarrow> isrlfm (disj p q)" |
|
3551 using disj_def by (cases p, auto) |
|
3552 |
|
3553 |
|
3554 lemma rsplit0_cs: |
|
3555 shows "\<forall> (p,n,s) \<in> set (rsplit0 t). |
|
3556 (Ifm (x#bs) p \<longrightarrow> (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p" |
|
3557 (is "\<forall> (p,n,s) \<in> ?SS t. (?I p \<longrightarrow> ?N t = ?N (CN 0 n s)) \<and> _ \<and> _ ") |
|
3558 proof(induct t rule: rsplit0.induct) |
|
3559 case (5 a) |
|
3560 let ?p = "\<lambda> (p,n,s) j. fp p n s j" |
|
3561 let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))" |
|
3562 let ?J = "\<lambda> n. if n>0 then iupt (0,n) else iupt (n,0)" |
|
3563 let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" |
|
3564 have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith |
|
3565 have U1: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = |
|
3566 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)]))" by auto |
|
3567 have U2': "\<forall> (p,n,s) \<in> {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0}. |
|
3568 ?ff (p,n,s) = map (?f(p,n,s)) (iupt(0,n))" by auto |
|
3569 hence U2: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = |
|
3570 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). |
|
3571 set (map (?f(p,n,s)) (iupt(0,n)))))" |
|
3572 proof- |
|
3573 fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g |
|
3574 assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c" |
|
3575 thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))" |
|
3576 by (auto simp add: split_def) |
|
3577 qed |
|
3578 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)) (iupt(n,0))" |
|
3579 by auto |
|
3580 hence U3: "(UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) = |
|
3581 (UNION {(p,n,s). (p,n,s)\<in> ?SS a\<and>n<0} (\<lambda>(p,n,s). set (map (?f(p,n,s)) (iupt(n,0)))))" |
|
3582 proof- |
|
3583 fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g |
|
3584 assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c" |
|
3585 thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))" |
|
3586 by (auto simp add: split_def) |
|
3587 qed |
|
3588 have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))" |
|
3589 by (auto simp add: foldl_conv_concat) |
|
3590 also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by auto |
|
3591 also have "\<dots> = |
|
3592 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un |
|
3593 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un |
|
3594 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))" |
|
3595 using int_cases[rule_format] by blast |
|
3596 also have "\<dots> = |
|
3597 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un |
|
3598 (UNION {(p,n,s). (p,n,s)\<in> ?SS a\<and>n>0} (\<lambda>(p,n,s). set(map(?f(p,n,s)) (iupt(0,n))))) Un |
|
3599 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). |
|
3600 set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3) |
|
3601 also have "\<dots> = |
|
3602 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un |
|
3603 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un |
|
3604 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))" |
|
3605 by (simp only: set_map iupt_set set.simps) |
|
3606 also have "\<dots> = |
|
3607 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un |
|
3608 (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 |
|
3609 (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 |
|
3610 finally |
|
3611 have FS: "?SS (Floor a) = |
|
3612 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un |
|
3613 (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 |
|
3614 (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 |
|
3615 show ?case |
|
3616 proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -) |
|
3617 fix p n s |
|
3618 let ?ths = "(?I p \<longrightarrow> (?N (Floor a) = ?N (CN 0 n s))) \<and> numbound0 s \<and> isrlfm p" |
|
3619 assume "(\<exists>ba. (p, 0, ba) \<in> set (rsplit0 a) \<and> n = 0 \<and> s = Floor ba) \<or> |
|
3620 (\<exists>ab ac ba. |
|
3621 (ab, ac, ba) \<in> set (rsplit0 a) \<and> |
|
3622 0 < ac \<and> |
|
3623 (\<exists>j. p = fp ab ac ba j \<and> |
|
3624 n = 0 \<and> s = Add (Floor ba) (C j) \<and> 0 \<le> j \<and> j \<le> ac)) \<or> |
|
3625 (\<exists>ab ac ba. |
|
3626 (ab, ac, ba) \<in> set (rsplit0 a) \<and> |
|
3627 ac < 0 \<and> |
|
3628 (\<exists>j. p = fp ab ac ba j \<and> |
|
3629 n = 0 \<and> s = Add (Floor ba) (C j) \<and> ac \<le> j \<and> j \<le> 0))" |
|
3630 moreover |
|
3631 {fix s' |
|
3632 assume "(p, 0, s') \<in> ?SS a" and "n = 0" and "s = Floor s'" |
|
3633 hence ?ths using prems by auto} |
|
3634 moreover |
|
3635 { fix p' n' s' j |
|
3636 assume pns: "(p', n', s') \<in> ?SS a" |
|
3637 and np: "0 < n'" |
|
3638 and p_def: "p = ?p (p',n',s') j" |
|
3639 and n0: "n = 0" |
|
3640 and s_def: "s = (Add (Floor s') (C j))" |
|
3641 and jp: "0 \<le> j" and jn: "j \<le> n'" |
|
3642 from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow> |
|
3643 Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and> |
|
3644 numbound0 s' \<and> isrlfm p'" by blast |
|
3645 hence nb: "numbound0 s'" by simp |
|
3646 from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numsub_nb) |
|
3647 let ?nxs = "CN 0 n' s'" |
|
3648 let ?l = "floor (?N s') + j" |
|
3649 from H |
|
3650 have "?I (?p (p',n',s') j) \<longrightarrow> |
|
3651 (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" |
|
3652 by (simp add: fp_def np algebra_simps numsub numadd numfloor) |
|
3653 also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))" |
|
3654 using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp |
|
3655 moreover |
|
3656 have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp |
|
3657 ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" |
|
3658 by blast |
|
3659 with s_def n0 p_def nb nf have ?ths by auto} |
|
3660 moreover |
|
3661 {fix p' n' s' j |
|
3662 assume pns: "(p', n', s') \<in> ?SS a" |
|
3663 and np: "n' < 0" |
|
3664 and p_def: "p = ?p (p',n',s') j" |
|
3665 and n0: "n = 0" |
|
3666 and s_def: "s = (Add (Floor s') (C j))" |
|
3667 and jp: "n' \<le> j" and jn: "j \<le> 0" |
|
3668 from prems pns have H:"(Ifm ((x\<Colon>real) # (bs\<Colon>real list)) p' \<longrightarrow> |
|
3669 Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \<and> |
|
3670 numbound0 s' \<and> isrlfm p'" by blast |
|
3671 hence nb: "numbound0 s'" by simp |
|
3672 from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by (simp add: numneg_nb) |
|
3673 let ?nxs = "CN 0 n' s'" |
|
3674 let ?l = "floor (?N s') + j" |
|
3675 from H |
|
3676 have "?I (?p (p',n',s') j) \<longrightarrow> |
|
3677 (((?N ?nxs \<ge> real ?l) \<and> (?N ?nxs < real (?l + 1))) \<and> (?N a = ?N ?nxs ))" |
|
3678 by (simp add: np fp_def algebra_simps numneg numfloor numadd numsub) |
|
3679 also have "\<dots> \<longrightarrow> ((floor (?N ?nxs) = ?l) \<and> (?N a = ?N ?nxs ))" |
|
3680 using floor_int_eq[where x="?N ?nxs" and n="?l"] by simp |
|
3681 moreover |
|
3682 have "\<dots> \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" by simp |
|
3683 ultimately have "?I (?p (p',n',s') j) \<longrightarrow> (?N (Floor a) = ?N ((Add (Floor s') (C j))))" |
|
3684 by blast |
|
3685 with s_def n0 p_def nb nf have ?ths by auto} |
|
3686 ultimately show ?ths by auto |
|
3687 qed |
|
3688 next |
|
3689 case (3 a b) then show ?case |
|
3690 apply auto |
|
3691 apply (erule_tac x = "(aa, aaa, ba)" in ballE) apply simp_all |
|
3692 apply (erule_tac x = "(ab, ac, baa)" in ballE) apply simp_all |
|
3693 done |
|
3694 qed (auto simp add: Let_def split_def algebra_simps conj_rl) |
|
3695 |
|
3696 lemma real_in_int_intervals: |
|
3697 assumes xb: "real m \<le> x \<and> x < real ((n::int) + 1)" |
|
3698 shows "\<exists> j\<in> {m.. n}. real j \<le> x \<and> x < real (j+1)" (is "\<exists> j\<in> ?N. ?P j") |
|
3699 by (rule bexI[where P="?P" and x="floor x" and A="?N"]) |
|
3700 (auto simp add: floor_less_eq[where x="x" and a="n+1", simplified] xb[simplified] floor_mono2[where x="real m" and y="x", OF conjunct1[OF xb], simplified floor_real_of_int[where n="m"]]) |
|
3701 |
|
3702 lemma rsplit0_complete: |
|
3703 assumes xp:"0 \<le> x" and x1:"x < 1" |
|
3704 shows "\<exists> (p,n,s) \<in> set (rsplit0 t). Ifm (x#bs) p" (is "\<exists> (p,n,s) \<in> ?SS t. ?I p") |
|
3705 proof(induct t rule: rsplit0.induct) |
|
3706 case (2 a b) |
|
3707 from prems have "\<exists> (pa,na,sa) \<in> ?SS a. ?I pa" by auto |
|
3708 then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\<in> ?SS a \<and> ?I pa" by blast |
|
3709 from prems have "\<exists> (pb,nb,sb) \<in> ?SS b. ?I pb" by auto |
|
3710 then obtain "pb" "nb" "sb" where pb: "(pb,nb,sb)\<in> ?SS b \<and> ?I pb" by blast |
|
3711 from pa pb have th: "((pa,na,sa),(pb,nb,sb)) \<in> set[(x,y). x\<leftarrow>rsplit0 a, y\<leftarrow>rsplit0 b]" |
|
3712 by (auto) |
|
3713 let ?f="(\<lambda> ((p,n,t),(q,m,s)). (And p q, n+m, Add t s))" |
|
3714 from imageI[OF th, where f="?f"] have "?f ((pa,na,sa),(pb,nb,sb)) \<in> ?SS (Add a b)" |
|
3715 by (simp add: Let_def) |
|
3716 hence "(And pa pb, na +nb, Add sa sb) \<in> ?SS (Add a b)" by simp |
|
3717 moreover from pa pb have "?I (And pa pb)" by simp |
|
3718 ultimately show ?case by blast |
|
3719 next |
|
3720 case (5 a) |
|
3721 let ?p = "\<lambda> (p,n,s) j. fp p n s j" |
|
3722 let ?f = "(\<lambda> (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))" |
|
3723 let ?J = "\<lambda> n. if n>0 then iupt (0,n) else iupt (n,0)" |
|
3724 let ?ff=" (\<lambda> (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" |
|
3725 have int_cases: "\<forall> (i::int). i= 0 \<or> i < 0 \<or> i > 0" by arith |
|
3726 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 |
|
3727 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)) (iupt(0,n))" |
|
3728 by auto |
|
3729 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)) (iupt(0,n)))))" |
|
3730 proof- |
|
3731 fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g |
|
3732 assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c" |
|
3733 thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))" |
|
3734 by (auto simp add: split_def) |
|
3735 qed |
|
3736 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)) (iupt(n,0))" |
|
3737 by auto |
|
3738 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)) (iupt(n,0)))))" |
|
3739 proof- |
|
3740 fix M :: "('a\<times>'b\<times>'c) set" and f :: "('a\<times>'b\<times>'c) \<Rightarrow> 'd list" and g |
|
3741 assume "\<forall> (a,b,c) \<in> M. f (a,b,c) = g a b c" |
|
3742 thus "(UNION M (\<lambda> (a,b,c). set (f (a,b,c)))) = (UNION M (\<lambda> (a,b,c). set (g a b c)))" |
|
3743 by (auto simp add: split_def) |
|
3744 qed |
|
3745 |
|
3746 have "?SS (Floor a) = UNION (?SS a) (\<lambda>x. set (?ff x))" by (auto simp add: foldl_conv_concat) |
|
3747 also have "\<dots> = UNION (?SS a) (\<lambda> (p,n,s). set (?ff (p,n,s)))" by auto |
|
3748 also have "\<dots> = |
|
3749 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un |
|
3750 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (?ff (p,n,s)))) Un |
|
3751 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (?ff (p,n,s)))))" |
|
3752 using int_cases[rule_format] by blast |
|
3753 also have "\<dots> = |
|
3754 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). set [(p,0,Floor s)])) Un |
|
3755 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) (iupt(0,n))))) Un |
|
3756 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). set (map (?f(p,n,s)) (iupt(n,0))))))" by (simp only: U1 U2 U3) |
|
3757 also have "\<dots> = |
|
3758 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un |
|
3759 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n>0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un |
|
3760 (UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n<0} (\<lambda> (p,n,s). (?f(p,n,s)) ` {n .. 0})))" |
|
3761 by (simp only: set_map iupt_set set.simps) |
|
3762 also have "\<dots> = |
|
3763 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un |
|
3764 (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 |
|
3765 (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 |
|
3766 finally |
|
3767 have FS: "?SS (Floor a) = |
|
3768 ((UNION {(p,n,s). (p,n,s) \<in> ?SS a \<and> n=0} (\<lambda> (p,n,s). {(p,0,Floor s)})) Un |
|
3769 (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 |
|
3770 (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 |
|
3771 from prems have "\<exists> (p,n,s) \<in> ?SS a. ?I p" by auto |
|
3772 then obtain "p" "n" "s" where pns: "(p,n,s) \<in> ?SS a \<and> ?I p" by blast |
|
3773 let ?N = "\<lambda> t. Inum (x#bs) t" |
|
3774 from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \<and> numbound0 s \<and> isrlfm p" |
|
3775 by auto |
|
3776 |
|
3777 have "n=0 \<or> n >0 \<or> n <0" by arith |
|
3778 moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto } |
|
3779 moreover |
|
3780 { |
|
3781 assume np: "n > 0" |
|
3782 from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) \<le> ?N s" by simp |
|
3783 also from mult_left_mono[OF xp] np have "?N s \<le> real n * x + ?N s" by simp |
|
3784 finally have "?N (Floor s) \<le> real n * x + ?N s" . |
|
3785 moreover |
|
3786 {from mult_strict_left_mono[OF x1] np |
|
3787 have "real n *x + ?N s < real n + ?N s" by simp |
|
3788 also from real_of_int_floor_add_one_gt[where r="?N s"] |
|
3789 have "\<dots> < real n + ?N (Floor s) + 1" by simp |
|
3790 finally have "real n *x + ?N s < ?N (Floor s) + real (n+1)" by simp} |
|
3791 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 |
|
3792 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 |
|
3793 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 |
|
3794 |
|
3795 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" |
|
3796 by(simp only: myl[rule_format, where b="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)"]) |
|
3797 hence "\<exists> j\<in> {0.. n}. ?I (?p (p,n,s) j)" |
|
3798 using pns by (simp add: fp_def np algebra_simps numsub numadd) |
|
3799 then obtain "j" where j_def: "j\<in> {0 .. n} \<and> ?I (?p (p,n,s) j)" by blast |
|
3800 hence "\<exists>x \<in> {?p (p,n,s) j |j. 0\<le> j \<and> j \<le> n }. ?I x" by auto |
|
3801 hence ?case using pns |
|
3802 by (simp only: FS,simp add: bex_Un) |
|
3803 (rule disjI2, rule disjI1,rule exI [where x="p"], |
|
3804 rule exI [where x="n"],rule exI [where x="s"],simp_all add: np) |
|
3805 } |
|
3806 moreover |
|
3807 { assume nn: "n < 0" hence np: "-n >0" by simp |
|
3808 from real_of_int_floor_le[where r="?N s"] have "?N (Floor s) + 1 > ?N s" by simp |
|
3809 moreover from mult_left_mono_neg[OF xp] nn have "?N s \<ge> real n * x + ?N s" by simp |
|
3810 ultimately have "?N (Floor s) + 1 > real n * x + ?N s" by arith |
|
3811 moreover |
|
3812 {from mult_strict_left_mono_neg[OF x1, where c="real n"] nn |
|
3813 have "real n *x + ?N s \<ge> real n + ?N s" by simp |
|
3814 moreover from real_of_int_floor_le[where r="?N s"] have "real n + ?N s \<ge> real n + ?N (Floor s)" by simp |
|
3815 ultimately have "real n *x + ?N s \<ge> ?N (Floor s) + real n" |
|
3816 by (simp only: algebra_simps)} |
|
3817 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 |
|
3818 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 |
|
3819 have th1: "\<forall> (a::real). (- a > 0) = (a < 0)" by auto |
|
3820 have th2: "\<forall> (a::real). (0 \<ge> - a) = (a \<ge> 0)" by auto |
|
3821 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 |
|
3822 |
|
3823 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" |
|
3824 by(simp only: myl[rule_format, where b="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)"]) |
|
3825 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]) |
|
3826 hence "\<exists> j\<in> {n.. 0}. ?I (?p (p,n,s) j)" |
|
3827 using pns by (simp add: fp_def nn diff_def add_ac mult_ac numfloor numadd numneg |
|
3828 del: diff_less_0_iff_less diff_le_0_iff_le) |
|
3829 then obtain "j" where j_def: "j\<in> {n .. 0} \<and> ?I (?p (p,n,s) j)" by blast |
|
3830 hence "\<exists>x \<in> {?p (p,n,s) j |j. n\<le> j \<and> j \<le> 0 }. ?I x" by auto |
|
3831 hence ?case using pns |
|
3832 by (simp only: FS,simp add: bex_Un) |
|
3833 (rule disjI2, rule disjI2,rule exI [where x="p"], |
|
3834 rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn) |
|
3835 } |
|
3836 ultimately show ?case by blast |
|
3837 qed (auto simp add: Let_def split_def) |
|
3838 |
|
3839 (* Linearize a formula where Bound 0 ranges over [0,1) *) |
|
3840 |
|
3841 constdefs rsplit :: "(int \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm" |
|
3842 "rsplit f a \<equiv> foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) (rsplit0 a)) F" |
|
3843 |
|
3844 lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\<exists> x \<in> set xs. Ifm bs (f x))" |
|
3845 by(induct xs, simp_all) |
|
3846 |
|
3847 lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\<forall> x \<in> set xs. Ifm bs (f x))" |
|
3848 by(induct xs, simp_all) |
|
3849 |
|
3850 lemma foldr_disj_map_rlfm: |
|
3851 assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)" |
|
3852 and \<phi>: "\<forall> (\<phi>,n,s) \<in> set xs. numbound0 s \<and> isrlfm \<phi>" |
|
3853 shows "isrlfm (foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) xs) F)" |
|
3854 using lf \<phi> by (induct xs, auto) |
|
3855 |
|
3856 lemma rsplit_ex: "Ifm bs (rsplit f a) = (\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). Ifm bs (conj \<phi> (f n s)))" |
|
3857 using foldr_disj_map[where xs="rsplit0 a"] rsplit_def by (simp add: split_def) |
|
3858 |
|
3859 lemma rsplit_l: assumes lf: "\<forall> n s. numbound0 s \<longrightarrow> isrlfm (f n s)" |
|
3860 shows "isrlfm (rsplit f a)" |
|
3861 proof- |
|
3862 from rsplit0_cs[where t="a"] have th: "\<forall> (\<phi>,n,s) \<in> set (rsplit0 a). numbound0 s \<and> isrlfm \<phi>" by blast |
|
3863 from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp |
|
3864 qed |
|
3865 |
|
3866 lemma rsplit: |
|
3867 assumes xp: "x \<ge> 0" and x1: "x < 1" |
|
3868 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))" |
|
3869 shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)" |
|
3870 proof(auto) |
|
3871 let ?I = "\<lambda>x p. Ifm (x#bs) p" |
|
3872 let ?N = "\<lambda> x t. Inum (x#bs) t" |
|
3873 assume "?I x (rsplit f a)" |
|
3874 hence "\<exists> (\<phi>,n,s) \<in> set (rsplit0 a). ?I x (And \<phi> (f n s))" using rsplit_ex by simp |
|
3875 then obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and "?I x (And \<phi> (f n s))" by blast |
|
3876 hence \<phi>: "?I x \<phi>" and fns: "?I x (f n s)" by auto |
|
3877 from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \<phi> |
|
3878 have th: "(?N x a = ?N x (CN 0 n s)) \<and> numbound0 s" by auto |
|
3879 from f[rule_format, OF th] fns show "?I x (g a)" by simp |
|
3880 next |
|
3881 let ?I = "\<lambda>x p. Ifm (x#bs) p" |
|
3882 let ?N = "\<lambda> x t. Inum (x#bs) t" |
|
3883 assume ga: "?I x (g a)" |
|
3884 from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] |
|
3885 obtain "\<phi>" "n" "s" where fnsS:"(\<phi>,n,s) \<in> set (rsplit0 a)" and fx: "?I x \<phi>" by blast |
|
3886 from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx |
|
3887 have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto |
|
3888 with ga f have "?I x (f n s)" by auto |
|
3889 with rsplit_ex fnsS fx show "?I x (rsplit f a)" by auto |
|
3890 qed |
|
3891 |
|
3892 definition lt :: "int \<Rightarrow> num \<Rightarrow> fm" where |
|
3893 lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) |
|
3894 else (Gt (CN 0 (-c) (Neg t))))" |
|
3895 |
|
3896 definition le :: "int \<Rightarrow> num \<Rightarrow> fm" where |
|
3897 le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) |
|
3898 else (Ge (CN 0 (-c) (Neg t))))" |
|
3899 |
|
3900 definition gt :: "int \<Rightarrow> num \<Rightarrow> fm" where |
|
3901 gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) |
|
3902 else (Lt (CN 0 (-c) (Neg t))))" |
|
3903 |
|
3904 definition ge :: "int \<Rightarrow> num \<Rightarrow> fm" where |
|
3905 ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) |
|
3906 else (Le (CN 0 (-c) (Neg t))))" |
|
3907 |
|
3908 definition eq :: "int \<Rightarrow> num \<Rightarrow> fm" where |
|
3909 eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) |
|
3910 else (Eq (CN 0 (-c) (Neg t))))" |
|
3911 |
|
3912 definition neq :: "int \<Rightarrow> num \<Rightarrow> fm" where |
|
3913 neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) |
|
3914 else (NEq (CN 0 (-c) (Neg t))))" |
|
3915 |
|
3916 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)" |
|
3917 (is "\<forall> a n s . ?N a = ?N (CN 0 n s) \<and> _\<longrightarrow> ?I (lt n s) = ?I (Lt a)") |
|
3918 proof(clarify) |
|
3919 fix a n s |
|
3920 assume H: "?N a = ?N (CN 0 n s)" |
|
3921 show "?I (lt n s) = ?I (Lt a)" using H by (cases "n=0", (simp add: lt_def)) |
|
3922 (cases "n > 0", simp_all add: lt_def algebra_simps myless[rule_format, where b="0"]) |
|
3923 qed |
|
3924 |
|
3925 lemma lt_l: "isrlfm (rsplit lt a)" |
|
3926 by (rule rsplit_l[where f="lt" and a="a"], auto simp add: lt_def, |
|
3927 case_tac s, simp_all, case_tac "nat", simp_all) |
|
3928 |
|
3929 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)") |
|
3930 proof(clarify) |
|
3931 fix a n s |
|
3932 assume H: "?N a = ?N (CN 0 n s)" |
|
3933 show "?I (le n s) = ?I (Le a)" using H by (cases "n=0", (simp add: le_def)) |
|
3934 (cases "n > 0", simp_all add: le_def algebra_simps myl[rule_format, where b="0"]) |
|
3935 qed |
|
3936 |
|
3937 lemma le_l: "isrlfm (rsplit le a)" |
|
3938 by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) |
|
3939 (case_tac s, simp_all, case_tac "nat",simp_all) |
|
3940 |
|
3941 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)") |
|
3942 proof(clarify) |
|
3943 fix a n s |
|
3944 assume H: "?N a = ?N (CN 0 n s)" |
|
3945 show "?I (gt n s) = ?I (Gt a)" using H by (cases "n=0", (simp add: gt_def)) |
|
3946 (cases "n > 0", simp_all add: gt_def algebra_simps myless[rule_format, where b="0"]) |
|
3947 qed |
|
3948 lemma gt_l: "isrlfm (rsplit gt a)" |
|
3949 by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) |
|
3950 (case_tac s, simp_all, case_tac "nat", simp_all) |
|
3951 |
|
3952 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)") |
|
3953 proof(clarify) |
|
3954 fix a n s |
|
3955 assume H: "?N a = ?N (CN 0 n s)" |
|
3956 show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def)) |
|
3957 (cases "n > 0", simp_all add: ge_def algebra_simps myl[rule_format, where b="0"]) |
|
3958 qed |
|
3959 lemma ge_l: "isrlfm (rsplit ge a)" |
|
3960 by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) |
|
3961 (case_tac s, simp_all, case_tac "nat", simp_all) |
|
3962 |
|
3963 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)") |
|
3964 proof(clarify) |
|
3965 fix a n s |
|
3966 assume H: "?N a = ?N (CN 0 n s)" |
|
3967 show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def algebra_simps) |
|
3968 qed |
|
3969 lemma eq_l: "isrlfm (rsplit eq a)" |
|
3970 by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) |
|
3971 (case_tac s, simp_all, case_tac"nat", simp_all) |
|
3972 |
|
3973 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)") |
|
3974 proof(clarify) |
|
3975 fix a n s bs |
|
3976 assume H: "?N a = ?N (CN 0 n s)" |
|
3977 show "?I (neq n s) = ?I (NEq a)" using H by (auto simp add: neq_def algebra_simps) |
|
3978 qed |
|
3979 |
|
3980 lemma neq_l: "isrlfm (rsplit neq a)" |
|
3981 by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) |
|
3982 (case_tac s, simp_all, case_tac"nat", simp_all) |
|
3983 |
|
3984 lemma small_le: |
|
3985 assumes u0:"0 \<le> u" and u1: "u < 1" |
|
3986 shows "(-u \<le> real (n::int)) = (0 \<le> n)" |
|
3987 using u0 u1 by auto |
|
3988 |
|
3989 lemma small_lt: |
|
3990 assumes u0:"0 \<le> u" and u1: "u < 1" |
|
3991 shows "(real (n::int) < real (m::int) - u) = (n < m)" |
|
3992 using u0 u1 by auto |
|
3993 |
|
3994 lemma rdvd01_cs: |
|
3995 assumes up: "u \<ge> 0" and u1: "u<1" and np: "real n > 0" |
|
3996 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") |
|
3997 proof- |
|
3998 let ?ss = "s - real (floor s)" |
|
3999 from real_of_int_floor_add_one_gt[where r="s", simplified myless[rule_format,where a="s"]] |
|
4000 real_of_int_floor_le[where r="s"] have ss0:"?ss \<ge> 0" and ss1:"?ss < 1" |
|
4001 by (auto simp add: myl[rule_format, where b="s", symmetric] myless[rule_format, where a="?ss"]) |
|
4002 from np have n0: "real n \<ge> 0" by simp |
|
4003 from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] |
|
4004 have nu0:"real n * u - s \<ge> -s" and nun:"real n * u -s < real n - s" by auto |
|
4005 from int_rdvd_real[where i="i" and x="real (n::int) * u - s"] |
|
4006 have "real i rdvd real n * u - s = |
|
4007 (i dvd floor (real n * u -s) \<and> (real (floor (real n * u - s)) = real n * u - s ))" |
|
4008 (is "_ = (?DE)" is "_ = (?D \<and> ?E)") by simp |
|
4009 also have "\<dots> = (?DE \<and> real(floor (real n * u - s) + floor s)\<ge> -?ss |
|
4010 \<and> real(floor (real n * u - s) + floor s)< real n - ?ss)" (is "_=(?DE \<and>real ?a \<ge> _ \<and> real ?a < _)") |
|
4011 using nu0 nun by auto |
|
4012 also have "\<dots> = (?DE \<and> ?a \<ge> 0 \<and> ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1]) |
|
4013 also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. ?a = j))" by simp |
|
4014 also have "\<dots> = (?DE \<and> (\<exists> j\<in> {0 .. (n - 1)}. real (\<lfloor>real n * u - s\<rfloor>) = real j - real \<lfloor>s\<rfloor> ))" |
|
4015 by (simp only: algebra_simps real_of_int_diff[symmetric] real_of_int_inject del: real_of_int_diff) |
|
4016 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>"] |
|
4017 by (auto cong: conj_cong) |
|
4018 also have "\<dots> = ?rhs" by(simp cong: conj_cong) (simp add: algebra_simps ) |
|
4019 finally show ?thesis . |
|
4020 qed |
|
4021 |
|
4022 definition |
|
4023 DVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" |
|
4024 where |
|
4025 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))))) (iupt(0,n - 1))) F)" |
|
4026 |
|
4027 definition |
|
4028 NDVDJ:: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" |
|
4029 where |
|
4030 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))))) (iupt(0,n - 1))) T)" |
|
4031 |
|
4032 lemma DVDJ_DVD: |
|
4033 assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0" |
|
4034 shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))" |
|
4035 proof- |
|
4036 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))))" |
|
4037 let ?s= "Inum (x#bs) s" |
|
4038 from foldr_disj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"] |
|
4039 have "Ifm (x#bs) (DVDJ i n s) = (\<exists> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" |
|
4040 by (simp add: iupt_set np DVDJ_def del: iupt.simps) |
|
4041 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_def[symmetric]) |
|
4042 also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] |
|
4043 have "\<dots> = (real i rdvd real n * x - (-?s))" by simp |
|
4044 finally show ?thesis by simp |
|
4045 qed |
|
4046 |
|
4047 lemma NDVDJ_NDVD: |
|
4048 assumes xp:"x\<ge> 0" and x1: "x < 1" and np:"real n > 0" |
|
4049 shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))" |
|
4050 proof- |
|
4051 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))))" |
|
4052 let ?s= "Inum (x#bs) s" |
|
4053 from foldr_conj_map[where xs="iupt(0,n - 1)" and bs="x#bs" and f="?f"] |
|
4054 have "Ifm (x#bs) (NDVDJ i n s) = (\<forall> j\<in> {0 .. (n - 1)}. Ifm (x#bs) (?f j))" |
|
4055 by (simp add: iupt_set np NDVDJ_def del: iupt.simps) |
|
4056 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_def[symmetric]) |
|
4057 also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] |
|
4058 have "\<dots> = (\<not> (real i rdvd real n * x - (-?s)))" by simp |
|
4059 finally show ?thesis by simp |
|
4060 qed |
|
4061 |
|
4062 lemma foldr_disj_map_rlfm2: |
|
4063 assumes lf: "\<forall> n . isrlfm (f n)" |
|
4064 shows "isrlfm (foldr disj (map f xs) F)" |
|
4065 using lf by (induct xs, auto) |
|
4066 lemma foldr_And_map_rlfm2: |
|
4067 assumes lf: "\<forall> n . isrlfm (f n)" |
|
4068 shows "isrlfm (foldr conj (map f xs) T)" |
|
4069 using lf by (induct xs, auto) |
|
4070 |
|
4071 lemma DVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" |
|
4072 shows "isrlfm (DVDJ i n s)" |
|
4073 proof- |
|
4074 let ?f="\<lambda>j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) |
|
4075 (Dvd i (Sub (C j) (Floor (Neg s))))" |
|
4076 have th: "\<forall> j. isrlfm (?f j)" using nb np by auto |
|
4077 from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp |
|
4078 qed |
|
4079 |
|
4080 lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" |
|
4081 shows "isrlfm (NDVDJ i n s)" |
|
4082 proof- |
|
4083 let ?f="\<lambda>j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) |
|
4084 (NDvd i (Sub (C j) (Floor (Neg s))))" |
|
4085 have th: "\<forall> j. isrlfm (?f j)" using nb np by auto |
|
4086 from NDVDJ_def foldr_And_map_rlfm2[OF th] show ?thesis by auto |
|
4087 qed |
|
4088 |
|
4089 definition DVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where |
|
4090 DVD_def: "DVD i c t = |
|
4091 (if i=0 then eq c t else |
|
4092 if c = 0 then (Dvd i t) else if c >0 then DVDJ (abs i) c t else DVDJ (abs i) (-c) (Neg t))" |
|
4093 |
|
4094 definition NDVD :: "int \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where |
|
4095 "NDVD i c t = |
|
4096 (if i=0 then neq c t else |
|
4097 if c = 0 then (NDvd i t) else if c >0 then NDVDJ (abs i) c t else NDVDJ (abs i) (-c) (Neg t))" |
|
4098 |
|
4099 lemma DVD_mono: |
|
4100 assumes xp: "0\<le> x" and x1: "x < 1" |
|
4101 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)" |
|
4102 (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (DVD i n s) = ?I (Dvd i a)") |
|
4103 proof(clarify) |
|
4104 fix a n s |
|
4105 assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" |
|
4106 let ?th = "?I (DVD i n s) = ?I (Dvd i a)" |
|
4107 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 |
|
4108 moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] |
|
4109 by (simp add: DVD_def rdvd_left_0_eq)} |
|
4110 moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H DVD_def) } |
|
4111 moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th |
|
4112 by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 |
|
4113 rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } |
|
4114 moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)} |
|
4115 ultimately show ?th by blast |
|
4116 qed |
|
4117 |
|
4118 lemma NDVD_mono: assumes xp: "0\<le> x" and x1: "x < 1" |
|
4119 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)" |
|
4120 (is "\<forall> a n s. ?N a = ?N (CN 0 n s) \<and> _ \<longrightarrow> ?I (NDVD i n s) = ?I (NDvd i a)") |
|
4121 proof(clarify) |
|
4122 fix a n s |
|
4123 assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" |
|
4124 let ?th = "?I (NDVD i n s) = ?I (NDvd i a)" |
|
4125 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 |
|
4126 moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] |
|
4127 by (simp add: NDVD_def rdvd_left_0_eq)} |
|
4128 moreover {assume inz: "i\<noteq>0" and "n=0" hence ?th by (simp add: H NDVD_def) } |
|
4129 moreover {assume inz: "i\<noteq>0" and "n<0" hence ?th |
|
4130 by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 |
|
4131 rdvd_minus[where d="i" and t="real n * x + Inum (x # bs) s"]) } |
|
4132 moreover {assume inz: "i\<noteq>0" and "n>0" hence ?th |
|
4133 by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)} |
|
4134 ultimately show ?th by blast |
|
4135 qed |
|
4136 |
|
4137 lemma DVD_l: "isrlfm (rsplit (DVD i) a)" |
|
4138 by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) |
|
4139 (case_tac s, simp_all, case_tac "nat", simp_all) |
|
4140 |
|
4141 lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)" |
|
4142 by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) |
|
4143 (case_tac s, simp_all, case_tac "nat", simp_all) |
|
4144 |
|
4145 consts rlfm :: "fm \<Rightarrow> fm" |
|
4146 recdef rlfm "measure fmsize" |
|
4147 "rlfm (And p q) = conj (rlfm p) (rlfm q)" |
|
4148 "rlfm (Or p q) = disj (rlfm p) (rlfm q)" |
|
4149 "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)" |
|
4150 "rlfm (Iff p q) = disj (conj(rlfm p) (rlfm q)) (conj(rlfm (NOT p)) (rlfm (NOT q)))" |
|
4151 "rlfm (Lt a) = rsplit lt a" |
|
4152 "rlfm (Le a) = rsplit le a" |
|
4153 "rlfm (Gt a) = rsplit gt a" |
|
4154 "rlfm (Ge a) = rsplit ge a" |
|
4155 "rlfm (Eq a) = rsplit eq a" |
|
4156 "rlfm (NEq a) = rsplit neq a" |
|
4157 "rlfm (Dvd i a) = rsplit (\<lambda> t. DVD i t) a" |
|
4158 "rlfm (NDvd i a) = rsplit (\<lambda> t. NDVD i t) a" |
|
4159 "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))" |
|
4160 "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))" |
|
4161 "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))" |
|
4162 "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))" |
|
4163 "rlfm (NOT (NOT p)) = rlfm p" |
|
4164 "rlfm (NOT T) = F" |
|
4165 "rlfm (NOT F) = T" |
|
4166 "rlfm (NOT (Lt a)) = simpfm (rlfm (Ge a))" |
|
4167 "rlfm (NOT (Le a)) = simpfm (rlfm (Gt a))" |
|
4168 "rlfm (NOT (Gt a)) = simpfm (rlfm (Le a))" |
|
4169 "rlfm (NOT (Ge a)) = simpfm (rlfm (Lt a))" |
|
4170 "rlfm (NOT (Eq a)) = simpfm (rlfm (NEq a))" |
|
4171 "rlfm (NOT (NEq a)) = simpfm (rlfm (Eq a))" |
|
4172 "rlfm (NOT (Dvd i a)) = simpfm (rlfm (NDvd i a))" |
|
4173 "rlfm (NOT (NDvd i a)) = simpfm (rlfm (Dvd i a))" |
|
4174 "rlfm p = p" (hints simp add: fmsize_pos) |
|
4175 |
|
4176 lemma bound0at_l : "\<lbrakk>isatom p ; bound0 p\<rbrakk> \<Longrightarrow> isrlfm p" |
|
4177 by (induct p rule: isrlfm.induct, auto) |
|
4178 lemma zgcd_le1: assumes ip: "0 < i" shows "zgcd i j \<le> i" |
|
4179 proof- |
|
4180 from zgcd_zdvd1 have th: "zgcd i j dvd i" by blast |
|
4181 from zdvd_imp_le[OF th ip] show ?thesis . |
|
4182 qed |
|
4183 |
|
4184 |
|
4185 lemma simpfm_rl: "isrlfm p \<Longrightarrow> isrlfm (simpfm p)" |
|
4186 proof (induct p) |
|
4187 case (Lt a) |
|
4188 hence "bound0 (Lt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)" |
|
4189 by (cases a,simp_all, case_tac "nat", simp_all) |
|
4190 moreover |
|
4191 {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))" |
|
4192 using simpfm_bound0 by blast |
|
4193 have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def) |
|
4194 with bn bound0at_l have ?case by blast} |
|
4195 moreover |
|
4196 {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" |
|
4197 { |
|
4198 assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0" |
|
4199 with numgcd_pos[where t="CN 0 c (simpnum e)"] |
|
4200 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp |
|
4201 from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" |
|
4202 by (simp add: numgcd_def zgcd_le1) |
|
4203 from prems have th': "c\<noteq>0" by auto |
|
4204 from prems have cp: "c \<ge> 0" by simp |
|
4205 from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] |
|
4206 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp |
|
4207 } |
|
4208 with prems have ?case |
|
4209 by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} |
|
4210 ultimately show ?case by blast |
|
4211 next |
|
4212 case (Le a) |
|
4213 hence "bound0 (Le a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)" |
|
4214 by (cases a,simp_all, case_tac "nat", simp_all) |
|
4215 moreover |
|
4216 {assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))" |
|
4217 using simpfm_bound0 by blast |
|
4218 have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def) |
|
4219 with bn bound0at_l have ?case by blast} |
|
4220 moreover |
|
4221 {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" |
|
4222 { |
|
4223 assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0" |
|
4224 with numgcd_pos[where t="CN 0 c (simpnum e)"] |
|
4225 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp |
|
4226 from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" |
|
4227 by (simp add: numgcd_def zgcd_le1) |
|
4228 from prems have th': "c\<noteq>0" by auto |
|
4229 from prems have cp: "c \<ge> 0" by simp |
|
4230 from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] |
|
4231 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp |
|
4232 } |
|
4233 with prems have ?case |
|
4234 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} |
|
4235 ultimately show ?case by blast |
|
4236 next |
|
4237 case (Gt a) |
|
4238 hence "bound0 (Gt a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)" |
|
4239 by (cases a,simp_all, case_tac "nat", simp_all) |
|
4240 moreover |
|
4241 {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))" |
|
4242 using simpfm_bound0 by blast |
|
4243 have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def) |
|
4244 with bn bound0at_l have ?case by blast} |
|
4245 moreover |
|
4246 {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" |
|
4247 { |
|
4248 assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0" |
|
4249 with numgcd_pos[where t="CN 0 c (simpnum e)"] |
|
4250 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp |
|
4251 from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" |
|
4252 by (simp add: numgcd_def zgcd_le1) |
|
4253 from prems have th': "c\<noteq>0" by auto |
|
4254 from prems have cp: "c \<ge> 0" by simp |
|
4255 from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] |
|
4256 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp |
|
4257 } |
|
4258 with prems have ?case |
|
4259 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} |
|
4260 ultimately show ?case by blast |
|
4261 next |
|
4262 case (Ge a) |
|
4263 hence "bound0 (Ge a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)" |
|
4264 by (cases a,simp_all, case_tac "nat", simp_all) |
|
4265 moreover |
|
4266 {assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))" |
|
4267 using simpfm_bound0 by blast |
|
4268 have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def) |
|
4269 with bn bound0at_l have ?case by blast} |
|
4270 moreover |
|
4271 {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" |
|
4272 { |
|
4273 assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0" |
|
4274 with numgcd_pos[where t="CN 0 c (simpnum e)"] |
|
4275 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp |
|
4276 from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" |
|
4277 by (simp add: numgcd_def zgcd_le1) |
|
4278 from prems have th': "c\<noteq>0" by auto |
|
4279 from prems have cp: "c \<ge> 0" by simp |
|
4280 from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] |
|
4281 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp |
|
4282 } |
|
4283 with prems have ?case |
|
4284 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} |
|
4285 ultimately show ?case by blast |
|
4286 next |
|
4287 case (Eq a) |
|
4288 hence "bound0 (Eq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)" |
|
4289 by (cases a,simp_all, case_tac "nat", simp_all) |
|
4290 moreover |
|
4291 {assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))" |
|
4292 using simpfm_bound0 by blast |
|
4293 have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def) |
|
4294 with bn bound0at_l have ?case by blast} |
|
4295 moreover |
|
4296 {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" |
|
4297 { |
|
4298 assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0" |
|
4299 with numgcd_pos[where t="CN 0 c (simpnum e)"] |
|
4300 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp |
|
4301 from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" |
|
4302 by (simp add: numgcd_def zgcd_le1) |
|
4303 from prems have th': "c\<noteq>0" by auto |
|
4304 from prems have cp: "c \<ge> 0" by simp |
|
4305 from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] |
|
4306 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp |
|
4307 } |
|
4308 with prems have ?case |
|
4309 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} |
|
4310 ultimately show ?case by blast |
|
4311 next |
|
4312 case (NEq a) |
|
4313 hence "bound0 (NEq a) \<or> (\<exists> c e. a = CN 0 c e \<and> c > 0 \<and> numbound0 e)" |
|
4314 by (cases a,simp_all, case_tac "nat", simp_all) |
|
4315 moreover |
|
4316 {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))" |
|
4317 using simpfm_bound0 by blast |
|
4318 have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def) |
|
4319 with bn bound0at_l have ?case by blast} |
|
4320 moreover |
|
4321 {fix c e assume "a = CN 0 c e" and "c>0" and "numbound0 e" |
|
4322 { |
|
4323 assume cn1:"numgcd (CN 0 c (simpnum e)) \<noteq> 1" and cnz:"numgcd (CN 0 c (simpnum e)) \<noteq> 0" |
|
4324 with numgcd_pos[where t="CN 0 c (simpnum e)"] |
|
4325 have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp |
|
4326 from prems have th:"numgcd (CN 0 c (simpnum e)) \<le> c" |
|
4327 by (simp add: numgcd_def zgcd_le1) |
|
4328 from prems have th': "c\<noteq>0" by auto |
|
4329 from prems have cp: "c \<ge> 0" by simp |
|
4330 from zdiv_mono2[OF cp th1 th, simplified zdiv_self[OF th']] |
|
4331 have "0 < c div numgcd (CN 0 c (simpnum e))" by simp |
|
4332 } |
|
4333 with prems have ?case |
|
4334 by (simp add: Let_def reducecoeff_def simpnum_numbound0 reducecoeffh_numbound0)} |
|
4335 ultimately show ?case by blast |
|
4336 next |
|
4337 case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))" |
|
4338 using simpfm_bound0 by blast |
|
4339 have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) |
|
4340 with bn bound0at_l show ?case by blast |
|
4341 next |
|
4342 case (NDvd i a) hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))" |
|
4343 using simpfm_bound0 by blast |
|
4344 have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) |
|
4345 with bn bound0at_l show ?case by blast |
|
4346 qed(auto simp add: conj_def imp_def disj_def iff_def Let_def simpfm_bound0 numadd_nb numneg_nb) |
|
4347 |
|
4348 lemma rlfm_I: |
|
4349 assumes qfp: "qfree p" |
|
4350 and xp: "0 \<le> x" and x1: "x < 1" |
|
4351 shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \<and> isrlfm (rlfm p)" |
|
4352 using qfp |
|
4353 by (induct p rule: rlfm.induct) |
|
4354 (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 |
|
4355 rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l |
|
4356 rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl) |
|
4357 lemma rlfm_l: |
|
4358 assumes qfp: "qfree p" |
|
4359 shows "isrlfm (rlfm p)" |
|
4360 using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l |
|
4361 by (induct p rule: rlfm.induct,auto simp add: simpfm_rl) |
|
4362 |
|
4363 (* Operations needed for Ferrante and Rackoff *) |
|
4364 lemma rminusinf_inf: |
|
4365 assumes lp: "isrlfm p" |
|
4366 shows "\<exists> z. \<forall> x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p") |
|
4367 using lp |
|
4368 proof (induct p rule: minusinf.induct) |
|
4369 case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto |
|
4370 next |
|
4371 case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto |
|
4372 next |
|
4373 case (3 c e) |
|
4374 from prems have nb: "numbound0 e" by simp |
|
4375 from prems have cp: "real c > 0" by simp |
|
4376 fix a |
|
4377 let ?e="Inum (a#bs) e" |
|
4378 let ?z = "(- ?e) / real c" |
|
4379 {fix x |
|
4380 assume xz: "x < ?z" |
|
4381 hence "(real c * x < - ?e)" |
|
4382 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) |
|
4383 hence "real c * x + ?e < 0" by arith |
|
4384 hence "real c * x + ?e \<noteq> 0" by simp |
|
4385 with xz have "?P ?z x (Eq (CN 0 c e))" |
|
4386 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4387 hence "\<forall> x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp |
|
4388 thus ?case by blast |
|
4389 next |
|
4390 case (4 c e) |
|
4391 from prems have nb: "numbound0 e" by simp |
|
4392 from prems have cp: "real c > 0" by simp |
|
4393 fix a |
|
4394 let ?e="Inum (a#bs) e" |
|
4395 let ?z = "(- ?e) / real c" |
|
4396 {fix x |
|
4397 assume xz: "x < ?z" |
|
4398 hence "(real c * x < - ?e)" |
|
4399 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) |
|
4400 hence "real c * x + ?e < 0" by arith |
|
4401 hence "real c * x + ?e \<noteq> 0" by simp |
|
4402 with xz have "?P ?z x (NEq (CN 0 c e))" |
|
4403 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4404 hence "\<forall> x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp |
|
4405 thus ?case by blast |
|
4406 next |
|
4407 case (5 c e) |
|
4408 from prems have nb: "numbound0 e" by simp |
|
4409 from prems have cp: "real c > 0" by simp |
|
4410 fix a |
|
4411 let ?e="Inum (a#bs) e" |
|
4412 let ?z = "(- ?e) / real c" |
|
4413 {fix x |
|
4414 assume xz: "x < ?z" |
|
4415 hence "(real c * x < - ?e)" |
|
4416 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) |
|
4417 hence "real c * x + ?e < 0" by arith |
|
4418 with xz have "?P ?z x (Lt (CN 0 c e))" |
|
4419 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4420 hence "\<forall> x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp |
|
4421 thus ?case by blast |
|
4422 next |
|
4423 case (6 c e) |
|
4424 from prems have nb: "numbound0 e" by simp |
|
4425 from prems have cp: "real c > 0" by simp |
|
4426 fix a |
|
4427 let ?e="Inum (a#bs) e" |
|
4428 let ?z = "(- ?e) / real c" |
|
4429 {fix x |
|
4430 assume xz: "x < ?z" |
|
4431 hence "(real c * x < - ?e)" |
|
4432 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) |
|
4433 hence "real c * x + ?e < 0" by arith |
|
4434 with xz have "?P ?z x (Le (CN 0 c e))" |
|
4435 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4436 hence "\<forall> x < ?z. ?P ?z x (Le (CN 0 c e))" by simp |
|
4437 thus ?case by blast |
|
4438 next |
|
4439 case (7 c e) |
|
4440 from prems have nb: "numbound0 e" by simp |
|
4441 from prems have cp: "real c > 0" by simp |
|
4442 fix a |
|
4443 let ?e="Inum (a#bs) e" |
|
4444 let ?z = "(- ?e) / real c" |
|
4445 {fix x |
|
4446 assume xz: "x < ?z" |
|
4447 hence "(real c * x < - ?e)" |
|
4448 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) |
|
4449 hence "real c * x + ?e < 0" by arith |
|
4450 with xz have "?P ?z x (Gt (CN 0 c e))" |
|
4451 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4452 hence "\<forall> x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp |
|
4453 thus ?case by blast |
|
4454 next |
|
4455 case (8 c e) |
|
4456 from prems have nb: "numbound0 e" by simp |
|
4457 from prems have cp: "real c > 0" by simp |
|
4458 fix a |
|
4459 let ?e="Inum (a#bs) e" |
|
4460 let ?z = "(- ?e) / real c" |
|
4461 {fix x |
|
4462 assume xz: "x < ?z" |
|
4463 hence "(real c * x < - ?e)" |
|
4464 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) |
|
4465 hence "real c * x + ?e < 0" by arith |
|
4466 with xz have "?P ?z x (Ge (CN 0 c e))" |
|
4467 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4468 hence "\<forall> x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp |
|
4469 thus ?case by blast |
|
4470 qed simp_all |
|
4471 |
|
4472 lemma rplusinf_inf: |
|
4473 assumes lp: "isrlfm p" |
|
4474 shows "\<exists> z. \<forall> x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "\<exists> z. \<forall> x. ?P z x p") |
|
4475 using lp |
|
4476 proof (induct p rule: isrlfm.induct) |
|
4477 case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto |
|
4478 next |
|
4479 case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto |
|
4480 next |
|
4481 case (3 c e) |
|
4482 from prems have nb: "numbound0 e" by simp |
|
4483 from prems have cp: "real c > 0" by simp |
|
4484 fix a |
|
4485 let ?e="Inum (a#bs) e" |
|
4486 let ?z = "(- ?e) / real c" |
|
4487 {fix x |
|
4488 assume xz: "x > ?z" |
|
4489 with mult_strict_right_mono [OF xz cp] cp |
|
4490 have "(real c * x > - ?e)" by (simp add: mult_ac) |
|
4491 hence "real c * x + ?e > 0" by arith |
|
4492 hence "real c * x + ?e \<noteq> 0" by simp |
|
4493 with xz have "?P ?z x (Eq (CN 0 c e))" |
|
4494 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4495 hence "\<forall> x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp |
|
4496 thus ?case by blast |
|
4497 next |
|
4498 case (4 c e) |
|
4499 from prems have nb: "numbound0 e" by simp |
|
4500 from prems have cp: "real c > 0" by simp |
|
4501 fix a |
|
4502 let ?e="Inum (a#bs) e" |
|
4503 let ?z = "(- ?e) / real c" |
|
4504 {fix x |
|
4505 assume xz: "x > ?z" |
|
4506 with mult_strict_right_mono [OF xz cp] cp |
|
4507 have "(real c * x > - ?e)" by (simp add: mult_ac) |
|
4508 hence "real c * x + ?e > 0" by arith |
|
4509 hence "real c * x + ?e \<noteq> 0" by simp |
|
4510 with xz have "?P ?z x (NEq (CN 0 c e))" |
|
4511 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4512 hence "\<forall> x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp |
|
4513 thus ?case by blast |
|
4514 next |
|
4515 case (5 c e) |
|
4516 from prems have nb: "numbound0 e" by simp |
|
4517 from prems have cp: "real c > 0" by simp |
|
4518 fix a |
|
4519 let ?e="Inum (a#bs) e" |
|
4520 let ?z = "(- ?e) / real c" |
|
4521 {fix x |
|
4522 assume xz: "x > ?z" |
|
4523 with mult_strict_right_mono [OF xz cp] cp |
|
4524 have "(real c * x > - ?e)" by (simp add: mult_ac) |
|
4525 hence "real c * x + ?e > 0" by arith |
|
4526 with xz have "?P ?z x (Lt (CN 0 c e))" |
|
4527 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4528 hence "\<forall> x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp |
|
4529 thus ?case by blast |
|
4530 next |
|
4531 case (6 c e) |
|
4532 from prems have nb: "numbound0 e" by simp |
|
4533 from prems have cp: "real c > 0" by simp |
|
4534 fix a |
|
4535 let ?e="Inum (a#bs) e" |
|
4536 let ?z = "(- ?e) / real c" |
|
4537 {fix x |
|
4538 assume xz: "x > ?z" |
|
4539 with mult_strict_right_mono [OF xz cp] cp |
|
4540 have "(real c * x > - ?e)" by (simp add: mult_ac) |
|
4541 hence "real c * x + ?e > 0" by arith |
|
4542 with xz have "?P ?z x (Le (CN 0 c e))" |
|
4543 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4544 hence "\<forall> x > ?z. ?P ?z x (Le (CN 0 c e))" by simp |
|
4545 thus ?case by blast |
|
4546 next |
|
4547 case (7 c e) |
|
4548 from prems have nb: "numbound0 e" by simp |
|
4549 from prems have cp: "real c > 0" by simp |
|
4550 fix a |
|
4551 let ?e="Inum (a#bs) e" |
|
4552 let ?z = "(- ?e) / real c" |
|
4553 {fix x |
|
4554 assume xz: "x > ?z" |
|
4555 with mult_strict_right_mono [OF xz cp] cp |
|
4556 have "(real c * x > - ?e)" by (simp add: mult_ac) |
|
4557 hence "real c * x + ?e > 0" by arith |
|
4558 with xz have "?P ?z x (Gt (CN 0 c e))" |
|
4559 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4560 hence "\<forall> x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp |
|
4561 thus ?case by blast |
|
4562 next |
|
4563 case (8 c e) |
|
4564 from prems have nb: "numbound0 e" by simp |
|
4565 from prems have cp: "real c > 0" by simp |
|
4566 fix a |
|
4567 let ?e="Inum (a#bs) e" |
|
4568 let ?z = "(- ?e) / real c" |
|
4569 {fix x |
|
4570 assume xz: "x > ?z" |
|
4571 with mult_strict_right_mono [OF xz cp] cp |
|
4572 have "(real c * x > - ?e)" by (simp add: mult_ac) |
|
4573 hence "real c * x + ?e > 0" by arith |
|
4574 with xz have "?P ?z x (Ge (CN 0 c e))" |
|
4575 using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } |
|
4576 hence "\<forall> x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp |
|
4577 thus ?case by blast |
|
4578 qed simp_all |
|
4579 |
|
4580 lemma rminusinf_bound0: |
|
4581 assumes lp: "isrlfm p" |
|
4582 shows "bound0 (minusinf p)" |
|
4583 using lp |
|
4584 by (induct p rule: minusinf.induct) simp_all |
|
4585 |
|
4586 lemma rplusinf_bound0: |
|
4587 assumes lp: "isrlfm p" |
|
4588 shows "bound0 (plusinf p)" |
|
4589 using lp |
|
4590 by (induct p rule: plusinf.induct) simp_all |
|
4591 |
|
4592 lemma rminusinf_ex: |
|
4593 assumes lp: "isrlfm p" |
|
4594 and ex: "Ifm (a#bs) (minusinf p)" |
|
4595 shows "\<exists> x. Ifm (x#bs) p" |
|
4596 proof- |
|
4597 from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex |
|
4598 have th: "\<forall> x. Ifm (x#bs) (minusinf p)" by auto |
|
4599 from rminusinf_inf[OF lp, where bs="bs"] |
|
4600 obtain z where z_def: "\<forall>x<z. Ifm (x # bs) (minusinf p) = Ifm (x # bs) p" by blast |
|
4601 from th have "Ifm ((z - 1)#bs) (minusinf p)" by simp |
|
4602 moreover have "z - 1 < z" by simp |
|
4603 ultimately show ?thesis using z_def by auto |
|
4604 qed |
|
4605 |
|
4606 lemma rplusinf_ex: |
|
4607 assumes lp: "isrlfm p" |
|
4608 and ex: "Ifm (a#bs) (plusinf p)" |
|
4609 shows "\<exists> x. Ifm (x#bs) p" |
|
4610 proof- |
|
4611 from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex |
|
4612 have th: "\<forall> x. Ifm (x#bs) (plusinf p)" by auto |
|
4613 from rplusinf_inf[OF lp, where bs="bs"] |
|
4614 obtain z where z_def: "\<forall>x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast |
|
4615 from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp |
|
4616 moreover have "z + 1 > z" by simp |
|
4617 ultimately show ?thesis using z_def by auto |
|
4618 qed |
|
4619 |
|
4620 consts |
|
4621 \<Upsilon>:: "fm \<Rightarrow> (num \<times> int) list" |
|
4622 \<upsilon> :: "fm \<Rightarrow> (num \<times> int) \<Rightarrow> fm " |
|
4623 recdef \<Upsilon> "measure size" |
|
4624 "\<Upsilon> (And p q) = (\<Upsilon> p @ \<Upsilon> q)" |
|
4625 "\<Upsilon> (Or p q) = (\<Upsilon> p @ \<Upsilon> q)" |
|
4626 "\<Upsilon> (Eq (CN 0 c e)) = [(Neg e,c)]" |
|
4627 "\<Upsilon> (NEq (CN 0 c e)) = [(Neg e,c)]" |
|
4628 "\<Upsilon> (Lt (CN 0 c e)) = [(Neg e,c)]" |
|
4629 "\<Upsilon> (Le (CN 0 c e)) = [(Neg e,c)]" |
|
4630 "\<Upsilon> (Gt (CN 0 c e)) = [(Neg e,c)]" |
|
4631 "\<Upsilon> (Ge (CN 0 c e)) = [(Neg e,c)]" |
|
4632 "\<Upsilon> p = []" |
|
4633 |
|
4634 recdef \<upsilon> "measure size" |
|
4635 "\<upsilon> (And p q) = (\<lambda> (t,n). And (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))" |
|
4636 "\<upsilon> (Or p q) = (\<lambda> (t,n). Or (\<upsilon> p (t,n)) (\<upsilon> q (t,n)))" |
|
4637 "\<upsilon> (Eq (CN 0 c e)) = (\<lambda> (t,n). Eq (Add (Mul c t) (Mul n e)))" |
|
4638 "\<upsilon> (NEq (CN 0 c e)) = (\<lambda> (t,n). NEq (Add (Mul c t) (Mul n e)))" |
|
4639 "\<upsilon> (Lt (CN 0 c e)) = (\<lambda> (t,n). Lt (Add (Mul c t) (Mul n e)))" |
|
4640 "\<upsilon> (Le (CN 0 c e)) = (\<lambda> (t,n). Le (Add (Mul c t) (Mul n e)))" |
|
4641 "\<upsilon> (Gt (CN 0 c e)) = (\<lambda> (t,n). Gt (Add (Mul c t) (Mul n e)))" |
|
4642 "\<upsilon> (Ge (CN 0 c e)) = (\<lambda> (t,n). Ge (Add (Mul c t) (Mul n e)))" |
|
4643 "\<upsilon> p = (\<lambda> (t,n). p)" |
|
4644 |
|
4645 lemma \<upsilon>_I: assumes lp: "isrlfm p" |
|
4646 and np: "real n > 0" and nbt: "numbound0 t" |
|
4647 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> _") |
|
4648 using lp |
|
4649 proof(induct p rule: \<upsilon>.induct) |
|
4650 case (5 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ |
|
4651 have "?I ?u (Lt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) < 0)" |
|
4652 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp |
|
4653 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) < 0)" |
|
4654 by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" |
|
4655 and b="0", simplified divide_zero_left]) (simp only: algebra_simps) |
|
4656 also have "\<dots> = (real c *?t + ?n* (?N x e) < 0)" |
|
4657 using np by simp |
|
4658 finally show ?case using nbt nb by (simp add: algebra_simps) |
|
4659 next |
|
4660 case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ |
|
4661 have "?I ?u (Le (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<le> 0)" |
|
4662 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp |
|
4663 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<le> 0)" |
|
4664 by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" |
|
4665 and b="0", simplified divide_zero_left]) (simp only: algebra_simps) |
|
4666 also have "\<dots> = (real c *?t + ?n* (?N x e) \<le> 0)" |
|
4667 using np by simp |
|
4668 finally show ?case using nbt nb by (simp add: algebra_simps) |
|
4669 next |
|
4670 case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ |
|
4671 have "?I ?u (Gt (CN 0 c e)) = (real c *(?t/?n) + (?N x e) > 0)" |
|
4672 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp |
|
4673 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) > 0)" |
|
4674 by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" |
|
4675 and b="0", simplified divide_zero_left]) (simp only: algebra_simps) |
|
4676 also have "\<dots> = (real c *?t + ?n* (?N x e) > 0)" |
|
4677 using np by simp |
|
4678 finally show ?case using nbt nb by (simp add: algebra_simps) |
|
4679 next |
|
4680 case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ |
|
4681 have "?I ?u (Ge (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<ge> 0)" |
|
4682 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp |
|
4683 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<ge> 0)" |
|
4684 by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" |
|
4685 and b="0", simplified divide_zero_left]) (simp only: algebra_simps) |
|
4686 also have "\<dots> = (real c *?t + ?n* (?N x e) \<ge> 0)" |
|
4687 using np by simp |
|
4688 finally show ?case using nbt nb by (simp add: algebra_simps) |
|
4689 next |
|
4690 case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ |
|
4691 from np have np: "real n \<noteq> 0" by simp |
|
4692 have "?I ?u (Eq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) = 0)" |
|
4693 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp |
|
4694 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) = 0)" |
|
4695 by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" |
|
4696 and b="0", simplified divide_zero_left]) (simp only: algebra_simps) |
|
4697 also have "\<dots> = (real c *?t + ?n* (?N x e) = 0)" |
|
4698 using np by simp |
|
4699 finally show ?case using nbt nb by (simp add: algebra_simps) |
|
4700 next |
|
4701 case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ |
|
4702 from np have np: "real n \<noteq> 0" by simp |
|
4703 have "?I ?u (NEq (CN 0 c e)) = (real c *(?t/?n) + (?N x e) \<noteq> 0)" |
|
4704 using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp |
|
4705 also have "\<dots> = (?n*(real c *(?t/?n)) + ?n*(?N x e) \<noteq> 0)" |
|
4706 by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" |
|
4707 and b="0", simplified divide_zero_left]) (simp only: algebra_simps) |
|
4708 also have "\<dots> = (real c *?t + ?n* (?N x e) \<noteq> 0)" |
|
4709 using np by simp |
|
4710 finally show ?case using nbt nb by (simp add: algebra_simps) |
|
4711 qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2) |
|
4712 |
|
4713 lemma \<Upsilon>_l: |
|
4714 assumes lp: "isrlfm p" |
|
4715 shows "\<forall> (t,k) \<in> set (\<Upsilon> p). numbound0 t \<and> k >0" |
|
4716 using lp |
|
4717 by(induct p rule: \<Upsilon>.induct) auto |
|
4718 |
|
4719 lemma rminusinf_\<Upsilon>: |
|
4720 assumes lp: "isrlfm p" |
|
4721 and nmi: "\<not> (Ifm (a#bs) (minusinf p))" (is "\<not> (Ifm (a#bs) (?M p))") |
|
4722 and ex: "Ifm (x#bs) p" (is "?I x p") |
|
4723 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") |
|
4724 proof- |
|
4725 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") |
|
4726 using lp nmi ex |
|
4727 by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) |
|
4728 then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<ge> ?N a s" by blast |
|
4729 from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto |
|
4730 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" |
|
4731 by (auto simp add: mult_commute) |
|
4732 thus ?thesis using smU by auto |
|
4733 qed |
|
4734 |
|
4735 lemma rplusinf_\<Upsilon>: |
|
4736 assumes lp: "isrlfm p" |
|
4737 and nmi: "\<not> (Ifm (a#bs) (plusinf p))" (is "\<not> (Ifm (a#bs) (?M p))") |
|
4738 and ex: "Ifm (x#bs) p" (is "?I x p") |
|
4739 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") |
|
4740 proof- |
|
4741 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") |
|
4742 using lp nmi ex |
|
4743 by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) |
|
4744 then obtain s m where smU: "(s,m) \<in> set (\<Upsilon> p)" and mx: "real m * x \<le> ?N a s" by blast |
|
4745 from \<Upsilon>_l[OF lp] smU have mp: "real m > 0" by auto |
|
4746 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" |
|
4747 by (auto simp add: mult_commute) |
|
4748 thus ?thesis using smU by auto |
|
4749 qed |
|
4750 |
|
4751 lemma lin_dense: |
|
4752 assumes lp: "isrlfm p" |
|
4753 and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (t,n). Inum (x#bs) t / real n) ` set (\<Upsilon> p)" |
|
4754 (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (t,n). ?N x t / real n ) ` (?U p)") |
|
4755 and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" |
|
4756 and ly: "l < y" and yu: "y < u" |
|
4757 shows "Ifm (y#bs) p" |
|
4758 using lp px noS |
|
4759 proof (induct p rule: isrlfm.induct) |
|
4760 case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ |
|
4761 from prems have "x * real c + ?N x e < 0" by (simp add: algebra_simps) |
|
4762 hence pxc: "x < (- ?N x e) / real c" |
|
4763 by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) |
|
4764 from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto |
|
4765 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto |
|
4766 hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto |
|
4767 moreover {assume y: "y < (-?N x e)/ real c" |
|
4768 hence "y * real c < - ?N x e" |
|
4769 by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) |
|
4770 hence "real c * y + ?N x e < 0" by (simp add: algebra_simps) |
|
4771 hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} |
|
4772 moreover {assume y: "y > (- ?N x e) / real c" |
|
4773 with yu have eu: "u > (- ?N x e) / real c" by auto |
|
4774 with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto) |
|
4775 with lx pxc have "False" by auto |
|
4776 hence ?case by simp } |
|
4777 ultimately show ?case by blast |
|
4778 next |
|
4779 case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp + |
|
4780 from prems have "x * real c + ?N x e \<le> 0" by (simp add: algebra_simps) |
|
4781 hence pxc: "x \<le> (- ?N x e) / real c" |
|
4782 by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) |
|
4783 from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto |
|
4784 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto |
|
4785 hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto |
|
4786 moreover {assume y: "y < (-?N x e)/ real c" |
|
4787 hence "y * real c < - ?N x e" |
|
4788 by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) |
|
4789 hence "real c * y + ?N x e < 0" by (simp add: algebra_simps) |
|
4790 hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} |
|
4791 moreover {assume y: "y > (- ?N x e) / real c" |
|
4792 with yu have eu: "u > (- ?N x e) / real c" by auto |
|
4793 with noSc ly yu have "(- ?N x e) / real c \<le> l" by (cases "(- ?N x e) / real c > l", auto) |
|
4794 with lx pxc have "False" by auto |
|
4795 hence ?case by simp } |
|
4796 ultimately show ?case by blast |
|
4797 next |
|
4798 case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ |
|
4799 from prems have "x * real c + ?N x e > 0" by (simp add: algebra_simps) |
|
4800 hence pxc: "x > (- ?N x e) / real c" |
|
4801 by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) |
|
4802 from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto |
|
4803 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto |
|
4804 hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto |
|
4805 moreover {assume y: "y > (-?N x e)/ real c" |
|
4806 hence "y * real c > - ?N x e" |
|
4807 by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) |
|
4808 hence "real c * y + ?N x e > 0" by (simp add: algebra_simps) |
|
4809 hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} |
|
4810 moreover {assume y: "y < (- ?N x e) / real c" |
|
4811 with ly have eu: "l < (- ?N x e) / real c" by auto |
|
4812 with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto) |
|
4813 with xu pxc have "False" by auto |
|
4814 hence ?case by simp } |
|
4815 ultimately show ?case by blast |
|
4816 next |
|
4817 case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ |
|
4818 from prems have "x * real c + ?N x e \<ge> 0" by (simp add: algebra_simps) |
|
4819 hence pxc: "x \<ge> (- ?N x e) / real c" |
|
4820 by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) |
|
4821 from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto |
|
4822 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto |
|
4823 hence "y < (- ?N x e) / real c \<or> y > (-?N x e) / real c" by auto |
|
4824 moreover {assume y: "y > (-?N x e)/ real c" |
|
4825 hence "y * real c > - ?N x e" |
|
4826 by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) |
|
4827 hence "real c * y + ?N x e > 0" by (simp add: algebra_simps) |
|
4828 hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} |
|
4829 moreover {assume y: "y < (- ?N x e) / real c" |
|
4830 with ly have eu: "l < (- ?N x e) / real c" by auto |
|
4831 with noSc ly yu have "(- ?N x e) / real c \<ge> u" by (cases "(- ?N x e) / real c > l", auto) |
|
4832 with xu pxc have "False" by auto |
|
4833 hence ?case by simp } |
|
4834 ultimately show ?case by blast |
|
4835 next |
|
4836 case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ |
|
4837 from cp have cnz: "real c \<noteq> 0" by simp |
|
4838 from prems have "x * real c + ?N x e = 0" by (simp add: algebra_simps) |
|
4839 hence pxc: "x = (- ?N x e) / real c" |
|
4840 by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) |
|
4841 from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto |
|
4842 with lx xu have yne: "x \<noteq> - ?N x e / real c" by auto |
|
4843 with pxc show ?case by simp |
|
4844 next |
|
4845 case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ |
|
4846 from cp have cnz: "real c \<noteq> 0" by simp |
|
4847 from prems have noSc:"\<forall> t. l < t \<and> t < u \<longrightarrow> t \<noteq> (- ?N x e) / real c" by auto |
|
4848 with ly yu have yne: "y \<noteq> - ?N x e / real c" by auto |
|
4849 hence "y* real c \<noteq> -?N x e" |
|
4850 by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp |
|
4851 hence "y* real c + ?N x e \<noteq> 0" by (simp add: algebra_simps) |
|
4852 thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] |
|
4853 by (simp add: algebra_simps) |
|
4854 qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"]) |
|
4855 |
|
4856 lemma finite_set_intervals: |
|
4857 assumes px: "P (x::real)" |
|
4858 and lx: "l \<le> x" and xu: "x \<le> u" |
|
4859 and linS: "l\<in> S" and uinS: "u \<in> S" |
|
4860 and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u" |
|
4861 shows "\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a \<le> x \<and> x \<le> b \<and> P x" |
|
4862 proof- |
|
4863 let ?Mx = "{y. y\<in> S \<and> y \<le> x}" |
|
4864 let ?xM = "{y. y\<in> S \<and> x \<le> y}" |
|
4865 let ?a = "Max ?Mx" |
|
4866 let ?b = "Min ?xM" |
|
4867 have MxS: "?Mx \<subseteq> S" by blast |
|
4868 hence fMx: "finite ?Mx" using fS finite_subset by auto |
|
4869 from lx linS have linMx: "l \<in> ?Mx" by blast |
|
4870 hence Mxne: "?Mx \<noteq> {}" by blast |
|
4871 have xMS: "?xM \<subseteq> S" by blast |
|
4872 hence fxM: "finite ?xM" using fS finite_subset by auto |
|
4873 from xu uinS have linxM: "u \<in> ?xM" by blast |
|
4874 hence xMne: "?xM \<noteq> {}" by blast |
|
4875 have ax:"?a \<le> x" using Mxne fMx by auto |
|
4876 have xb:"x \<le> ?b" using xMne fxM by auto |
|
4877 have "?a \<in> ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a \<in> S" using MxS by blast |
|
4878 have "?b \<in> ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b \<in> S" using xMS by blast |
|
4879 have noy:"\<forall> y. ?a < y \<and> y < ?b \<longrightarrow> y \<notin> S" |
|
4880 proof(clarsimp) |
|
4881 fix y |
|
4882 assume ay: "?a < y" and yb: "y < ?b" and yS: "y \<in> S" |
|
4883 from yS have "y\<in> ?Mx \<or> y\<in> ?xM" by auto |
|
4884 moreover {assume "y \<in> ?Mx" hence "y \<le> ?a" using Mxne fMx by auto with ay have "False" by simp} |
|
4885 moreover {assume "y \<in> ?xM" hence "y \<ge> ?b" using xMne fxM by auto with yb have "False" by simp} |
|
4886 ultimately show "False" by blast |
|
4887 qed |
|
4888 from ainS binS noy ax xb px show ?thesis by blast |
|
4889 qed |
|
4890 |
|
4891 lemma finite_set_intervals2: |
|
4892 assumes px: "P (x::real)" |
|
4893 and lx: "l \<le> x" and xu: "x \<le> u" |
|
4894 and linS: "l\<in> S" and uinS: "u \<in> S" |
|
4895 and fS:"finite S" and lS: "\<forall> x\<in> S. l \<le> x" and Su: "\<forall> x\<in> S. x \<le> u" |
|
4896 shows "(\<exists> s\<in> S. P s) \<or> (\<exists> a \<in> S. \<exists> b \<in> S. (\<forall> y. a < y \<and> y < b \<longrightarrow> y \<notin> S) \<and> a < x \<and> x < b \<and> P x)" |
|
4897 proof- |
|
4898 from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] |
|
4899 obtain a and b where |
|
4900 as: "a\<in> S" and bs: "b\<in> S" and noS:"\<forall>y. a < y \<and> y < b \<longrightarrow> y \<notin> S" and axb: "a \<le> x \<and> x \<le> b \<and> P x" by auto |
|
4901 from axb have "x= a \<or> x= b \<or> (a < x \<and> x < b)" by auto |
|
4902 thus ?thesis using px as bs noS by blast |
|
4903 qed |
|
4904 |
|
4905 lemma rinf_\<Upsilon>: |
|
4906 assumes lp: "isrlfm p" |
|
4907 and nmi: "\<not> (Ifm (x#bs) (minusinf p))" (is "\<not> (Ifm (x#bs) (?M p))") |
|
4908 and npi: "\<not> (Ifm (x#bs) (plusinf p))" (is "\<not> (Ifm (x#bs) (?P p))") |
|
4909 and ex: "\<exists> x. Ifm (x#bs) p" (is "\<exists> x. ?I x p") |
|
4910 shows "\<exists> (l,n) \<in> set (\<Upsilon> p). \<exists> (s,m) \<in> set (\<Upsilon> p). ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" |
|
4911 proof- |
|
4912 let ?N = "\<lambda> x t. Inum (x#bs) t" |
|
4913 let ?U = "set (\<Upsilon> p)" |
|
4914 from ex obtain a where pa: "?I a p" by blast |
|
4915 from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi |
|
4916 have nmi': "\<not> (?I a (?M p))" by simp |
|
4917 from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi |
|
4918 have npi': "\<not> (?I a (?P p))" by simp |
|
4919 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" |
|
4920 proof- |
|
4921 let ?M = "(\<lambda> (t,c). ?N a t / real c) ` ?U" |
|
4922 have fM: "finite ?M" by auto |
|
4923 from rminusinf_\<Upsilon>[OF lp nmi pa] rplusinf_\<Upsilon>[OF lp npi pa] |
|
4924 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 |
|
4925 then obtain "t" "n" "s" "m" where |
|
4926 tnU: "(t,n) \<in> ?U" and smU: "(s,m) \<in> ?U" |
|
4927 and xs1: "a \<le> ?N x s / real m" and tx1: "a \<ge> ?N x t / real n" by blast |
|
4928 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 |
|
4929 from tnU have Mne: "?M \<noteq> {}" by auto |
|
4930 hence Une: "?U \<noteq> {}" by simp |
|
4931 let ?l = "Min ?M" |
|
4932 let ?u = "Max ?M" |
|
4933 have linM: "?l \<in> ?M" using fM Mne by simp |
|
4934 have uinM: "?u \<in> ?M" using fM Mne by simp |
|
4935 have tnM: "?N a t / real n \<in> ?M" using tnU by auto |
|
4936 have smM: "?N a s / real m \<in> ?M" using smU by auto |
|
4937 have lM: "\<forall> t\<in> ?M. ?l \<le> t" using Mne fM by auto |
|
4938 have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto |
|
4939 have "?l \<le> ?N a t / real n" using tnM Mne by simp hence lx: "?l \<le> a" using tx by simp |
|
4940 have "?N a s / real m \<le> ?u" using smM Mne by simp hence xu: "a \<le> ?u" using xs by simp |
|
4941 from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu] |
|
4942 have "(\<exists> s\<in> ?M. ?I s p) \<or> |
|
4943 (\<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)" . |
|
4944 moreover { fix u assume um: "u\<in> ?M" and pu: "?I u p" |
|
4945 hence "\<exists> (tu,nu) \<in> ?U. u = ?N a tu / real nu" by auto |
|
4946 then obtain "tu" "nu" where tuU: "(tu,nu) \<in> ?U" and tuu:"u= ?N a tu / real nu" by blast |
|
4947 have "(u + u) / 2 = u" by auto with pu tuu |
|
4948 have "?I (((?N a tu / real nu) + (?N a tu / real nu)) / 2) p" by simp |
|
4949 with tuU have ?thesis by blast} |
|
4950 moreover{ |
|
4951 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" |
|
4952 then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M" |
|
4953 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" |
|
4954 by blast |
|
4955 from t1M have "\<exists> (t1u,t1n) \<in> ?U. t1 = ?N a t1u / real t1n" by auto |
|
4956 then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) \<in> ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast |
|
4957 from t2M have "\<exists> (t2u,t2n) \<in> ?U. t2 = ?N a t2u / real t2n" by auto |
|
4958 then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) \<in> ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast |
|
4959 from t1x xt2 have t1t2: "t1 < t2" by simp |
|
4960 let ?u = "(t1 + t2) / 2" |
|
4961 from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto |
|
4962 from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" . |
|
4963 with t1uU t2uU t1u t2u have ?thesis by blast} |
|
4964 ultimately show ?thesis by blast |
|
4965 qed |
|
4966 then obtain "l" "n" "s" "m" where lnU: "(l,n) \<in> ?U" and smU:"(s,m) \<in> ?U" |
|
4967 and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast |
|
4968 from lnU smU \<Upsilon>_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto |
|
4969 from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] |
|
4970 numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu |
|
4971 have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp |
|
4972 with lnU smU |
|
4973 show ?thesis by auto |
|
4974 qed |
|
4975 (* The Ferrante - Rackoff Theorem *) |
|
4976 |
|
4977 theorem fr_eq: |
|
4978 assumes lp: "isrlfm p" |
|
4979 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))" |
|
4980 (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D") |
|
4981 proof |
|
4982 assume px: "\<exists> x. ?I x p" |
|
4983 have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast |
|
4984 moreover {assume "?M \<or> ?P" hence "?D" by blast} |
|
4985 moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P" |
|
4986 from rinf_\<Upsilon>[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast} |
|
4987 ultimately show "?D" by blast |
|
4988 next |
|
4989 assume "?D" |
|
4990 moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} |
|
4991 moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } |
|
4992 moreover {assume f:"?F" hence "?E" by blast} |
|
4993 ultimately show "?E" by blast |
|
4994 qed |
|
4995 |
|
4996 |
|
4997 lemma fr_eq\<upsilon>: |
|
4998 assumes lp: "isrlfm p" |
|
4999 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))))" |
|
5000 (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D") |
|
5001 proof |
|
5002 assume px: "\<exists> x. ?I x p" |
|
5003 have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast |
|
5004 moreover {assume "?M \<or> ?P" hence "?D" by blast} |
|
5005 moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P" |
|
5006 let ?f ="\<lambda> (t,n). Inum (x#bs) t / real n" |
|
5007 let ?N = "\<lambda> t. Inum (x#bs) t" |
|
5008 {fix t n s m assume "(t,n)\<in> set (\<Upsilon> p)" and "(s,m) \<in> set (\<Upsilon> p)" |
|
5009 with \<Upsilon>_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" |
|
5010 by auto |
|
5011 let ?st = "Add (Mul m t) (Mul n s)" |
|
5012 from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" |
|
5013 by (simp add: mult_commute) |
|
5014 from tnb snb have st_nb: "numbound0 ?st" by simp |
|
5015 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" |
|
5016 using mnp mp np by (simp add: algebra_simps add_divide_distrib) |
|
5017 from \<upsilon>_I[OF lp mnp st_nb, where x="x" and bs="bs"] |
|
5018 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])} |
|
5019 with rinf_\<Upsilon>[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} |
|
5020 ultimately show "?D" by blast |
|
5021 next |
|
5022 assume "?D" |
|
5023 moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} |
|
5024 moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } |
|
5025 moreover {fix t k s l assume "(t,k) \<in> set (\<Upsilon> p)" and "(s,l) \<in> set (\<Upsilon> p)" |
|
5026 and px:"?I x (\<upsilon> p (Add (Mul l t) (Mul k s), 2*k*l))" |
|
5027 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 |
|
5028 let ?st = "Add (Mul l t) (Mul k s)" |
|
5029 from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0" |
|
5030 by (simp add: mult_commute) |
|
5031 from tnb snb have st_nb: "numbound0 ?st" by simp |
|
5032 from \<upsilon>_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} |
|
5033 ultimately show "?E" by blast |
|
5034 qed |
|
5035 |
|
5036 text{* The overall Part *} |
|
5037 |
|
5038 lemma real_ex_int_real01: |
|
5039 shows "(\<exists> (x::real). P x) = (\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))" |
|
5040 proof(auto) |
|
5041 fix x |
|
5042 assume Px: "P x" |
|
5043 let ?i = "floor x" |
|
5044 let ?u = "x - real ?i" |
|
5045 have "x = real ?i + ?u" by simp |
|
5046 hence "P (real ?i + ?u)" using Px by simp |
|
5047 moreover have "real ?i \<le> x" using real_of_int_floor_le by simp hence "0 \<le> ?u" by arith |
|
5048 moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith |
|
5049 ultimately show "(\<exists> (i::int) (u::real). 0\<le> u \<and> u< 1 \<and> P (real i + u))" by blast |
|
5050 qed |
|
5051 |
|
5052 consts exsplitnum :: "num \<Rightarrow> num" |
|
5053 exsplit :: "fm \<Rightarrow> fm" |
|
5054 recdef exsplitnum "measure size" |
|
5055 "exsplitnum (C c) = (C c)" |
|
5056 "exsplitnum (Bound 0) = Add (Bound 0) (Bound 1)" |
|
5057 "exsplitnum (Bound n) = Bound (n+1)" |
|
5058 "exsplitnum (Neg a) = Neg (exsplitnum a)" |
|
5059 "exsplitnum (Add a b) = Add (exsplitnum a) (exsplitnum b) " |
|
5060 "exsplitnum (Sub a b) = Sub (exsplitnum a) (exsplitnum b) " |
|
5061 "exsplitnum (Mul c a) = Mul c (exsplitnum a)" |
|
5062 "exsplitnum (Floor a) = Floor (exsplitnum a)" |
|
5063 "exsplitnum (CN 0 c a) = CN 0 c (Add (Mul c (Bound 1)) (exsplitnum a))" |
|
5064 "exsplitnum (CN n c a) = CN (n+1) c (exsplitnum a)" |
|
5065 "exsplitnum (CF c s t) = CF c (exsplitnum s) (exsplitnum t)" |
|
5066 |
|
5067 recdef exsplit "measure size" |
|
5068 "exsplit (Lt a) = Lt (exsplitnum a)" |
|
5069 "exsplit (Le a) = Le (exsplitnum a)" |
|
5070 "exsplit (Gt a) = Gt (exsplitnum a)" |
|
5071 "exsplit (Ge a) = Ge (exsplitnum a)" |
|
5072 "exsplit (Eq a) = Eq (exsplitnum a)" |
|
5073 "exsplit (NEq a) = NEq (exsplitnum a)" |
|
5074 "exsplit (Dvd i a) = Dvd i (exsplitnum a)" |
|
5075 "exsplit (NDvd i a) = NDvd i (exsplitnum a)" |
|
5076 "exsplit (And p q) = And (exsplit p) (exsplit q)" |
|
5077 "exsplit (Or p q) = Or (exsplit p) (exsplit q)" |
|
5078 "exsplit (Imp p q) = Imp (exsplit p) (exsplit q)" |
|
5079 "exsplit (Iff p q) = Iff (exsplit p) (exsplit q)" |
|
5080 "exsplit (NOT p) = NOT (exsplit p)" |
|
5081 "exsplit p = p" |
|
5082 |
|
5083 lemma exsplitnum: |
|
5084 "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t" |
|
5085 by(induct t rule: exsplitnum.induct) (simp_all add: algebra_simps) |
|
5086 |
|
5087 lemma exsplit: |
|
5088 assumes qfp: "qfree p" |
|
5089 shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p" |
|
5090 using qfp exsplitnum[where x="x" and y="y" and bs="bs"] |
|
5091 by(induct p rule: exsplit.induct) simp_all |
|
5092 |
|
5093 lemma splitex: |
|
5094 assumes qf: "qfree p" |
|
5095 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") |
|
5096 proof- |
|
5097 have "?rhs = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm (x#(real i)#bs) (exsplit p))" |
|
5098 by (simp add: myless[rule_format, where b="1"] myless[rule_format, where b="0"] add_ac diff_def) |
|
5099 also have "\<dots> = (\<exists> (i::int). \<exists> x. 0\<le> x \<and> x < 1 \<and> Ifm ((real i + x) #bs) p)" |
|
5100 by (simp only: exsplit[OF qf] add_ac) |
|
5101 also have "\<dots> = (\<exists> x. Ifm (x#bs) p)" |
|
5102 by (simp only: real_ex_int_real01[where P="\<lambda> x. Ifm (x#bs) p"]) |
|
5103 finally show ?thesis by simp |
|
5104 qed |
|
5105 |
|
5106 (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *) |
|
5107 |
|
5108 constdefs ferrack01:: "fm \<Rightarrow> fm" |
|
5109 "ferrack01 p \<equiv> (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p); |
|
5110 U = remdups(map simp_num_pair |
|
5111 (map (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) |
|
5112 (alluopairs (\<Upsilon> p')))) |
|
5113 in decr (evaldjf (\<upsilon> p') U ))" |
|
5114 |
|
5115 lemma fr_eq_01: |
|
5116 assumes qf: "qfree p" |
|
5117 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)))" |
|
5118 (is "(\<exists> x. ?I x ?q) = ?F") |
|
5119 proof- |
|
5120 let ?rq = "rlfm ?q" |
|
5121 let ?M = "?I x (minusinf ?rq)" |
|
5122 let ?P = "?I x (plusinf ?rq)" |
|
5123 have MF: "?M = False" |
|
5124 apply (simp add: Let_def reducecoeff_def numgcd_def zgcd_def rsplit_def ge_def lt_def conj_def disj_def) |
|
5125 by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all) |
|
5126 have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def zgcd_def rsplit_def ge_def lt_def conj_def disj_def) |
|
5127 by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C -1)))", simp_all) |
|
5128 have "(\<exists> x. ?I x ?q ) = |
|
5129 ((?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))))" |
|
5130 (is "(\<exists> x. ?I x ?q) = (?M \<or> ?P \<or> ?F)" is "?E = ?D") |
|
5131 proof |
|
5132 assume "\<exists> x. ?I x ?q" |
|
5133 then obtain x where qx: "?I x ?q" by blast |
|
5134 hence xp: "0\<le> x" and x1: "x< 1" and px: "?I x p" |
|
5135 by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf]) |
|
5136 from qx have "?I x ?rq " |
|
5137 by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) |
|
5138 hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto |
|
5139 from qf have qfq:"isrlfm ?rq" |
|
5140 by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) |
|
5141 with lqx fr_eq\<upsilon>[OF qfq] show "?M \<or> ?P \<or> ?F" by blast |
|
5142 next |
|
5143 assume D: "?D" |
|
5144 let ?U = "set (\<Upsilon> ?rq )" |
|
5145 from MF PF D have "?F" by auto |
|
5146 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 |
|
5147 from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] |
|
5148 by (auto simp add: rsplit_def lt_def ge_def) |
|
5149 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) |
|
5150 let ?st = "Add (Mul m t) (Mul n s)" |
|
5151 from tnb snb have stnb: "numbound0 ?st" by simp |
|
5152 from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" |
|
5153 by (simp add: mult_commute) |
|
5154 from conjunct1[OF \<upsilon>_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx |
|
5155 have "\<exists> x. ?I x ?rq" by auto |
|
5156 thus "?E" |
|
5157 using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def) |
|
5158 qed |
|
5159 with MF PF show ?thesis by blast |
|
5160 qed |
|
5161 |
|
5162 lemma \<Upsilon>_cong_aux: |
|
5163 assumes Ul: "\<forall> (t,n) \<in> set U. numbound0 t \<and> n >0" |
|
5164 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))" |
|
5165 (is "?lhs = ?rhs") |
|
5166 proof(auto) |
|
5167 fix t n s m |
|
5168 assume "((t,n),(s,m)) \<in> set (alluopairs U)" |
|
5169 hence th: "((t,n),(s,m)) \<in> (set U \<times> set U)" |
|
5170 using alluopairs_set1[where xs="U"] by blast |
|
5171 let ?N = "\<lambda> t. Inum (x#bs) t" |
|
5172 let ?st= "Add (Mul m t) (Mul n s)" |
|
5173 from Ul th have mnz: "m \<noteq> 0" by auto |
|
5174 from Ul th have nnz: "n \<noteq> 0" by auto |
|
5175 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" |
|
5176 using mnz nnz by (simp add: algebra_simps add_divide_distrib) |
|
5177 |
|
5178 thus "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) / |
|
5179 (2 * real n * real m) |
|
5180 \<in> (\<lambda>((t, n), s, m). |
|
5181 (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` |
|
5182 (set U \<times> set U)"using mnz nnz th |
|
5183 apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) |
|
5184 by (rule_tac x="(s,m)" in bexI,simp_all) |
|
5185 (rule_tac x="(t,n)" in bexI,simp_all) |
|
5186 next |
|
5187 fix t n s m |
|
5188 assume tnU: "(t,n) \<in> set U" and smU:"(s,m) \<in> set U" |
|
5189 let ?N = "\<lambda> t. Inum (x#bs) t" |
|
5190 let ?st= "Add (Mul m t) (Mul n s)" |
|
5191 from Ul smU have mnz: "m \<noteq> 0" by auto |
|
5192 from Ul tnU have nnz: "n \<noteq> 0" by auto |
|
5193 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" |
|
5194 using mnz nnz by (simp add: algebra_simps add_divide_distrib) |
|
5195 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" |
|
5196 have Pc:"\<forall> a b. ?P a b = ?P b a" |
|
5197 by auto |
|
5198 from Ul alluopairs_set1 have Up:"\<forall> ((t,n),(s,m)) \<in> set (alluopairs U). n \<noteq> 0 \<and> m \<noteq> 0" by blast |
|
5199 from alluopairs_ex[OF Pc, where xs="U"] tnU smU |
|
5200 have th':"\<exists> ((t',n'),(s',m')) \<in> set (alluopairs U). ?P (t',n') (s',m')" |
|
5201 by blast |
|
5202 then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \<in> set (alluopairs U)" |
|
5203 and Pts': "?P (t',n') (s',m')" by blast |
|
5204 from ts'_U Up have mnz': "m' \<noteq> 0" and nnz': "n'\<noteq> 0" by auto |
|
5205 let ?st' = "Add (Mul m' t') (Mul n' s')" |
|
5206 have st': "(?N t' / real n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')" |
|
5207 using mnz' nnz' by (simp add: algebra_simps add_divide_distrib) |
|
5208 from Pts' have |
|
5209 "(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 |
|
5210 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') |
|
5211 finally show "(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 |
|
5212 \<in> (\<lambda>(t, n). Inum (x # bs) t / real n) ` |
|
5213 (\<lambda>((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` |
|
5214 set (alluopairs U)" |
|
5215 using ts'_U by blast |
|
5216 qed |
|
5217 |
|
5218 lemma \<Upsilon>_cong: |
|
5219 assumes lp: "isrlfm p" |
|
5220 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)") |
|
5221 and U: "\<forall> (t,n) \<in> U. numbound0 t \<and> n > 0" |
|
5222 and U': "\<forall> (t,n) \<in> U'. numbound0 t \<and> n > 0" |
|
5223 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)))" |
|
5224 (is "?lhs = ?rhs") |
|
5225 proof |
|
5226 assume ?lhs |
|
5227 then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and |
|
5228 Pst: "Ifm (x#bs) (\<upsilon> p (Add (Mul m t) (Mul n s),2*n*m))" by blast |
|
5229 let ?N = "\<lambda> t. Inum (x#bs) t" |
|
5230 from tnU smU U have tnb: "numbound0 t" and np: "n > 0" |
|
5231 and snb: "numbound0 s" and mp:"m > 0" by auto |
|
5232 let ?st= "Add (Mul m t) (Mul n s)" |
|
5233 from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" |
|
5234 by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) |
|
5235 from tnb snb have stnb: "numbound0 ?st" by simp |
|
5236 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" |
|
5237 using mp np by (simp add: algebra_simps add_divide_distrib) |
|
5238 from tnU smU UU' have "?g ((t,n),(s,m)) \<in> ?f ` U'" by blast |
|
5239 hence "\<exists> (t',n') \<in> U'. ?g ((t,n),(s,m)) = ?f (t',n')" |
|
5240 by auto (rule_tac x="(a,b)" in bexI, auto) |
|
5241 then obtain t' n' where tnU': "(t',n') \<in> U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast |
|
5242 from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto |
|
5243 from \<upsilon>_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst |
|
5244 have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp |
|
5245 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]] |
|
5246 have "Ifm (x # bs) (\<upsilon> p (t', n')) " by (simp only: st) |
|
5247 then show ?rhs using tnU' by auto |
|
5248 next |
|
5249 assume ?rhs |
|
5250 then obtain t' n' where tnU': "(t',n') \<in> U'" and Pt': "Ifm (x # bs) (\<upsilon> p (t', n'))" |
|
5251 by blast |
|
5252 from tnU' UU' have "?f (t',n') \<in> ?g ` (U\<times>U)" by blast |
|
5253 hence "\<exists> ((t,n),(s,m)) \<in> (U\<times>U). ?f (t',n') = ?g ((t,n),(s,m))" |
|
5254 by auto (rule_tac x="(a,b)" in bexI, auto) |
|
5255 then obtain t n s m where tnU: "(t,n) \<in> U" and smU:"(s,m) \<in> U" and |
|
5256 th: "?f (t',n') = ?g((t,n),(s,m)) "by blast |
|
5257 let ?N = "\<lambda> t. Inum (x#bs) t" |
|
5258 from tnU smU U have tnb: "numbound0 t" and np: "n > 0" |
|
5259 and snb: "numbound0 s" and mp:"m > 0" by auto |
|
5260 let ?st= "Add (Mul m t) (Mul n s)" |
|
5261 from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" |
|
5262 by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) |
|
5263 from tnb snb have stnb: "numbound0 ?st" by simp |
|
5264 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" |
|
5265 using mp np by (simp add: algebra_simps add_divide_distrib) |
|
5266 from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto |
|
5267 from \<upsilon>_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' |
|
5268 have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp |
|
5269 with \<upsilon>_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast |
|
5270 qed |
|
5271 |
|
5272 lemma ferrack01: |
|
5273 assumes qf: "qfree p" |
|
5274 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> _") |
|
5275 proof- |
|
5276 let ?I = "\<lambda> x p. Ifm (x#bs) p" |
|
5277 fix x |
|
5278 let ?N = "\<lambda> t. Inum (x#bs) t" |
|
5279 let ?q = "rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)" |
|
5280 let ?U = "\<Upsilon> ?q" |
|
5281 let ?Up = "alluopairs ?U" |
|
5282 let ?g = "\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)" |
|
5283 let ?S = "map ?g ?Up" |
|
5284 let ?SS = "map simp_num_pair ?S" |
|
5285 let ?Y = "remdups ?SS" |
|
5286 let ?f= "(\<lambda> (t,n). ?N t / real n)" |
|
5287 let ?h = "\<lambda> ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2" |
|
5288 let ?F = "\<lambda> p. \<exists> a \<in> set (\<Upsilon> p). \<exists> b \<in> set (\<Upsilon> p). ?I x (\<upsilon> p (?g(a,b)))" |
|
5289 let ?ep = "evaldjf (\<upsilon> ?q) ?Y" |
|
5290 from rlfm_l[OF qf] have lq: "isrlfm ?q" |
|
5291 by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def zgcd_def) |
|
5292 from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \<le> (set ?U \<times> set ?U)" by simp |
|
5293 from \<Upsilon>_l[OF lq] have U_l: "\<forall> (t,n) \<in> set ?U. numbound0 t \<and> n > 0" . |
|
5294 from U_l UpU |
|
5295 have Up_: "\<forall> ((t,n),(s,m)) \<in> set ?Up. numbound0 t \<and> n> 0 \<and> numbound0 s \<and> m > 0" by auto |
|
5296 hence Snb: "\<forall> (t,n) \<in> set ?S. numbound0 t \<and> n > 0 " |
|
5297 by (auto simp add: mult_pos_pos) |
|
5298 have Y_l: "\<forall> (t,n) \<in> set ?Y. numbound0 t \<and> n > 0" |
|
5299 proof- |
|
5300 { fix t n assume tnY: "(t,n) \<in> set ?Y" |
|
5301 hence "(t,n) \<in> set ?SS" by simp |
|
5302 hence "\<exists> (t',n') \<in> set ?S. simp_num_pair (t',n') = (t,n)" |
|
5303 by (auto simp add: split_def) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all) |
|
5304 then obtain t' n' where tn'S: "(t',n') \<in> set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast |
|
5305 from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto |
|
5306 from simp_num_pair_l[OF tnb np tns] |
|
5307 have "numbound0 t \<and> n > 0" . } |
|
5308 thus ?thesis by blast |
|
5309 qed |
|
5310 |
|
5311 have YU: "(?f ` set ?Y) = (?h ` (set ?U \<times> set ?U))" |
|
5312 proof- |
|
5313 from simp_num_pair_ci[where bs="x#bs"] have |
|
5314 "\<forall>x. (?f o simp_num_pair) x = ?f x" by auto |
|
5315 hence th: "?f o simp_num_pair = ?f" using ext by blast |
|
5316 have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_compose) |
|
5317 also have "\<dots> = (?f ` set ?S)" by (simp add: th) |
|
5318 also have "\<dots> = ((?f o ?g) ` set ?Up)" |
|
5319 by (simp only: set_map o_def image_compose[symmetric]) |
|
5320 also have "\<dots> = (?h ` (set ?U \<times> set ?U))" |
|
5321 using \<Upsilon>_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast |
|
5322 finally show ?thesis . |
|
5323 qed |
|
5324 have "\<forall> (t,n) \<in> set ?Y. bound0 (\<upsilon> ?q (t,n))" |
|
5325 proof- |
|
5326 { fix t n assume tnY: "(t,n) \<in> set ?Y" |
|
5327 with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto |
|
5328 from \<upsilon>_I[OF lq np tnb] |
|
5329 have "bound0 (\<upsilon> ?q (t,n))" by simp} |
|
5330 thus ?thesis by blast |
|
5331 qed |
|
5332 hence ep_nb: "bound0 ?ep" using evaldjf_bound0[where xs="?Y" and f="\<upsilon> ?q"] |
|
5333 by auto |
|
5334 |
|
5335 from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q" |
|
5336 by (simp only: split_def fst_conv snd_conv) |
|
5337 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] |
|
5338 by (simp only: split_def fst_conv snd_conv) |
|
5339 also have "\<dots> = (Ifm (x#bs) ?ep)" |
|
5340 using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\<upsilon> ?q",symmetric] |
|
5341 by (simp only: split_def pair_collapse) |
|
5342 also have "\<dots> = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast |
|
5343 finally have lr: "?lhs = ?rhs" by (simp only: ferrack01_def Let_def) |
|
5344 from decr_qf[OF ep_nb] have "qfree (ferrack01 p)" by (simp only: Let_def ferrack01_def) |
|
5345 with lr show ?thesis by blast |
|
5346 qed |
|
5347 |
|
5348 lemma cp_thm': |
|
5349 assumes lp: "iszlfm p (real (i::int)#bs)" |
|
5350 and up: "d\<beta> p 1" and dd: "d\<delta> p d" and dp: "d > 0" |
|
5351 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))" |
|
5352 using cp_thm[OF lp up dd dp] by auto |
|
5353 |
|
5354 constdefs unit:: "fm \<Rightarrow> fm \<times> num list \<times> int" |
|
5355 "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; |
|
5356 B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q)) |
|
5357 in if length B \<le> length a then (q,B,d) else (mirror q, a,d))" |
|
5358 |
|
5359 lemma unit: assumes qf: "qfree p" |
|
5360 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)" |
|
5361 proof- |
|
5362 fix q B d |
|
5363 assume qBd: "unit p = (q,B,d)" |
|
5364 let ?thes = "((\<exists> (x::int). Ifm (real x#bs) p) = (\<exists> (x::int). Ifm (real x#bs) q)) \<and> |
|
5365 Inum (real i#bs) ` set B = Inum (real i#bs) ` set (\<beta> q) \<and> |
|
5366 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)" |
|
5367 let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p" |
|
5368 let ?p' = "zlfm p" |
|
5369 let ?l = "\<zeta> ?p'" |
|
5370 let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (a\<beta> ?p' ?l)" |
|
5371 let ?d = "\<delta> ?q" |
|
5372 let ?B = "set (\<beta> ?q)" |
|
5373 let ?B'= "remdups (map simpnum (\<beta> ?q))" |
|
5374 let ?A = "set (\<alpha> ?q)" |
|
5375 let ?A'= "remdups (map simpnum (\<alpha> ?q))" |
|
5376 from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] |
|
5377 have pp': "\<forall> i. ?I i ?p' = ?I i p" by auto |
|
5378 from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]] |
|
5379 have lp': "\<forall> (i::int). iszlfm ?p' (real i#bs)" by simp |
|
5380 hence lp'': "iszlfm ?p' (real (i::int)#bs)" by simp |
|
5381 from lp' \<zeta>[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d\<beta> ?p' ?l" by auto |
|
5382 from a\<beta>_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp' |
|
5383 have pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by (simp add: int_rdvd_iff) |
|
5384 from lp'' lp a\<beta>[OF lp'' dl lp] have lq:"iszlfm ?q (real i#bs)" and uq: "d\<beta> ?q 1" |
|
5385 by (auto simp add: isint_def) |
|
5386 from \<delta>[OF lq] have dp:"?d >0" and dd: "d\<delta> ?q ?d" by blast+ |
|
5387 let ?N = "\<lambda> t. Inum (real (i::int)#bs) t" |
|
5388 have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_compose) |
|
5389 also have "\<dots> = ?N ` ?B" using simpnum_ci[where bs="real i #bs"] by auto |
|
5390 finally have BB': "?N ` set ?B' = ?N ` ?B" . |
|
5391 have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_compose) |
|
5392 also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] by auto |
|
5393 finally have AA': "?N ` set ?A' = ?N ` ?A" . |
|
5394 from \<beta>_numbound0[OF lq] have B_nb:"\<forall> b\<in> set ?B'. numbound0 b" |
|
5395 by (simp add: simpnum_numbound0) |
|
5396 from \<alpha>_l[OF lq] have A_nb: "\<forall> b\<in> set ?A'. numbound0 b" |
|
5397 by (simp add: simpnum_numbound0) |
|
5398 {assume "length ?B' \<le> length ?A'" |
|
5399 hence q:"q=?q" and "B = ?B'" and d:"d = ?d" |
|
5400 using qBd by (auto simp add: Let_def unit_def) |
|
5401 with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<beta> q)" |
|
5402 and bn: "\<forall>b\<in> set B. numbound0 b" by simp+ |
|
5403 with pq_ex dp uq dd lq q d have ?thes by simp} |
|
5404 moreover |
|
5405 {assume "\<not> (length ?B' \<le> length ?A')" |
|
5406 hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" |
|
5407 using qBd by (auto simp add: Let_def unit_def) |
|
5408 with AA' mirror\<alpha>\<beta>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<beta> q)" |
|
5409 and bn: "\<forall>b\<in> set B. numbound0 b" by simp+ |
|
5410 from mirror_ex[OF lq] pq_ex q |
|
5411 have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp |
|
5412 from lq uq q mirror_d\<beta> [where p="?q" and bs="bs" and a="real i"] |
|
5413 have lq': "iszlfm q (real i#bs)" and uq: "d\<beta> q 1" by auto |
|
5414 from \<delta>[OF lq'] mirror_\<delta>[OF lq] q d have dq:"d\<delta> q d " by auto |
|
5415 from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp |
|
5416 } |
|
5417 ultimately show ?thes by blast |
|
5418 qed |
|
5419 (* Cooper's Algorithm *) |
|
5420 |
|
5421 constdefs cooper :: "fm \<Rightarrow> fm" |
|
5422 "cooper p \<equiv> |
|
5423 (let (q,B,d) = unit p; js = iupt (1,d); |
|
5424 mq = simpfm (minusinf q); |
|
5425 md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) js |
|
5426 in if md = T then T else |
|
5427 (let qd = evaldjf (\<lambda> t. simpfm (subst0 t q)) |
|
5428 (remdups (map (\<lambda> (b,j). simpnum (Add b (C j))) |
|
5429 [(b,j). b\<leftarrow>B,j\<leftarrow>js])) |
|
5430 in decr (disj md qd)))" |
|
5431 lemma cooper: assumes qf: "qfree p" |
|
5432 shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (cooper p))) \<and> qfree (cooper p)" |
|
5433 (is "(?lhs = ?rhs) \<and> _") |
|
5434 proof- |
|
5435 |
|
5436 let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p" |
|
5437 let ?q = "fst (unit p)" |
|
5438 let ?B = "fst (snd(unit p))" |
|
5439 let ?d = "snd (snd (unit p))" |
|
5440 let ?js = "iupt (1,?d)" |
|
5441 let ?mq = "minusinf ?q" |
|
5442 let ?smq = "simpfm ?mq" |
|
5443 let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js" |
|
5444 fix i |
|
5445 let ?N = "\<lambda> t. Inum (real (i::int)#bs) t" |
|
5446 let ?bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]" |
|
5447 let ?sbjs = "map (\<lambda> (b,j). simpnum (Add b (C j))) ?bjs" |
|
5448 let ?qd = "evaldjf (\<lambda> t. simpfm (subst0 t ?q)) (remdups ?sbjs)" |
|
5449 have qbf:"unit p = (?q,?B,?d)" by simp |
|
5450 from unit[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and |
|
5451 B:"?N ` set ?B = ?N ` set (\<beta> ?q)" and |
|
5452 uq:"d\<beta> ?q 1" and dd: "d\<delta> ?q ?d" and dp: "?d > 0" and |
|
5453 lq: "iszlfm ?q (real i#bs)" and |
|
5454 Bn: "\<forall> b\<in> set ?B. numbound0 b" by auto |
|
5455 from zlin_qfree[OF lq] have qfq: "qfree ?q" . |
|
5456 from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". |
|
5457 have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp |
|
5458 hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" |
|
5459 by (auto simp only: subst0_bound0[OF qfmq]) |
|
5460 hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))" |
|
5461 by (auto simp add: simpfm_bound0) |
|
5462 from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp |
|
5463 from Bn jsnb have "\<forall> (b,j) \<in> set ?bjs. numbound0 (Add b (C j))" |
|
5464 by simp |
|
5465 hence "\<forall> (b,j) \<in> set ?bjs. numbound0 (simpnum (Add b (C j)))" |
|
5466 using simpnum_numbound0 by blast |
|
5467 hence "\<forall> t \<in> set ?sbjs. numbound0 t" by simp |
|
5468 hence "\<forall> t \<in> set (remdups ?sbjs). bound0 (subst0 t ?q)" |
|
5469 using subst0_bound0[OF qfq] by auto |
|
5470 hence th': "\<forall> t \<in> set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))" |
|
5471 using simpfm_bound0 by blast |
|
5472 from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp |
|
5473 from mdb qdb |
|
5474 have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all) |
|
5475 from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B |
|
5476 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 |
|
5477 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))" apply (simp only: iupt_set simpfm) by auto |
|
5478 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 |
|
5479 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) |
|
5480 also have "\<dots>= ((\<exists> j\<in> set ?js. ?I j ?smq) \<or> (\<exists> t \<in> set ?sbjs. Ifm (?N t #bs) ?q))" |
|
5481 by (auto simp add: split_def) |
|
5482 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) |
|
5483 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) |
|
5484 finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by (simp add: disj) |
|
5485 hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp |
|
5486 {assume mdT: "?md = T" |
|
5487 hence cT:"cooper p = T" |
|
5488 by (simp only: cooper_def unit_def split_def Let_def if_True) simp |
|
5489 from mdT mdqd have lhs:"?lhs" by (auto simp add: disj) |
|
5490 from mdT have "?rhs" by (simp add: cooper_def unit_def split_def) |
|
5491 with lhs cT have ?thesis by simp } |
|
5492 moreover |
|
5493 {assume mdT: "?md \<noteq> T" hence "cooper p = decr (disj ?md ?qd)" |
|
5494 by (simp only: cooper_def unit_def split_def Let_def if_False) |
|
5495 with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } |
|
5496 ultimately show ?thesis by blast |
|
5497 qed |
|
5498 |
|
5499 lemma DJcooper: |
|
5500 assumes qf: "qfree p" |
|
5501 shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ cooper p))) \<and> qfree (DJ cooper p)" |
|
5502 proof- |
|
5503 from cooper have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (cooper p)" by blast |
|
5504 from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast |
|
5505 have "Ifm bs (DJ cooper p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (cooper q))" |
|
5506 by (simp add: DJ_def evaldjf_ex) |
|
5507 also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs) q)" |
|
5508 using cooper disjuncts_qf[OF qf] by blast |
|
5509 also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto) |
|
5510 finally show ?thesis using thqf by blast |
|
5511 qed |
|
5512 |
|
5513 (* Redy and Loveland *) |
|
5514 |
|
5515 lemma \<sigma>\<rho>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" |
|
5516 shows "Ifm (a#bs) (\<sigma>\<rho> p (t,c)) = Ifm (a#bs) (\<sigma>\<rho> p (t',c))" |
|
5517 using lp |
|
5518 by (induct p rule: iszlfm.induct, auto simp add: tt') |
|
5519 |
|
5520 lemma \<sigma>_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" |
|
5521 shows "Ifm (a#bs) (\<sigma> p c t) = Ifm (a#bs) (\<sigma> p c t')" |
|
5522 by (simp add: \<sigma>_def tt' \<sigma>\<rho>_cong[OF lp tt']) |
|
5523 |
|
5524 lemma \<rho>_cong: assumes lp: "iszlfm p (a#bs)" |
|
5525 and RR: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R = (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" |
|
5526 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))))" |
|
5527 (is "?lhs = ?rhs") |
|
5528 proof |
|
5529 let ?d = "\<delta> p" |
|
5530 assume ?lhs then obtain e c j where ecR: "(e,c) \<in> R" and jD:"j \<in> {1 .. c*?d}" |
|
5531 and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast |
|
5532 from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" by auto |
|
5533 hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" using RR by simp |
|
5534 hence "\<exists> (e',c') \<in> set (\<rho> p). Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto |
|
5535 then obtain e' c' where ecRo:"(e',c') \<in> set (\<rho> p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'" |
|
5536 and cc':"c = c'" by blast |
|
5537 from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp |
|
5538 |
|
5539 from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp |
|
5540 from ecRo jD px' cc' show ?rhs apply auto |
|
5541 by (rule_tac x="(e', c')" in bexI,simp_all) |
|
5542 (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) |
|
5543 next |
|
5544 let ?d = "\<delta> p" |
|
5545 assume ?rhs then obtain e c j where ecR: "(e,c) \<in> set (\<rho> p)" and jD:"j \<in> {1 .. c*?d}" |
|
5546 and px: "Ifm (a#bs) (\<sigma> p c (Add e (C j)))" (is "?sp c e j") by blast |
|
5547 from ecR have "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" by auto |
|
5548 hence "(Inum (a#bs) e,c) \<in> (\<lambda>(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp |
|
5549 hence "\<exists> (e',c') \<in> R. Inum (a#bs) e = Inum (a#bs) e' \<and> c = c'" by auto |
|
5550 then obtain e' c' where ecRo:"(e',c') \<in> R" and ee':"Inum (a#bs) e = Inum (a#bs) e'" |
|
5551 and cc':"c = c'" by blast |
|
5552 from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp |
|
5553 from \<sigma>_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp |
|
5554 from ecRo jD px' cc' show ?lhs apply auto |
|
5555 by (rule_tac x="(e', c')" in bexI,simp_all) |
|
5556 (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) |
|
5557 qed |
|
5558 |
|
5559 lemma rl_thm': |
|
5560 assumes lp: "iszlfm p (real (i::int)#bs)" |
|
5561 and R: "(\<lambda>(b,k). (Inum (a#bs) b,k)) ` R = (\<lambda>(b,k). (Inum (a#bs) b,k)) ` set (\<rho> p)" |
|
5562 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)))))" |
|
5563 using rl_thm[OF lp] \<rho>_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp |
|
5564 |
|
5565 constdefs chooset:: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int" |
|
5566 "chooset p \<equiv> (let q = zlfm p ; d = \<delta> q; |
|
5567 B = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<rho> q)) ; |
|
5568 a = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<alpha>\<rho> q)) |
|
5569 in if length B \<le> length a then (q,B,d) else (mirror q, a,d))" |
|
5570 |
|
5571 lemma chooset: assumes qf: "qfree p" |
|
5572 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)" |
|
5573 proof- |
|
5574 fix q B d |
|
5575 assume qBd: "chooset p = (q,B,d)" |
|
5576 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)" |
|
5577 let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p" |
|
5578 let ?q = "zlfm p" |
|
5579 let ?d = "\<delta> ?q" |
|
5580 let ?B = "set (\<rho> ?q)" |
|
5581 let ?f = "\<lambda> (t,k). (simpnum t,k)" |
|
5582 let ?B'= "remdups (map ?f (\<rho> ?q))" |
|
5583 let ?A = "set (\<alpha>\<rho> ?q)" |
|
5584 let ?A'= "remdups (map ?f (\<alpha>\<rho> ?q))" |
|
5585 from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] |
|
5586 have pp': "\<forall> i. ?I i ?q = ?I i p" by auto |
|
5587 hence pq_ex:"(\<exists> (x::int). ?I x p) = (\<exists> x. ?I x ?q)" by simp |
|
5588 from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real i"] |
|
5589 have lq: "iszlfm ?q (real (i::int)#bs)" . |
|
5590 from \<delta>[OF lq] have dp:"?d >0" by blast |
|
5591 let ?N = "\<lambda> (t,c). (Inum (real (i::int)#bs) t,c)" |
|
5592 have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_compose) |
|
5593 also have "\<dots> = ?N ` ?B" |
|
5594 by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def) |
|
5595 finally have BB': "?N ` set ?B' = ?N ` ?B" . |
|
5596 have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_compose) |
|
5597 also have "\<dots> = ?N ` ?A" using simpnum_ci[where bs="real i #bs"] |
|
5598 by(simp add: split_def image_compose simpnum_ci[where bs="real i #bs"] image_def) |
|
5599 finally have AA': "?N ` set ?A' = ?N ` ?A" . |
|
5600 from \<rho>_l[OF lq] have B_nb:"\<forall> (e,c)\<in> set ?B'. numbound0 e \<and> c > 0" |
|
5601 by (simp add: simpnum_numbound0 split_def) |
|
5602 from \<alpha>\<rho>_l[OF lq] have A_nb: "\<forall> (e,c)\<in> set ?A'. numbound0 e \<and> c > 0" |
|
5603 by (simp add: simpnum_numbound0 split_def) |
|
5604 {assume "length ?B' \<le> length ?A'" |
|
5605 hence q:"q=?q" and "B = ?B'" and d:"d = ?d" |
|
5606 using qBd by (auto simp add: Let_def chooset_def) |
|
5607 with BB' B_nb have b: "?N ` (set B) = ?N ` set (\<rho> q)" |
|
5608 and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto |
|
5609 with pq_ex dp lq q d have ?thes by simp} |
|
5610 moreover |
|
5611 {assume "\<not> (length ?B' \<le> length ?A')" |
|
5612 hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" |
|
5613 using qBd by (auto simp add: Let_def chooset_def) |
|
5614 with AA' mirror_\<alpha>\<rho>[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\<rho> q)" |
|
5615 and bn: "\<forall>(e,c)\<in> set B. numbound0 e \<and> c > 0" by auto |
|
5616 from mirror_ex[OF lq] pq_ex q |
|
5617 have pqm_eq:"(\<exists> (x::int). ?I x p) = (\<exists> (x::int). ?I x q)" by simp |
|
5618 from lq q mirror_l [where p="?q" and bs="bs" and a="real i"] |
|
5619 have lq': "iszlfm q (real i#bs)" by auto |
|
5620 from mirror_\<delta>[OF lq] pqm_eq b bn lq' dp q dp d have ?thes by simp |
|
5621 } |
|
5622 ultimately show ?thes by blast |
|
5623 qed |
|
5624 |
|
5625 constdefs stage:: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm" |
|
5626 "stage p d \<equiv> (\<lambda> (e,c). evaldjf (\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))) (iupt (1,c*d)))" |
|
5627 lemma stage: |
|
5628 shows "Ifm bs (stage p d (e,c)) = (\<exists> j\<in>{1 .. c*d}. Ifm bs (\<sigma> p c (Add e (C j))))" |
|
5629 by (unfold stage_def split_def ,simp only: evaldjf_ex iupt_set simpfm) simp |
|
5630 |
|
5631 lemma stage_nb: assumes lp: "iszlfm p (a#bs)" and cp: "c >0" and nb:"numbound0 e" |
|
5632 shows "bound0 (stage p d (e,c))" |
|
5633 proof- |
|
5634 let ?f = "\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))" |
|
5635 have th: "\<forall> j\<in> set (iupt(1,c*d)). bound0 (?f j)" |
|
5636 proof |
|
5637 fix j |
|
5638 from nb have nb':"numbound0 (Add e (C j))" by simp |
|
5639 from simpfm_bound0[OF \<sigma>_nb[OF lp nb', where k="c"]] |
|
5640 show "bound0 (simpfm (\<sigma> p c (Add e (C j))))" . |
|
5641 qed |
|
5642 from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp |
|
5643 qed |
|
5644 |
|
5645 constdefs redlove:: "fm \<Rightarrow> fm" |
|
5646 "redlove p \<equiv> |
|
5647 (let (q,B,d) = chooset p; |
|
5648 mq = simpfm (minusinf q); |
|
5649 md = evaldjf (\<lambda> j. simpfm (subst0 (C j) mq)) (iupt (1,d)) |
|
5650 in if md = T then T else |
|
5651 (let qd = evaldjf (stage q d) B |
|
5652 in decr (disj md qd)))" |
|
5653 |
|
5654 lemma redlove: assumes qf: "qfree p" |
|
5655 shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (redlove p))) \<and> qfree (redlove p)" |
|
5656 (is "(?lhs = ?rhs) \<and> _") |
|
5657 proof- |
|
5658 |
|
5659 let ?I = "\<lambda> (x::int) p. Ifm (real x#bs) p" |
|
5660 let ?q = "fst (chooset p)" |
|
5661 let ?B = "fst (snd(chooset p))" |
|
5662 let ?d = "snd (snd (chooset p))" |
|
5663 let ?js = "iupt (1,?d)" |
|
5664 let ?mq = "minusinf ?q" |
|
5665 let ?smq = "simpfm ?mq" |
|
5666 let ?md = "evaldjf (\<lambda> j. simpfm (subst0 (C j) ?smq)) ?js" |
|
5667 fix i |
|
5668 let ?N = "\<lambda> (t,k). (Inum (real (i::int)#bs) t,k)" |
|
5669 let ?qd = "evaldjf (stage ?q ?d) ?B" |
|
5670 have qbf:"chooset p = (?q,?B,?d)" by simp |
|
5671 from chooset[OF qf qbf] have pq_ex: "(\<exists>(x::int). ?I x p) = (\<exists> (x::int). ?I x ?q)" and |
|
5672 B:"?N ` set ?B = ?N ` set (\<rho> ?q)" and dd: "\<delta> ?q = ?d" and dp: "?d > 0" and |
|
5673 lq: "iszlfm ?q (real i#bs)" and |
|
5674 Bn: "\<forall> (e,c)\<in> set ?B. numbound0 e \<and> c > 0" by auto |
|
5675 from zlin_qfree[OF lq] have qfq: "qfree ?q" . |
|
5676 from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". |
|
5677 have jsnb: "\<forall> j \<in> set ?js. numbound0 (C j)" by simp |
|
5678 hence "\<forall> j\<in> set ?js. bound0 (subst0 (C j) ?smq)" |
|
5679 by (auto simp only: subst0_bound0[OF qfmq]) |
|
5680 hence th: "\<forall> j\<in> set ?js. bound0 (simpfm (subst0 (C j) ?smq))" |
|
5681 by (auto simp add: simpfm_bound0) |
|
5682 from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp |
|
5683 from Bn stage_nb[OF lq] have th:"\<forall> x \<in> set ?B. bound0 (stage ?q ?d x)" by auto |
|
5684 from evaldjf_bound0[OF th] have qdb: "bound0 ?qd" . |
|
5685 from mdb qdb |
|
5686 have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \<or> ?qd=T", simp_all) |
|
5687 from trans [OF pq_ex rl_thm'[OF lq B]] dd |
|
5688 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 |
|
5689 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) )))" |
|
5690 by (simp add: simpfm stage split_def) |
|
5691 also have "\<dots> = ((\<exists> j\<in> {1 .. ?d}. ?I i (subst0 (C j) ?smq)) \<or> ?I i ?qd)" |
|
5692 by (simp add: evaldjf_ex subst0_I[OF qfmq]) |
|
5693 finally have mdqd:"?lhs = (?I i ?md \<or> ?I i ?qd)" by (simp only: evaldjf_ex iupt_set simpfm) |
|
5694 also have "\<dots> = (?I i (disj ?md ?qd))" by (simp add: disj) |
|
5695 also have "\<dots> = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) |
|
5696 finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . |
|
5697 {assume mdT: "?md = T" |
|
5698 hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def) |
|
5699 from mdT have lhs:"?lhs" using mdqd by simp |
|
5700 from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def) |
|
5701 with lhs cT have ?thesis by simp } |
|
5702 moreover |
|
5703 {assume mdT: "?md \<noteq> T" hence "redlove p = decr (disj ?md ?qd)" |
|
5704 by (simp add: redlove_def chooset_def split_def Let_def) |
|
5705 with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } |
|
5706 ultimately show ?thesis by blast |
|
5707 qed |
|
5708 |
|
5709 lemma DJredlove: |
|
5710 assumes qf: "qfree p" |
|
5711 shows "((\<exists> (x::int). Ifm (real x#bs) p) = (Ifm bs (DJ redlove p))) \<and> qfree (DJ redlove p)" |
|
5712 proof- |
|
5713 from redlove have cqf: "\<forall> p. qfree p \<longrightarrow> qfree (redlove p)" by blast |
|
5714 from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast |
|
5715 have "Ifm bs (DJ redlove p) = (\<exists> q\<in> set (disjuncts p). Ifm bs (redlove q))" |
|
5716 by (simp add: DJ_def evaldjf_ex) |
|
5717 also have "\<dots> = (\<exists> q \<in> set(disjuncts p). \<exists> (x::int). Ifm (real x#bs) q)" |
|
5718 using redlove disjuncts_qf[OF qf] by blast |
|
5719 also have "\<dots> = (\<exists> (x::int). Ifm (real x#bs) p)" by (induct p rule: disjuncts.induct, auto) |
|
5720 finally show ?thesis using thqf by blast |
|
5721 qed |
|
5722 |
|
5723 |
|
5724 lemma exsplit_qf: assumes qf: "qfree p" |
|
5725 shows "qfree (exsplit p)" |
|
5726 using qf by (induct p rule: exsplit.induct, auto) |
|
5727 |
|
5728 definition mircfr :: "fm \<Rightarrow> fm" where |
|
5729 "mircfr = DJ cooper o ferrack01 o simpfm o exsplit" |
|
5730 |
|
5731 definition mirlfr :: "fm \<Rightarrow> fm" where |
|
5732 "mirlfr = DJ redlove o ferrack01 o simpfm o exsplit" |
|
5733 |
|
5734 lemma mircfr: "\<forall> bs p. qfree p \<longrightarrow> qfree (mircfr p) \<and> Ifm bs (mircfr p) = Ifm bs (E p)" |
|
5735 proof(clarsimp simp del: Ifm.simps) |
|
5736 fix bs p |
|
5737 assume qf: "qfree p" |
|
5738 show "qfree (mircfr p)\<and>(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)") |
|
5739 proof- |
|
5740 let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" |
|
5741 have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)" |
|
5742 using splitex[OF qf] by simp |
|
5743 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+ |
|
5744 with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def) |
|
5745 qed |
|
5746 qed |
|
5747 |
|
5748 lemma mirlfr: "\<forall> bs p. qfree p \<longrightarrow> qfree(mirlfr p) \<and> Ifm bs (mirlfr p) = Ifm bs (E p)" |
|
5749 proof(clarsimp simp del: Ifm.simps) |
|
5750 fix bs p |
|
5751 assume qf: "qfree p" |
|
5752 show "qfree (mirlfr p)\<and>(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \<and> (?lhs = ?rhs)") |
|
5753 proof- |
|
5754 let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" |
|
5755 have "?rhs = (\<exists> (i::int). \<exists> x. Ifm (x#real i#bs) ?es)" |
|
5756 using splitex[OF qf] by simp |
|
5757 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+ |
|
5758 with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def) |
|
5759 qed |
|
5760 qed |
|
5761 |
|
5762 definition mircfrqe:: "fm \<Rightarrow> fm" where |
|
5763 "mircfrqe p = qelim (prep p) mircfr" |
|
5764 |
|
5765 definition mirlfrqe:: "fm \<Rightarrow> fm" where |
|
5766 "mirlfrqe p = qelim (prep p) mirlfr" |
|
5767 |
|
5768 theorem mircfrqe: "(Ifm bs (mircfrqe p) = Ifm bs p) \<and> qfree (mircfrqe p)" |
|
5769 using qelim_ci[OF mircfr] prep by (auto simp add: mircfrqe_def) |
|
5770 |
|
5771 theorem mirlfrqe: "(Ifm bs (mirlfrqe p) = Ifm bs p) \<and> qfree (mirlfrqe p)" |
|
5772 using qelim_ci[OF mirlfr] prep by (auto simp add: mirlfrqe_def) |
|
5773 |
|
5774 definition |
|
5775 "test1 (u\<Colon>unit) = mircfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))" |
|
5776 |
|
5777 definition |
|
5778 "test2 (u\<Colon>unit) = mircfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))" |
|
5779 |
|
5780 definition |
|
5781 "test3 (u\<Colon>unit) = mirlfrqe (A (And (Le (Sub (Floor (Bound 0)) (Bound 0))) (Le (Add (Bound 0) (Floor (Neg (Bound 0)))))))" |
|
5782 |
|
5783 definition |
|
5784 "test4 (u\<Colon>unit) = mirlfrqe (A (Iff (Eq (Add (Floor (Bound 0)) (Floor (Neg (Bound 0))))) (Eq (Sub (Floor (Bound 0)) (Bound 0)))))" |
|
5785 |
|
5786 definition |
|
5787 "test5 (u\<Colon>unit) = mircfrqe (A(E(And (Ge(Sub (Bound 1) (Bound 0))) (Eq (Add (Floor (Bound 1)) (Floor (Neg(Bound 0))))))))" |
|
5788 |
|
5789 ML {* @{code test1} () *} |
|
5790 ML {* @{code test2} () *} |
|
5791 ML {* @{code test3} () *} |
|
5792 ML {* @{code test4} () *} |
|
5793 ML {* @{code test5} () *} |
|
5794 |
|
5795 (*export_code mircfrqe mirlfrqe |
|
5796 in SML module_name Mir file "raw_mir.ML"*) |
|
5797 |
|
5798 oracle mirfr_oracle = {* fn (proofs, ct) => |
|
5799 let |
|
5800 |
|
5801 fun num_of_term vs (t as Free (xn, xT)) = (case AList.lookup (op =) vs t |
|
5802 of NONE => error "Variable not found in the list!" |
|
5803 | SOME n => @{code Bound} n) |
|
5804 | num_of_term vs @{term "real (0::int)"} = @{code C} 0 |
|
5805 | num_of_term vs @{term "real (1::int)"} = @{code C} 1 |
|
5806 | num_of_term vs @{term "0::real"} = @{code C} 0 |
|
5807 | num_of_term vs @{term "1::real"} = @{code C} 1 |
|
5808 | num_of_term vs (Bound i) = @{code Bound} i |
|
5809 | num_of_term vs (@{term "uminus :: real \<Rightarrow> real"} $ t') = @{code Neg} (num_of_term vs t') |
|
5810 | num_of_term vs (@{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) = |
|
5811 @{code Add} (num_of_term vs t1, num_of_term vs t2) |
|
5812 | num_of_term vs (@{term "op - :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) = |
|
5813 @{code Sub} (num_of_term vs t1, num_of_term vs t2) |
|
5814 | num_of_term vs (@{term "op * :: real \<Rightarrow> real \<Rightarrow> real"} $ t1 $ t2) = |
|
5815 (case (num_of_term vs t1) |
|
5816 of @{code C} i => @{code Mul} (i, num_of_term vs t2) |
|
5817 | _ => error "num_of_term: unsupported Multiplication") |
|
5818 | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "number_of :: int \<Rightarrow> int"} $ t')) = |
|
5819 @{code C} (HOLogic.dest_numeral t') |
|
5820 | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ t')) = |
|
5821 @{code Floor} (num_of_term vs t') |
|
5822 | num_of_term vs (@{term "real :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ t')) = |
|
5823 @{code Neg} (@{code Floor} (@{code Neg} (num_of_term vs t'))) |
|
5824 | num_of_term vs (@{term "number_of :: int \<Rightarrow> real"} $ t') = |
|
5825 @{code C} (HOLogic.dest_numeral t') |
|
5826 | num_of_term vs t = error ("num_of_term: unknown term " ^ Syntax.string_of_term @{context} t); |
|
5827 |
|
5828 fun fm_of_term vs @{term True} = @{code T} |
|
5829 | fm_of_term vs @{term False} = @{code F} |
|
5830 | fm_of_term vs (@{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = |
|
5831 @{code Lt} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) |
|
5832 | fm_of_term vs (@{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = |
|
5833 @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) |
|
5834 | fm_of_term vs (@{term "op = :: real \<Rightarrow> real \<Rightarrow> bool"} $ t1 $ t2) = |
|
5835 @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) |
|
5836 | fm_of_term vs (@{term "op rdvd"} $ (@{term "real :: int \<Rightarrow> real"} $ (@{term "number_of :: int \<Rightarrow> int"} $ t1)) $ t2) = |
|
5837 @{code Dvd} (HOLogic.dest_numeral t1, num_of_term vs t2) |
|
5838 | fm_of_term vs (@{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ t1 $ t2) = |
|
5839 @{code Iff} (fm_of_term vs t1, fm_of_term vs t2) |
|
5840 | fm_of_term vs (@{term "op &"} $ t1 $ t2) = |
|
5841 @{code And} (fm_of_term vs t1, fm_of_term vs t2) |
|
5842 | fm_of_term vs (@{term "op |"} $ t1 $ t2) = |
|
5843 @{code Or} (fm_of_term vs t1, fm_of_term vs t2) |
|
5844 | fm_of_term vs (@{term "op -->"} $ t1 $ t2) = |
|
5845 @{code Imp} (fm_of_term vs t1, fm_of_term vs t2) |
|
5846 | fm_of_term vs (@{term "Not"} $ t') = |
|
5847 @{code NOT} (fm_of_term vs t') |
|
5848 | fm_of_term vs (Const ("Ex", _) $ Abs (xn, xT, p)) = |
|
5849 @{code E} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) |
|
5850 | fm_of_term vs (Const ("All", _) $ Abs (xn, xT, p)) = |
|
5851 @{code A} (fm_of_term (map (fn (v, n) => (v, n + 1)) vs) p) |
|
5852 | fm_of_term vs t = error ("fm_of_term : unknown term " ^ Syntax.string_of_term @{context} t); |
|
5853 |
|
5854 fun term_of_num vs (@{code C} i) = @{term "real :: int \<Rightarrow> real"} $ HOLogic.mk_number HOLogic.intT i |
|
5855 | term_of_num vs (@{code Bound} n) = fst (the (find_first (fn (_, m) => n = m) vs)) |
|
5856 | term_of_num vs (@{code Neg} (@{code Floor} (@{code Neg} t'))) = |
|
5857 @{term "real :: int \<Rightarrow> real"} $ (@{term "ceiling :: real \<Rightarrow> int"} $ term_of_num vs t') |
|
5858 | term_of_num vs (@{code Neg} t') = @{term "uminus :: real \<Rightarrow> real"} $ term_of_num vs t' |
|
5859 | term_of_num vs (@{code Add} (t1, t2)) = @{term "op + :: real \<Rightarrow> real \<Rightarrow> real"} $ |
|
5860 term_of_num vs t1 $ term_of_num vs t2 |
|
5861 | term_of_num vs (@{code Sub} (t1, t2)) = @{term "op - :: real \<Rightarrow> real \<Rightarrow> real"} $ |
|
5862 term_of_num vs t1 $ term_of_num vs t2 |
|
5863 | term_of_num vs (@{code Mul} (i, t2)) = @{term "op * :: real \<Rightarrow> real \<Rightarrow> real"} $ |
|
5864 term_of_num vs (@{code C} i) $ term_of_num vs t2 |
|
5865 | term_of_num vs (@{code Floor} t) = @{term "real :: int \<Rightarrow> real"} $ (@{term "floor :: real \<Rightarrow> int"} $ term_of_num vs t) |
|
5866 | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t)) |
|
5867 | term_of_num vs (@{code CF} (c, t, s)) = term_of_num vs (@{code Add} (@{code Mul} (c, @{code Floor} t), s)); |
|
5868 |
|
5869 fun term_of_fm vs @{code T} = HOLogic.true_const |
|
5870 | term_of_fm vs @{code F} = HOLogic.false_const |
|
5871 | term_of_fm vs (@{code Lt} t) = |
|
5872 @{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"} |
|
5873 | term_of_fm vs (@{code Le} t) = |
|
5874 @{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"} |
|
5875 | term_of_fm vs (@{code Gt} t) = |
|
5876 @{term "op < :: real \<Rightarrow> real \<Rightarrow> bool"} $ @{term "0::real"} $ term_of_num vs t |
|
5877 | term_of_fm vs (@{code Ge} t) = |
|
5878 @{term "op \<le> :: real \<Rightarrow> real \<Rightarrow> bool"} $ @{term "0::real"} $ term_of_num vs t |
|
5879 | term_of_fm vs (@{code Eq} t) = |
|
5880 @{term "op = :: real \<Rightarrow> real \<Rightarrow> bool"} $ term_of_num vs t $ @{term "0::real"} |
|
5881 | term_of_fm vs (@{code NEq} t) = |
|
5882 term_of_fm vs (@{code NOT} (@{code Eq} t)) |
|
5883 | term_of_fm vs (@{code Dvd} (i, t)) = |
|
5884 @{term "op rdvd"} $ term_of_num vs (@{code C} i) $ term_of_num vs t |
|
5885 | term_of_fm vs (@{code NDvd} (i, t)) = |
|
5886 term_of_fm vs (@{code NOT} (@{code Dvd} (i, t))) |
|
5887 | term_of_fm vs (@{code NOT} t') = |
|
5888 HOLogic.Not $ term_of_fm vs t' |
|
5889 | term_of_fm vs (@{code And} (t1, t2)) = |
|
5890 HOLogic.conj $ term_of_fm vs t1 $ term_of_fm vs t2 |
|
5891 | term_of_fm vs (@{code Or} (t1, t2)) = |
|
5892 HOLogic.disj $ term_of_fm vs t1 $ term_of_fm vs t2 |
|
5893 | term_of_fm vs (@{code Imp} (t1, t2)) = |
|
5894 HOLogic.imp $ term_of_fm vs t1 $ term_of_fm vs t2 |
|
5895 | term_of_fm vs (@{code Iff} (t1, t2)) = |
|
5896 @{term "op = :: bool \<Rightarrow> bool \<Rightarrow> bool"} $ term_of_fm vs t1 $ term_of_fm vs t2; |
|
5897 |
|
5898 in |
|
5899 let |
|
5900 val thy = Thm.theory_of_cterm ct; |
|
5901 val t = Thm.term_of ct; |
|
5902 val fs = OldTerm.term_frees t; |
|
5903 val vs = fs ~~ (0 upto (length fs - 1)); |
|
5904 val qe = if proofs then @{code mirlfrqe} else @{code mircfrqe}; |
|
5905 val t' = (term_of_fm vs o qe o fm_of_term vs) t; |
|
5906 in (cterm_of thy o HOLogic.mk_Trueprop o HOLogic.mk_eq) (t, t') end |
|
5907 end; |
|
5908 *} |
|
5909 |
|
5910 use "mir_tac.ML" |
|
5911 setup "Mir_Tac.setup" |
|
5912 |
|
5913 lemma "ALL (x::real). (\<lfloor>x\<rfloor> = \<lceil>x\<rceil> = (x = real \<lfloor>x\<rfloor>))" |
|
5914 apply mir |
|
5915 done |
|
5916 |
|
5917 lemma "ALL (x::real). real (2::int)*x - (real (1::int)) < real \<lfloor>x\<rfloor> + real \<lceil>x\<rceil> \<and> real \<lfloor>x\<rfloor> + real \<lceil>x\<rceil> \<le> real (2::int)*x + (real (1::int))" |
|
5918 apply mir |
|
5919 done |
|
5920 |
|
5921 lemma "ALL (x::real). 2*\<lfloor>x\<rfloor> \<le> \<lfloor>2*x\<rfloor> \<and> \<lfloor>2*x\<rfloor> \<le> 2*\<lfloor>x+1\<rfloor>" |
|
5922 apply mir |
|
5923 done |
|
5924 |
|
5925 lemma "ALL (x::real). \<exists>y \<le> x. (\<lfloor>x\<rfloor> = \<lceil>y\<rceil>)" |
|
5926 apply mir |
|
5927 done |
|
5928 |
|
5929 lemma "ALL x y. \<lfloor>x\<rfloor> = \<lfloor>y\<rfloor> \<longrightarrow> 0 \<le> abs (y - x) \<and> abs (y - x) \<le> 1" |
|
5930 apply mir |
|
5931 done |
|
5932 |
|
5933 end |