diff -r f6b9f528c89c -r 6571c78c9667 src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Tue Nov 17 12:01:19 2015 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Tue Nov 17 12:32:08 2015 +0000 @@ -28,15 +28,15 @@ definition rdvd:: "real \ real \ bool" (infixl "rdvd" 50) where "x rdvd y \ (\k::int. y = x * real_of_int k)" -lemma int_rdvd_real: +lemma int_rdvd_real: "real_of_int (i::int) rdvd x = (i dvd (floor x) \ real_of_int (floor x) = x)" (is "?l = ?r") proof - assume "?l" + assume "?l" hence th: "\ k. x=real_of_int (i*k)" by (simp add: rdvd_def) hence th': "real_of_int (floor x) = x" by (auto simp del: of_int_mult) with th have "\ k. real_of_int (floor x) = real_of_int (i*k)" by simp hence "\ k. floor x = i*k" by presburger - thus ?r using th' by (simp add: dvd_def) + thus ?r using th' by (simp add: dvd_def) next assume "?r" hence "(i::int) dvd \x::real\" .. hence "\ k. real_of_int (floor x) = real_of_int (i*k)" @@ -55,7 +55,7 @@ by auto from iffD2[OF abs_dvd_iff] d2 have "(abs d) dvd (floor t)" by blast - with ti int_rdvd_real[symmetric] have "real_of_int (abs d) rdvd t" by blast + with ti int_rdvd_real[symmetric] have "real_of_int (abs d) rdvd t" by blast thus "abs (real_of_int d) rdvd t" by simp next assume "abs (real_of_int d) rdvd t" hence "real_of_int (abs d) rdvd t" by simp @@ -67,14 +67,14 @@ lemma rdvd_minus: "(real_of_int (d::int) rdvd t) = (real_of_int d rdvd -t)" apply (auto simp add: rdvd_def) - apply (rule_tac x="-k" in exI, simp) + apply (rule_tac x="-k" in exI, simp) apply (rule_tac x="-k" in exI, simp) done lemma rdvd_left_0_eq: "(0 rdvd t) = (t=0)" by (auto simp add: rdvd_def) -lemma rdvd_mult: +lemma rdvd_mult: assumes knz: "k\0" shows "(real_of_int (n::int) * real_of_int (k::int) rdvd x * real_of_int k) = (real_of_int n rdvd x)" using knz by (simp add: rdvd_def) @@ -83,7 +83,7 @@ (**** SHADOW SYNTAX AND SEMANTICS ****) (*********************************************************************************) -datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num +datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num | Mul int num | Floor num| CF int num num (* A size for num to make inductive proofs simpler*) @@ -132,17 +132,17 @@ proof- let ?I = "\ t. Inum bs t" assume ie: "isint e bs" - hence th: "real_of_int (floor (?I e)) = ?I e" by (simp add: isint_def) + hence th: "real_of_int (floor (?I e)) = ?I e" by (simp add: isint_def) have "real_of_int (floor (?I (Neg e))) = real_of_int (floor (- (real_of_int (floor (?I e)))))" by (simp add: th) also have "\ = - real_of_int (floor (?I e))" by simp finally show "isint (Neg e) bs" by (simp add: isint_def th) qed -lemma isint_sub: +lemma isint_sub: assumes ie: "isint e bs" shows "isint (Sub (C c) e) bs" proof- let ?I = "\ t. Inum bs t" - from ie have th: "real_of_int (floor (?I e)) = ?I e" by (simp add: isint_def) + from ie have th: "real_of_int (floor (?I e)) = ?I e" by (simp add: isint_def) have "real_of_int (floor (?I (Sub (C c) e))) = real_of_int (floor ((real_of_int (c -floor (?I e)))))" by (simp add: th) also have "\ = real_of_int (c- floor (?I e))" by simp finally show "isint (Sub (C c) e) bs" by (simp add: isint_def th) @@ -166,7 +166,7 @@ (* FORMULAE *) -datatype fm = +datatype fm = T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| Dvd int num| NDvd int num| NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm @@ -213,7 +213,7 @@ "prep (E F) = F" "prep (E (Or p q)) = Or (prep (E p)) (prep (E q))" "prep (E (Imp p q)) = Or (prep (E (NOT p))) (prep (E q))" - "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" + "prep (E (Iff p q)) = Or (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" "prep (E (NOT (And p q))) = Or (prep (E (NOT p))) (prep (E(NOT q)))" "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))" "prep (E (NOT (Iff p q))) = Or (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" @@ -241,10 +241,10 @@ fun qfree:: "fm \ bool" where "qfree (E p) = False" | "qfree (A p) = False" - | "qfree (NOT p) = qfree p" - | "qfree (And p q) = (qfree p \ qfree q)" - | "qfree (Or p q) = (qfree p \ qfree q)" - | "qfree (Imp p q) = (qfree p \ qfree q)" + | "qfree (NOT p) = qfree p" + | "qfree (And p q) = (qfree p \ qfree q)" + | "qfree (Or p q) = (qfree p \ qfree q)" + | "qfree (Imp p q) = (qfree p \ qfree q)" | "qfree (Iff p q) = (qfree p \ qfree q)" | "qfree p = True" @@ -255,20 +255,20 @@ | "numbound0 (CN n i a) = (n > 0 \ numbound0 a)" | "numbound0 (Neg a) = numbound0 a" | "numbound0 (Add a b) = (numbound0 a \ numbound0 b)" - | "numbound0 (Sub a b) = (numbound0 a \ numbound0 b)" + | "numbound0 (Sub a b) = (numbound0 a \ numbound0 b)" | "numbound0 (Mul i a) = numbound0 a" | "numbound0 (Floor a) = numbound0 a" - | "numbound0 (CF c a b) = (numbound0 a \ numbound0 b)" + | "numbound0 (CF c a b) = (numbound0 a \ numbound0 b)" lemma numbound0_I: assumes nb: "numbound0 a" shows "Inum (b#bs) a = Inum (b'#bs) a" using nb by (induct a) auto -lemma numbound0_gen: +lemma numbound0_gen: assumes nb: "numbound0 t" and ti: "isint t (x#bs)" shows "\ y. isint t (y#bs)" - using nb ti + using nb ti proof(clarify) fix y from numbound0_I[OF nb, where bs="bs" and b="y" and b'="x"] ti[simplified isint_def] @@ -308,7 +308,7 @@ | "numsubst0 t (CF i a b) = CF i (numsubst0 t a) (numsubst0 t b)" | "numsubst0 t (Neg a) = Neg (numsubst0 t a)" | "numsubst0 t (Add a b) = Add (numsubst0 t a) (numsubst0 t b)" - | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" + | "numsubst0 t (Sub a b) = Sub (numsubst0 t a) (numsubst0 t b)" | "numsubst0 t (Mul i a) = Mul i (numsubst0 t a)" | "numsubst0 t (Floor a) = Floor (numsubst0 t a)" @@ -358,7 +358,7 @@ | "decr (NEq a) = NEq (decrnum a)" | "decr (Dvd i a) = Dvd i (decrnum a)" | "decr (NDvd i a) = NDvd i (decrnum a)" -| "decr (NOT p) = NOT (decr p)" +| "decr (NOT p) = NOT (decr p)" | "decr (And p q) = And (decr p) (decr q)" | "decr (Or p q) = Or (decr p) (decr q)" | "decr (Imp p q) = Imp (decr p) (decr q)" @@ -404,20 +404,20 @@ definition djf:: "('a \ fm) \ 'a \ fm \ fm" where - "djf f p q = (if q=T then T else if q=F then f p else + "djf f p q = (if q=T then T else if q=F then f p else (let fp = f p in case fp of T \ T | F \ q | _ \ Or fp q))" definition evaldjf:: "('a \ fm) \ 'a list \ fm" where "evaldjf f ps = foldr (djf f) ps F" lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)" - by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) - (cases "f p", simp_all add: Let_def djf_def) + by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def) + (cases "f p", simp_all add: Let_def djf_def) lemma evaldjf_ex: "Ifm bs (evaldjf f ps) = (\ p \ set ps. Ifm bs (f p))" by (induct ps) (simp_all add: evaldjf_def djf_Or) -lemma evaldjf_bound0: +lemma evaldjf_bound0: assumes nb: "\ x\ set xs. bound0 (f x)" shows "bound0 (evaldjf f xs)" using nb @@ -427,7 +427,7 @@ apply auto done -lemma evaldjf_qf: +lemma evaldjf_qf: assumes nb: "\ x\ set xs. qfree (f x)" shows "qfree (evaldjf f xs)" using nb @@ -474,12 +474,12 @@ shows "Ifm bs (DJ f p) = Ifm bs (f p)" proof - have "Ifm bs (DJ f p) = (\ q \ set (disjuncts p). Ifm bs (f q))" - by (simp add: DJ_def evaldjf_ex) + by (simp add: DJ_def evaldjf_ex) also have "\ = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) finally show ?thesis . qed -lemma DJ_qf: assumes +lemma DJ_qf: assumes fqf: "\ p. qfree p \ qfree (f p)" shows "\p. qfree p \ qfree (DJ f p) " proof(clarify) @@ -487,7 +487,7 @@ have th: "DJ f p = evaldjf f (disjuncts p)" by (simp add: DJ_def) from disjuncts_qf[OF qf] have "\ q\ set (disjuncts p). qfree q" . with fqf have th':"\ q\ set (disjuncts p). qfree (f q)" by blast - + from evaldjf_qf[OF th'] th show "qfree (DJ f p)" by simp qed @@ -551,7 +551,7 @@ definition reducecoeff :: "num \ num" where "reducecoeff t = - (let g = numgcd t in + (let g = numgcd t in if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)" fun dvdnumcoeff:: "num \ int \ bool" where @@ -560,10 +560,10 @@ | "dvdnumcoeff (CF c s t) = (\ g. g dvd c \ (dvdnumcoeff t g))" | "dvdnumcoeff t = (\g. False)" -lemma dvdnumcoeff_trans: +lemma dvdnumcoeff_trans: assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'" shows "dvdnumcoeff t g" - using dgt' gdg + using dgt' gdg by (induct t rule: dvdnumcoeff.induct) (simp_all add: gdg dvd_trans[OF gdg]) declare dvd_trans [trans add] @@ -584,10 +584,10 @@ by (simp add: numgcd_def numgcdh_pos maxcoeff_pos) lemma reducecoeffh: - assumes gt: "dvdnumcoeff t g" and gp: "g > 0" + assumes gt: "dvdnumcoeff t g" and gp: "g > 0" shows "real_of_int g *(Inum bs (reducecoeffh t g)) = Inum bs t" using gt -proof(induct t rule: reducecoeffh.induct) +proof(induct t rule: reducecoeffh.induct) case (1 i) hence gd: "g dvd i" by simp from assms 1 show ?case by (simp add: real_of_int_div[OF gd]) next @@ -595,7 +595,7 @@ from assms 2 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps) next case (3 c s t) hence gd: "g dvd c" by simp - from assms 3 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps) + from assms 3 show ?case by (simp add: real_of_int_div[OF gd] algebra_simps) qed (auto simp add: numgcd_def gp) fun ismaxcoeff:: "num \ int \ bool" where @@ -614,7 +614,7 @@ have thh: "maxcoeff t \ max (abs c) (maxcoeff t)" by simp from ismaxcoeff_mono[OF H thh] show ?case by simp next - case (3 c t s) + case (3 c t s) hence H1:"ismaxcoeff s (maxcoeff s)" by auto have thh1: "maxcoeff s \ max \c\ (maxcoeff s)" by (simp add: max_def) from ismaxcoeff_mono[OF H1 thh1] show ?case by simp @@ -637,7 +637,7 @@ shows "dvdnumcoeff t (numgcdh t m)" using assms proof(induct t rule: numgcdh.induct) - case (2 n c t) + case (2 n c t) let ?g = "numgcdh t m" from 2 have th:"gcd c ?g > 1" by simp from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] @@ -651,11 +651,11 @@ have th': "gcd c ?g dvd ?g" by simp from dvdnumcoeff_trans[OF th' th] have ?case by simp hence ?case by simp } - moreover {assume "abs c > 1" and g0:"?g = 0" + moreover {assume "abs c > 1" and g0:"?g = 0" from numgcdh0[OF g0] have "m=0". with 2 g0 have ?case by simp } ultimately show ?case by blast next - case (3 c s t) + case (3 c s t) let ?g = "numgcdh t m" from 3 have th:"gcd c ?g > 1" by simp from zgcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] @@ -669,14 +669,14 @@ have th': "gcd c ?g dvd ?g" by simp from dvdnumcoeff_trans[OF th' th] have ?case by simp hence ?case by simp } - moreover {assume "abs c > 1" and g0:"?g = 0" + moreover {assume "abs c > 1" and g0:"?g = 0" from numgcdh0[OF g0] have "m=0". with 3 g0 have ?case by simp } ultimately show ?case by blast qed auto lemma dvdnumcoeff_aux2: assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) \ numgcd t > 0" - using assms + using assms proof (simp add: numgcd_def) let ?mc = "maxcoeff t" let ?g = "numgcdh t ?mc" @@ -691,12 +691,12 @@ let ?g = "numgcd t" have "?g \ 0" by (simp add: numgcd_pos) hence "?g = 0 \ ?g = 1 \ ?g > 1" by auto - moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} - moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} + moreover {assume "?g = 0" hence ?thesis by (simp add: numgcd0)} + moreover {assume "?g = 1" hence ?thesis by (simp add: reducecoeff_def)} moreover { assume g1:"?g > 1" from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff t ?g" and g0: "?g > 0" by blast+ - from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis - by (simp add: reducecoeff_def Let_def)} + from reducecoeffh[OF th1 g0, where bs="bs"] g1 have ?thesis + by (simp add: reducecoeff_def Let_def)} ultimately show ?thesis by blast qed @@ -709,15 +709,15 @@ consts numadd:: "num \ num \ num" recdef numadd "measure (\ (t,s). size t + size s)" "numadd (CN n1 c1 r1,CN n2 c2 r2) = - (if n1=n2 then + (if n1=n2 then (let c = c1 + c2 in (if c=0 then numadd(r1,r2) else CN n1 c (numadd (r1,r2)))) else if n1 \ n2 then CN n1 c1 (numadd (r1,CN n2 c2 r2)) else (CN n2 c2 (numadd (CN n1 c1 r1,r2))))" - "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))" - "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" - "numadd (CF c1 t1 r1,CF c2 t2 r2) = - (if t1 = t2 then + "numadd (CN n1 c1 r1,t) = CN n1 c1 (numadd (r1, t))" + "numadd (t,CN n2 c2 r2) = CN n2 c2 (numadd (t,r2))" + "numadd (CF c1 t1 r1,CF c2 t2 r2) = + (if t1 = t2 then (let c=c1+c2; s= numadd(r1,r2) in (if c=0 then s else CF c t1 s)) else if lex_bnd t1 t2 then CF c1 t1 (numadd(r1,CF c2 t2 r2)) else CF c2 t2 (numadd(CF c1 t1 r1,r2)))" @@ -775,7 +775,7 @@ lemma isint_CF: assumes si: "isint s bs" shows "isint (CF c t s) bs" proof- have cti: "isint (Mul c (Floor t)) bs" by (simp add: isint_Mul isint_Floor) - + have "?thesis = isint (Add (Mul c (Floor t)) s) bs" by (simp add: isint_def) also have "\" by (simp add: isint_add cti si) finally show ?thesis . @@ -783,11 +783,11 @@ fun split_int:: "num \ num \ num" where "split_int (C c) = (C 0, C c)" -| "split_int (CN n c b) = - (let (bv,bi) = split_int b +| "split_int (CN n c b) = + (let (bv,bi) = split_int b in (CN n c bv, bi))" -| "split_int (CF c a b) = - (let (bv,bi) = split_int b +| "split_int (CF c a b) = + (let (bv,bi) = split_int b in (bv, CF c a bi))" | "split_int a = (a,C 0)" @@ -801,7 +801,7 @@ from 2(2) have tibi: "ti = ?bi" by (simp add: Let_def split_def) from 2(2) b[symmetric] bii show ?case by (auto simp add: Let_def split_def) next - case (3 c a b tv ti) + case (3 c a b tv ti) let ?bv = "fst (split_int b)" let ?bi = "snd (split_int b)" have "split_int b = (?bv,?bi)" by simp @@ -817,8 +817,8 @@ definition numfloor:: "num \ num" where - "numfloor t = (let (tv,ti) = split_int t in - (case tv of C i \ numadd (tv,ti) + "numfloor t = (let (tv,ti) = split_int t in + (case tv of C i \ numadd (tv,ti) | _ \ numadd(CF 1 tv (C 0),ti)))" lemma numfloor[simp]: "Inum bs (numfloor t) = Inum bs (Floor t)" (is "?n t = ?N (Floor t)") @@ -827,17 +827,17 @@ let ?ti = "snd (split_int t)" have tvti:"split_int t = (?tv,?ti)" by simp {assume H: "\ v. ?tv \ C v" - hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" + hence th1: "?n t = ?N (Add (Floor ?tv) ?ti)" by (cases ?tv) (auto simp add: numfloor_def Let_def split_def) from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ - hence "?N (Floor t) = real_of_int (floor (?N (Add ?tv ?ti)))" by simp + hence "?N (Floor t) = real_of_int (floor (?N (Add ?tv ?ti)))" by simp also have "\ = real_of_int (floor (?N ?tv) + (floor (?N ?ti)))" by (simp,subst tii[simplified isint_iff, symmetric]) simp also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) finally have ?thesis using th1 by simp} - moreover {fix v assume H:"?tv = C v" + moreover {fix v assume H:"?tv = C v" from split_int[OF tvti] have "?N (Floor t) = ?N (Floor(Add ?tv ?ti))" and tii:"isint ?ti bs" by simp+ - hence "?N (Floor t) = real_of_int (floor (?N (Add ?tv ?ti)))" by simp + hence "?N (Floor t) = real_of_int (floor (?N (Add ?tv ?ti)))" by simp also have "\ = real_of_int (floor (?N ?tv) + (floor (?N ?ti)))" by (simp,subst tii[simplified isint_iff, symmetric]) simp also have "\ = ?N (Add (Floor ?tv) ?ti)" by (simp add: tii[simplified isint_iff]) @@ -906,7 +906,7 @@ with cnz have "max (abs c) (maxcoeff t) > 0" by arith with 2 show ?case by simp next - case (3 c s t) + case (3 c s t) hence cnz: "c \0" and mx: "max (abs c) (maxcoeff t) = 0" by simp+ have "max (abs c) (maxcoeff t) \ abs c" by simp with cnz have "max (abs c) (maxcoeff t) > 0" by arith @@ -922,10 +922,10 @@ definition simp_num_pair :: "(num \ int) \ num \ int" where "simp_num_pair \ (\ (t,n). (if n = 0 then (C 0, 0) else - (let t' = simpnum t ; g = numgcd t' in - if g > 1 then (let g' = gcd n g in - if g' = 1 then (t',n) - else (reducecoeffh t' g', n div g')) + (let t' = simpnum t ; g = numgcd t' in + if g > 1 then (let g' = gcd n g in + if g' = 1 then (t',n) + else (reducecoeffh t' g', n div g')) else (t',n))))" lemma simp_num_pair_ci: @@ -952,7 +952,7 @@ have gpdg: "?g' dvd ?g" by simp have gpdd: "?g' dvd n" by simp have gpdgp: "?g' dvd ?g'" by simp - from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] + from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] have th2:"real_of_int ?g' * ?t = Inum bs ?t'" by simp from nnz g1 g'1 have "?lhs = ?t / real_of_int (n div ?g')" by (simp add: simp_num_pair_def Let_def) also have "\ = (real_of_int ?g' * ?t) / (real_of_int ?g' * (real_of_int (n div ?g')))" by simp @@ -993,7 +993,7 @@ hence ?thesis using assms g1 g'1 by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0)} ultimately have ?thesis by blast } - ultimately have ?thesis by blast } + ultimately have ?thesis by blast } ultimately show ?thesis by blast qed @@ -1020,29 +1020,29 @@ by (induct p) auto definition conj :: "fm \ fm \ fm" where - "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else + "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else if p = q then p else And p q)" lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" by (cases "p=F \ q=F", simp_all add: conj_def) (cases p, simp_all) lemma conj_qf[simp]: "\qfree p ; qfree q\ \ qfree (conj p q)" - using conj_def by auto + using conj_def by auto lemma conj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" - using conj_def by auto + using conj_def by auto definition disj :: "fm \ fm \ fm" where - "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p + "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p else if p=q then p else Or p q)" lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)" by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) lemma disj_qf[simp]: "\qfree p ; qfree q\ \ qfree (disj p q)" - using disj_def by auto + using disj_def by auto lemma disj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" - using disj_def by auto + using disj_def by auto definition imp :: "fm \ fm \ fm" where - "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p + "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p else Imp p q)" lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" by (cases "p=F \ q=T",simp_all add: imp_def) @@ -1050,8 +1050,8 @@ using imp_def by (cases "p=F \ q=T",simp_all add: imp_def) definition iff :: "fm \ fm \ fm" where - "iff p q \ (if (p = q) then T else if (p = not q \ not p = q) then F else - 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 + "iff p q \ (if (p = q) then T else if (p = not q \ not p = q) then F else + 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 Iff p q)" lemma iff[simp]: "Ifm bs (iff p q) = Ifm bs (Iff p q)" @@ -1074,7 +1074,7 @@ lemma rdvd_left1_int: "real_of_int \t\ = t \ 1 rdvd t" by (simp add: rdvd_def,rule_tac x="\t\" in exI) simp -lemma rdvd_reduce: +lemma rdvd_reduce: assumes gd:"g dvd d" and gc:"g dvd c" and gp: "g > 0" shows "real_of_int (d::int) rdvd real_of_int (c::int)*t = (real_of_int (d div g) rdvd real_of_int (c div g)*t)" proof @@ -1095,13 +1095,13 @@ qed definition simpdvd :: "int \ num \ (int \ num)" where - "simpdvd d t \ - (let g = numgcd t in - if g > 1 then (let g' = gcd d g in - if g' = 1 then (d, t) - else (d div g',reducecoeffh t g')) + "simpdvd d t \ + (let g = numgcd t in + if g > 1 then (let g' = gcd d g in + if g' = 1 then (d, t) + else (d div g',reducecoeffh t g')) else (d, t))" -lemma simpdvd: +lemma simpdvd: assumes tnz: "nozerocoeff t" and dnz: "d \ 0" shows "Ifm bs (Dvd (fst (simpdvd d t)) (snd (simpdvd d t))) = Ifm bs (Dvd d t)" proof- @@ -1121,13 +1121,13 @@ have gpdg: "?g' dvd ?g" by simp have gpdd: "?g' dvd d" by simp have gpdgp: "?g' dvd ?g'" by simp - from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] + from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] have th2:"real_of_int ?g' * ?t = Inum bs t" by simp from assms g1 g0 g'1 have "Ifm bs (Dvd (fst (simpdvd d t)) (snd(simpdvd d t))) = Ifm bs (Dvd (d div ?g') ?tt)" by (simp add: simpdvd_def Let_def) also have "\ = (real_of_int d rdvd (Inum bs t))" - using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified div_self[OF gp0]] + using rdvd_reduce[OF gpdd gpdgp g'p, where t="?t", simplified div_self[OF gp0]] th2[symmetric] by simp finally have ?thesis by simp } ultimately have ?thesis by blast @@ -1141,7 +1141,7 @@ | "simpfm (Imp p q) = imp (simpfm p) (simpfm q)" | "simpfm (Iff p q) = iff (simpfm p) (simpfm q)" | "simpfm (NOT p) = not (simpfm p)" -| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \ if (v < 0) then T else F +| "simpfm (Lt a) = (let a' = simpnum a in case a' of C v \ if (v < 0) then T else F | _ \ Lt (reducecoeff a'))" | "simpfm (Le a) = (let a' = simpnum a in case a' of C v \ if (v \ 0) then T else F | _ \ Le (reducecoeff a'))" | "simpfm (Gt a) = (let a' = simpnum a in case a' of C v \ if (v > 0) then T else F | _ \ Gt (reducecoeff a'))" @@ -1151,7 +1151,7 @@ | "simpfm (Dvd i a) = (if i=0 then simpfm (Eq a) else if (abs i = 1) \ check_int a then T else let a' = simpnum a in case a' of C v \ if (i dvd v) then T else F | _ \ (let (d,t) = simpdvd i a' in Dvd d t))" -| "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) +| "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) else if (abs i = 1) \ check_int a then F else let a' = simpnum a in case a' of C v \ if (\(i dvd v)) then T else F | _ \ (let (d,t) = simpdvd i a' in NDvd d t))" | "simpfm p = p" @@ -1265,22 +1265,22 @@ case (12 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp have "i=0 \ (abs i = 1 \ check_int a) \ (i\0 \ ((abs i \ 1) \ (\ check_int a)))" by auto {assume "i=0" hence ?case using "12.hyps" by (simp add: rdvd_left_0_eq Let_def)} - moreover - {assume ai1: "abs i = 1" and ai: "check_int a" + moreover + {assume ai1: "abs i = 1" and ai: "check_int a" hence "i=1 \ i= - 1" by arith - moreover {assume i1: "i = 1" - from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] + moreover {assume i1: "i = 1" + from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] have ?case using i1 ai by simp } - moreover {assume i1: "i = - 1" - from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] + moreover {assume i1: "i = - 1" + from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] rdvd_abs1[where d="- 1" and t="Inum bs a"] have ?case using i1 ai by simp } ultimately have ?case by blast} - moreover + moreover {assume inz: "i\0" and cond: "(abs i \ 1) \ (\ check_int a)" {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond by (cases "abs i = 1", auto simp add: int_rdvd_iff) } - moreover {assume H:"\ (\ v. ?sa = C v)" + moreover {assume H:"\ (\ v. ?sa = C v)" 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) from simpnum_nz have nz:"nozerocoeff ?sa" by simp from simpdvd [OF nz inz] th have ?case using sa by simp} @@ -1290,23 +1290,23 @@ case (13 i a) let ?sa = "simpnum a" have sa: "Inum bs ?sa = Inum bs a" by simp have "i=0 \ (abs i = 1 \ check_int a) \ (i\0 \ ((abs i \ 1) \ (\ check_int a)))" by auto {assume "i=0" hence ?case using "13.hyps" by (simp add: rdvd_left_0_eq Let_def)} - moreover - {assume ai1: "abs i = 1" and ai: "check_int a" + moreover + {assume ai1: "abs i = 1" and ai: "check_int a" hence "i=1 \ i= - 1" by arith - moreover {assume i1: "i = 1" - from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] + moreover {assume i1: "i = 1" + from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] have ?case using i1 ai by simp } - moreover {assume i1: "i = - 1" - from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] + moreover {assume i1: "i = - 1" + from rdvd_left1_int[OF check_int[OF ai, simplified isint_iff]] rdvd_abs1[where d="- 1" and t="Inum bs a"] have ?case using i1 ai by simp } ultimately have ?case by blast} - moreover + moreover {assume inz: "i\0" and cond: "(abs i \ 1) \ (\ check_int a)" {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond by (cases "abs i = 1", auto simp add: int_rdvd_iff) } - moreover {assume H:"\ (\ v. ?sa = C v)" - hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond + moreover {assume H:"\ (\ v. ?sa = C v)" + hence th: "simpfm (NDvd i a) = NDvd (fst (simpdvd i ?sa)) (snd (simpdvd i ?sa))" using inz cond by (cases ?sa, auto simp add: Let_def split_def) from simpnum_nz have nz:"nozerocoeff ?sa" by simp from simpdvd [OF nz inz] th have ?case using sa by simp} @@ -1371,7 +1371,7 @@ "CJNB f p \ (let cjs = conjuncts p ; (yes,no) = List.partition bound0 cjs in conj (decr (list_conj yes)) (f (list_conj no)))" -lemma CJNB_qe: +lemma CJNB_qe: assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" shows "\ bs p. qfree p \ qfree (CJNB qe p) \ (Ifm bs ((CJNB qe p)) = Ifm bs (E p))" proof(clarify) @@ -1383,15 +1383,15 @@ let ?cno = "list_conj ?no" let ?cyes = "list_conj ?yes" have part: "List.partition bound0 ?cjs = (?yes,?no)" by simp - from partition_P[OF part] have "\ q\ set ?yes. bound0 q" by blast - hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) + from partition_P[OF part] have "\ q\ set ?yes. bound0 q" by blast + hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb) hence yes_qf: "qfree (decr ?cyes )" by (simp add: decr_qf) - from conjuncts_qf[OF qfp] partition_set[OF part] + from conjuncts_qf[OF qfp] partition_set[OF part] have " \q\ set ?no. qfree q" by auto hence no_qf: "qfree ?cno"by (simp add: list_conj_qf) - with qe have cno_qf:"qfree (qe ?cno )" + with qe have cno_qf:"qfree (qe ?cno )" and noE: "Ifm bs (qe ?cno) = Ifm bs (E ?cno)" by blast+ - from cno_qf yes_qf have qf: "qfree (CJNB qe p)" + from cno_qf yes_qf have qf: "qfree (CJNB qe p)" by (simp add: CJNB_def Let_def split_def) {fix bs from conjuncts have "Ifm bs p = (\q\ set ?cjs. Ifm bs q)" by blast @@ -1405,7 +1405,7 @@ by (auto simp add: decr[OF yes_nb] simp del: partition_filter_conv) also have "\ = (Ifm bs (conj (decr ?cyes) (qe ?cno)))" using qe[rule_format, OF no_qf] by auto - finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" + finally have "Ifm bs (E p) = Ifm bs (CJNB qe p)" by (simp add: Let_def CJNB_def split_def) with qf show "qfree (CJNB qe p) \ Ifm bs (CJNB qe p) = Ifm bs (E p)" by blast qed @@ -1414,8 +1414,8 @@ "qelim (E p) = (\ qe. DJ (CJNB qe) (qelim p qe))" | "qelim (A p) = (\ qe. not (qe ((qelim (NOT p) qe))))" | "qelim (NOT p) = (\ qe. not (qelim p qe))" -| "qelim (And p q) = (\ qe. conj (qelim p qe) (qelim q qe))" -| "qelim (Or p q) = (\ qe. disj (qelim p qe) (qelim q qe))" +| "qelim (And p q) = (\ qe. conj (qelim p qe) (qelim q qe))" +| "qelim (Or p q) = (\ qe. disj (qelim p qe) (qelim q qe))" | "qelim (Imp p q) = (\ qe. disj (qelim (NOT p) qe) (qelim q qe))" | "qelim (Iff p q) = (\ qe. iff (qelim p qe) (qelim q qe))" | "qelim p = (\ y. simpfm p)" @@ -1425,7 +1425,7 @@ lemma qelim_ci: assumes qe_inv: "\ bs p. qfree p \ qfree (qe p) \ (Ifm bs (qe p) = Ifm bs (E p))" shows "\ bs. qfree (qelim p qe) \ (Ifm bs (qelim p qe) = Ifm bs p)" - using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] + using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] by (induct p rule: qelim.induct) (auto simp del: simpfm.simps) @@ -1438,11 +1438,11 @@ | "zsplit0 (CN n c a) = zsplit0 (Add (Mul c (Bound n)) a)" | "zsplit0 (CF c a b) = zsplit0 (Add (Mul c (Floor a)) b)" | "zsplit0 (Neg a) = (let (i',a') = zsplit0 a in (-i', Neg a'))" -| "zsplit0 (Add a b) = (let (ia,a') = zsplit0 a ; - (ib,b') = zsplit0 b +| "zsplit0 (Add a b) = (let (ia,a') = zsplit0 a ; + (ib,b') = zsplit0 b in (ia+ib, Add a' b'))" -| "zsplit0 (Sub a b) = (let (ia,a') = zsplit0 a ; - (ib,b') = zsplit0 b +| "zsplit0 (Sub a b) = (let (ia,a') = zsplit0 a ; + (ib,b') = zsplit0 b in (ia-ib, Sub a' b'))" | "zsplit0 (Mul i a) = (let (i',a') = zsplit0 a in (i*i', Mul i a'))" | "zsplit0 (Floor a) = (let (i',a') = zsplit0 a in (i',Floor a'))" @@ -1453,18 +1453,18 @@ shows "\ n a. zsplit0 t = (n,a) \ (Inum ((real_of_int (x::int)) #bs) (CN 0 n a) = Inum (real_of_int x #bs) t) \ numbound0 a" (is "\ n a. ?S t = (n,a) \ (?I x (CN 0 n a) = ?I x t) \ ?N a") proof(induct t rule: zsplit0.induct) - case (1 c n a) thus ?case by auto + case (1 c n a) thus ?case by auto next case (2 m n a) thus ?case by (cases "m=0") auto next case (3 n i a n a') thus ?case by auto -next +next case (4 c a b n a') thus ?case by auto next case (5 t n a) let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" - have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \ n=-?nt" using 5 + have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Neg ?at \ n=-?nt" using 5 by (simp add: Let_def split_def) from abj 5 have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast from th2[simplified] th[simplified] show ?case by simp @@ -1474,15 +1474,15 @@ let ?as = "snd (zsplit0 s)" let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" - have abjs: "zsplit0 s = (?ns,?as)" by simp - moreover have abjt: "zsplit0 t = (?nt,?at)" by simp + have abjs: "zsplit0 s = (?ns,?as)" by simp + moreover have abjt: "zsplit0 t = (?nt,?at)" by simp ultimately have th: "a=Add ?as ?at \ n=?ns + ?nt" using 6 by (simp add: Let_def split_def) from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast from 6 have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real_of_int x # bs) (CN 0 xa xb) = Inum (real_of_int x # bs) t \ numbound0 xb)" by blast (*FIXME*) with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast from abjs 6 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast - from th3[simplified] th2[simplified] th[simplified] show ?case + from th3[simplified] th2[simplified] th[simplified] show ?case by (simp add: distrib_right) next case (7 s t n a) @@ -1490,15 +1490,15 @@ let ?as = "snd (zsplit0 s)" let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" - have abjs: "zsplit0 s = (?ns,?as)" by simp - moreover have abjt: "zsplit0 t = (?nt,?at)" by simp + have abjs: "zsplit0 s = (?ns,?as)" by simp + moreover have abjt: "zsplit0 t = (?nt,?at)" by simp ultimately have th: "a=Sub ?as ?at \ n=?ns - ?nt" using 7 by (simp add: Let_def split_def) from abjs[symmetric] have bluddy: "\ x y. (x,y) = zsplit0 s" by blast from 7 have "(\ x y. (x,y) = zsplit0 s) \ (\xa xb. zsplit0 t = (xa, xb) \ Inum (real_of_int x # bs) (CN 0 xa xb) = Inum (real_of_int x # bs) t \ numbound0 xb)" by blast (*FIXME*) with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) \ ?N ?at" by blast from abjs 7 have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) \ ?N ?as" by blast - from th3[simplified] th2[simplified] th[simplified] show ?case + from th3[simplified] th2[simplified] th[simplified] show ?case by (simp add: left_diff_distrib) next case (8 i t n a) @@ -1522,8 +1522,8 @@ have "?I x (Floor t) = ?I x (Floor (CN 0 ?nt ?at))" using th2 by simp also have "\ = real_of_int (floor ((real_of_int ?nt)* real_of_int(x) + ?I x ?at))" by simp also have "\ = real_of_int (floor (?I x ?at + real_of_int (?nt* x)))" by (simp add: ac_simps) - also have "\ = real_of_int (floor (?I x ?at) + (?nt* x))" - using floor_add_of_int[of "?I x ?at" "?nt* x"] by simp + also have "\ = real_of_int (floor (?I x ?at) + (?nt* x))" + using floor_add_of_int[of "?I x ?at" "?nt* x"] by simp also have "\ = real_of_int (?nt)*(real_of_int x) + real_of_int (floor (?I x ?at))" by (simp add: ac_simps) finally have "?I x (Floor t) = ?I x (CN 0 n a)" using th by simp with na show ?case by simp @@ -1533,17 +1533,17 @@ iszlfm :: "fm \ real list \ bool" (* Linearity test for fm *) zlfm :: "fm \ fm" (* Linearity transformation for fm *) recdef iszlfm "measure size" - "iszlfm (And p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" - "iszlfm (Or p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" + "iszlfm (And p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" + "iszlfm (Or p q) = (\ bs. iszlfm p bs \ iszlfm q bs)" "iszlfm (Eq (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" "iszlfm (NEq (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" "iszlfm (Lt (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" "iszlfm (Le (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" "iszlfm (Gt (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" "iszlfm (Ge (CN 0 c e)) = (\ bs. c>0 \ numbound0 e \ isint e bs)" - "iszlfm (Dvd i (CN 0 c e)) = + "iszlfm (Dvd i (CN 0 c e)) = (\ bs. c>0 \ i>0 \ numbound0 e \ isint e bs)" - "iszlfm (NDvd i (CN 0 c e))= + "iszlfm (NDvd i (CN 0 c e))= (\ bs. c>0 \ i>0 \ numbound0 e \ isint e bs)" "iszlfm p = (\ bs. isatom p \ (bound0 p))" @@ -1570,39 +1570,39 @@ "zlfm (Or p q) = disj (zlfm p) (zlfm q)" "zlfm (Imp p q) = disj (zlfm (NOT p)) (zlfm q)" "zlfm (Iff p q) = disj (conj (zlfm p) (zlfm q)) (conj (zlfm (NOT p)) (zlfm (NOT q)))" - "zlfm (Lt a) = (let (c,r) = zsplit0 a in - if c=0 then Lt r else - 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))) + "zlfm (Lt a) = (let (c,r) = zsplit0 a in + if c=0 then Lt r else + 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))) else Or (Gt (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" - "zlfm (Le a) = (let (c,r) = zsplit0 a in - if c=0 then Le r else - 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))) + "zlfm (Le a) = (let (c,r) = zsplit0 a in + if c=0 then Le r else + 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))) else Or (Ge (CN 0 (-c) (Floor(Neg r)))) (And (Eq(CN 0 (-c) (Floor(Neg r)))) (Lt (Add (Floor (Neg r)) r))))" - "zlfm (Gt a) = (let (c,r) = zsplit0 a in - if c=0 then Gt r else - if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) + "zlfm (Gt a) = (let (c,r) = zsplit0 a in + if c=0 then Gt r else + if c>0 then Or (Gt (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) else Or (Lt (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" - "zlfm (Ge a) = (let (c,r) = zsplit0 a in - if c=0 then Ge r else - if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) + "zlfm (Ge a) = (let (c,r) = zsplit0 a in + if c=0 then Ge r else + if c>0 then Or (Ge (CN 0 c (Floor r))) (And (Eq (CN 0 c (Floor r))) (Lt (Sub (Floor r) r))) else Or (Le (CN 0 (-c) (Neg (Floor r)))) (And (Eq(CN 0 (-c) (Neg (Floor r)))) (Lt (Sub (Floor r) r))))" - "zlfm (Eq a) = (let (c,r) = zsplit0 a in - if c=0 then Eq r else + "zlfm (Eq a) = (let (c,r) = zsplit0 a in + if c=0 then Eq r else if c>0 then (And (Eq (CN 0 c (Neg (Floor (Neg r))))) (Eq (Add (Floor (Neg r)) r))) else (And (Eq (CN 0 (-c) (Floor (Neg r)))) (Eq (Add (Floor (Neg r)) r))))" - "zlfm (NEq a) = (let (c,r) = zsplit0 a in - if c=0 then NEq r else + "zlfm (NEq a) = (let (c,r) = zsplit0 a in + if c=0 then NEq r else if c>0 then (Or (NEq (CN 0 c (Neg (Floor (Neg r))))) (NEq (Add (Floor (Neg r)) r))) else (Or (NEq (CN 0 (-c) (Floor (Neg r)))) (NEq (Add (Floor (Neg r)) r))))" - "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) - else (let (c,r) = zsplit0 a in - if c=0 then Dvd (abs i) r else - if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r))) + "zlfm (Dvd i a) = (if i=0 then zlfm (Eq a) + else (let (c,r) = zsplit0 a in + if c=0 then Dvd (abs i) r else + if c>0 then And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 c (Floor r))) else And (Eq (Sub (Floor r) r)) (Dvd (abs i) (CN 0 (-c) (Neg (Floor r))))))" - "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) - else (let (c,r) = zsplit0 a in - if c=0 then NDvd (abs i) r else - if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r))) + "zlfm (NDvd i a) = (if i=0 then zlfm (NEq a) + else (let (c,r) = zsplit0 a in + if c=0 then NDvd (abs i) r else + if c>0 then Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 c (Floor r))) else Or (NEq (Sub (Floor r) r)) (NDvd (abs i) (CN 0 (-c) (Neg (Floor r))))))" "zlfm (NOT (And p q)) = disj (zlfm (NOT p)) (zlfm (NOT q))" "zlfm (NOT (Or p q)) = conj (zlfm (NOT p)) (zlfm (NOT q))" @@ -1621,56 +1621,56 @@ "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)" "zlfm p = p" (hints simp add: fmsize_pos) -lemma split_int_less_real: +lemma split_int_less_real: "(real_of_int (a::int) < b) = (a < floor b \ (a = floor b \ real_of_int (floor b) < b))" proof( auto) assume alb: "real_of_int a < b" and agb: "\ a < floor b" from agb have "floor b \ a" by simp hence th: "b < real_of_int a + 1" by (simp only: floor_le_iff) - from floor_eq[OF alb th] show "a= floor b" by simp + from floor_eq[OF alb th] show "a= floor b" by simp next assume alb: "a < floor b" hence "real_of_int a < real_of_int (floor b)" by simp - moreover have "real_of_int (floor b) \ b" by simp ultimately show "real_of_int a < b" by arith + moreover have "real_of_int (floor b) \ b" by simp ultimately show "real_of_int a < b" by arith qed -lemma split_int_less_real': +lemma split_int_less_real': "(real_of_int (a::int) + b < 0) = (real_of_int a - real_of_int (floor(-b)) < 0 \ (real_of_int a - real_of_int (floor (-b)) = 0 \ real_of_int (floor (-b)) + b < 0))" -proof- +proof- have "(real_of_int a + b <0) = (real_of_int a < -b)" by arith - with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith + with split_int_less_real[where a="a" and b="-b"] show ?thesis by arith qed -lemma split_int_gt_real': +lemma split_int_gt_real': "(real_of_int (a::int) + b > 0) = (real_of_int a + real_of_int (floor b) > 0 \ (real_of_int a + real_of_int (floor b) = 0 \ real_of_int (floor b) - b < 0))" -proof- +proof- have th: "(real_of_int a + b >0) = (real_of_int (-a) + (-b)< 0)" by arith - show ?thesis using myless[of _ "real_of_int (floor b)"] - by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) + show ?thesis using myless[of _ "real_of_int (floor b)"] + by (simp only:th split_int_less_real'[where a="-a" and b="-b"]) (simp add: algebra_simps,arith) qed -lemma split_int_le_real: +lemma split_int_le_real: "(real_of_int (a::int) \ b) = (a \ floor b \ (a = floor b \ real_of_int (floor b) < b))" proof( auto) assume alb: "real_of_int a \ b" and agb: "\ a \ floor b" - from alb have "floor (real_of_int a) \ floor b " by (simp only: floor_mono) + from alb have "floor (real_of_int a) \ floor b " by (simp only: floor_mono) hence "a \ floor b" by simp with agb show "False" by simp next assume alb: "a \ floor b" hence "real_of_int a \ real_of_int (floor b)" by (simp only: floor_mono) - also have "\\ b" by simp finally show "real_of_int a \ b" . + also have "\\ b" by simp finally show "real_of_int a \ b" . qed -lemma split_int_le_real': +lemma split_int_le_real': "(real_of_int (a::int) + b \ 0) = (real_of_int a - real_of_int (floor(-b)) \ 0 \ (real_of_int a - real_of_int (floor (-b)) = 0 \ real_of_int (floor (-b)) + b < 0))" -proof- +proof- have "(real_of_int a + b \0) = (real_of_int a \ -b)" by arith - with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith + with split_int_le_real[where a="a" and b="-b"] show ?thesis by arith qed -lemma split_int_ge_real': +lemma split_int_ge_real': "(real_of_int (a::int) + b \ 0) = (real_of_int a + real_of_int (floor b) \ 0 \ (real_of_int a + real_of_int (floor b) = 0 \ real_of_int (floor b) - b < 0))" -proof- +proof- have th: "(real_of_int a + b \0) = (real_of_int (-a) + (-b) \ 0)" by arith show ?thesis by (simp only: th split_int_le_real'[where a="-a" and b="-b"]) (simp add: algebra_simps ,arith) @@ -1691,25 +1691,25 @@ (is "(?I (?l p) = ?I p) \ ?L (?l p)") using qfp proof(induct p rule: zlfm.induct) - case (5 a) + case (5 a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def,rename_tac nat a b,case_tac "nat", simp_all)} moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Lt a) = (real_of_int (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?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) finally have ?case using l by simp} moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Lt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Lt a) = (real_of_int (?c * i) + (?N ?r) < 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?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 ac_simps, arith) @@ -1720,46 +1720,46 @@ let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat",simp_all)} moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Le a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?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) finally have ?case using l by simp} moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Le a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Le a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?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 ac_simps, arith) finally have ?case using l by simp} ultimately show ?case by blast next - case (7 a) + case (7 a) let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Gt a) = (real_of_int (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?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) finally have ?case using l by simp} moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Gt a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Gt a) = (real_of_int (?c * i) + (?N ?r) > 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?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 ac_simps, arith) @@ -1770,21 +1770,21 @@ let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Ge a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?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) finally have ?case using l by simp} moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Ge a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Ge a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?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 ac_simps, arith) @@ -1795,21 +1795,21 @@ let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Eq a) = (real_of_int (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?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 of_int_mult[symmetric] del: of_int_mult) finally have ?case using l by simp} moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (Eq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (Eq a) = (real_of_int (?c * i) + (?N ?r) = 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?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 of_int_mult[symmetric] del: of_int_mult,arith) @@ -1820,21 +1820,21 @@ let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "?c = 0 \ (?c >0 \ ?c\0) \ (?c<0 \ ?c\0)" by arith moreover - {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] + {assume "?c=0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover - {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" + {assume cp: "?c > 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (NEq a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also have "\ = (?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 of_int_mult[symmetric] del: of_int_mult) finally have ?case using l by simp} moreover - {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" + {assume cn: "?c < 0" and cnz: "?c\0" hence l: "?L (?l (NEq a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) have "?I (NEq a) = (real_of_int (?c * i) + (?N ?r) \ 0)" using Ia by (simp add: Let_def split_def) also from cn cnz have "\ = (?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 of_int_mult[symmetric] del: of_int_mult,arith) @@ -1845,44 +1845,44 @@ let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith moreover - { assume j: "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) + { assume j: "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) hence ?case using 11 j by (simp del: zlfm.simps add: rdvd_left_0_eq)} moreover - {assume "?c=0" and "j\0" hence ?case + {assume "?c=0" and "j\0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover - {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" + {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r)))" + have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r)))" using Ia by (simp add: Let_def split_def) - also have "\ = (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r))" + also have "\ = (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r))" by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp - also have "\ = ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ - (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r))))" + also have "\ = ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ + (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r))))" by(simp only: int_rdvd_real[where i="abs j" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) - also have "\ = (?I (?l (Dvd j a)))" using cp cnz jnz - by (simp add: Let_def split_def int_rdvd_iff[symmetric] + also have "\ = (?I (?l (Dvd j a)))" using cp cnz jnz + by (simp add: Let_def split_def int_rdvd_iff[symmetric] del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by simp } moreover - {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" + {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r)))" + have "?I (Dvd j a) = (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r)))" using Ia by (simp add: Let_def split_def) - also have "\ = (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r))" + also have "\ = (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r))" by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp - also have "\ = ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ - (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r))))" + also have "\ = ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ + (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r))))" by(simp only: int_rdvd_real[where i="abs j" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (Dvd j a)))" using cn cnz jnz using rdvd_minus [where d="abs j" and t="real_of_int (?c*i + floor (?N ?r))", simplified, symmetric] - by (simp add: Let_def split_def int_rdvd_iff[symmetric] + by (simp add: Let_def split_def int_rdvd_iff[symmetric] del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by blast } ultimately show ?case by blast @@ -1891,44 +1891,44 @@ let ?c = "fst (zsplit0 a)" let ?r = "snd (zsplit0 a)" have spl: "zsplit0 a = (?c,?r)" by simp - from zsplit0_I[OF spl, where x="i" and bs="bs"] - have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto + from zsplit0_I[OF spl, where x="i" and bs="bs"] + have Ia:"Inum (real_of_int i # bs) a = Inum (real_of_int i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "\ t. Inum (real_of_int i#bs) t" have "j=0 \ (j\0 \ ?c = 0) \ (j\0 \ ?c >0 \ ?c\0) \ (j\ 0 \ ?c<0 \ ?c\0)" by arith moreover - {assume j: "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) + {assume j: "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) hence ?case using 12 j by (simp del: zlfm.simps add: rdvd_left_0_eq)} moreover - {assume "?c=0" and "j\0" hence ?case + {assume "?c=0" and "j\0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] rdvd_abs1[where d="j"] by (cases "?r", simp_all add: Let_def split_def, rename_tac nat a b, case_tac "nat", simp_all)} moreover - {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" + {assume cp: "?c > 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NDvd j a) = (\ (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r))))" + have "?I (NDvd j a) = (\ (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r))))" using Ia by (simp add: Let_def split_def) - also have "\ = (\ (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r)))" + also have "\ = (\ (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r)))" by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp - also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ - (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r)))))" + also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ + (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r)))))" by(simp only: int_rdvd_real[where i="abs j" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) - also have "\ = (?I (?l (NDvd j a)))" using cp cnz jnz - by (simp add: Let_def split_def int_rdvd_iff[symmetric] + also have "\ = (?I (?l (NDvd j a)))" using cp cnz jnz + by (simp add: Let_def split_def int_rdvd_iff[symmetric] del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by simp } moreover - {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" + {assume cn: "?c < 0" and cnz: "?c\0" and jnz: "j\0" hence l: "?L (?l (NDvd j a))" by (simp add: nb Let_def split_def isint_Floor isint_neg) - have "?I (NDvd j a) = (\ (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r))))" + have "?I (NDvd j a) = (\ (real_of_int j rdvd (real_of_int (?c * i) + (?N ?r))))" using Ia by (simp add: Let_def split_def) - also have "\ = (\ (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r)))" + also have "\ = (\ (real_of_int (abs j) rdvd real_of_int (?c*i) + (?N ?r)))" by (simp only: rdvd_abs1[where d="j" and t="real_of_int (?c*i) + ?N ?r", symmetric]) simp - also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ - (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r)))))" + also have "\ = (\ ((abs j) dvd (floor ((?N ?r) + real_of_int (?c*i))) \ + (real_of_int (floor ((?N ?r) + real_of_int (?c*i))) = (real_of_int (?c*i) + (?N ?r)))))" by(simp only: int_rdvd_real[where i="abs j" and x="real_of_int (?c*i) + (?N ?r)"]) (simp only: ac_simps) also have "\ = (?I (?l (NDvd j a)))" using cn cnz jnz using rdvd_minus [where d="abs j" and t="real_of_int (?c*i + floor (?N ?r))", simplified, symmetric] - by (simp add: Let_def split_def int_rdvd_iff[symmetric] + by (simp add: Let_def split_def int_rdvd_iff[symmetric] del: of_int_mult) (auto simp add: ac_simps) finally have ?case using l jnz by blast } ultimately show ?case by blast @@ -1940,8 +1940,8 @@ \d_\\ checks if a given l divides all the ds above\ fun minusinf:: "fm \ fm" where - "minusinf (And p q) = conj (minusinf p) (minusinf q)" -| "minusinf (Or p q) = disj (minusinf p) (minusinf q)" + "minusinf (And p q) = conj (minusinf p) (minusinf q)" +| "minusinf (Or p q) = disj (minusinf p) (minusinf q)" | "minusinf (Eq (CN 0 c e)) = F" | "minusinf (NEq (CN 0 c e)) = T" | "minusinf (Lt (CN 0 c e)) = T" @@ -1954,8 +1954,8 @@ by (induct p rule: minusinf.induct, auto) fun plusinf:: "fm \ fm" where - "plusinf (And p q) = conj (plusinf p) (plusinf q)" -| "plusinf (Or p q) = disj (plusinf p) (plusinf q)" + "plusinf (And p q) = conj (plusinf p) (plusinf q)" +| "plusinf (Or p q) = disj (plusinf p) (plusinf q)" | "plusinf (Eq (CN 0 c e)) = F" | "plusinf (NEq (CN 0 c e)) = T" | "plusinf (Lt (CN 0 c e)) = F" @@ -1965,20 +1965,20 @@ | "plusinf p = p" fun \ :: "fm \ int" where - "\ (And p q) = lcm (\ p) (\ q)" -| "\ (Or p q) = lcm (\ p) (\ q)" + "\ (And p q) = lcm (\ p) (\ q)" +| "\ (Or p q) = lcm (\ p) (\ q)" | "\ (Dvd i (CN 0 c e)) = i" | "\ (NDvd i (CN 0 c e)) = i" | "\ p = 1" fun d_\ :: "fm \ int \ bool" where - "d_\ (And p q) = (\ d. d_\ p d \ d_\ q d)" -| "d_\ (Or p q) = (\ d. d_\ p d \ d_\ q d)" + "d_\ (And p q) = (\ d. d_\ p d \ d_\ q d)" +| "d_\ (Or p q) = (\ d. d_\ p d \ d_\ q d)" | "d_\ (Dvd i (CN 0 c e)) = (\ d. i dvd d)" | "d_\ (NDvd i (CN 0 c e)) = (\ d. i dvd d)" | "d_\ p = (\ d. True)" -lemma delta_mono: +lemma delta_mono: assumes lin: "iszlfm p bs" and d: "d dvd d'" and ad: "d_\ p d" @@ -1996,17 +1996,17 @@ shows "d_\ p (\ p) \ \ p >0" using lin proof (induct p rule: iszlfm.induct) - case (1 p q) + case (1 p q) let ?d = "\ (And p q)" from 1 lcm_pos_int have dp: "?d >0" by simp - have d1: "\ p dvd \ (And p q)" using 1 by simp - hence th: "d_\ p ?d" + have d1: "\ p dvd \ (And p q)" using 1 by simp + hence th: "d_\ p ?d" using delta_mono 1 by (simp only: iszlfm.simps) blast - have "\ q dvd \ (And p q)" using 1 by simp + have "\ q dvd \ (And p q)" using 1 by simp hence th': "d_\ q ?d" using delta_mono 1 by (simp only: iszlfm.simps) blast - from th th' dp show ?case by simp + from th th' dp show ?case by simp next - case (2 p q) + case (2 p q) let ?d = "\ (And p q)" from 2 lcm_pos_int have dp: "?d >0" by simp have "\ p dvd \ (And p q)" using 2 by simp @@ -2041,101 +2041,101 @@ from z1_def z2_def have "\ x < ?z. ?I x (?M (Or f g)) = ?I x (Or f g)" by simp thus ?case by blast next - case (3 c e) + case (3 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 3 have nbe: "numbound0 e" by simp fix y have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Eq (CN 0 c e))) = ?I x (Eq (CN 0 c e))" - proof (simp add: less_floor_iff , rule allI, rule impI) + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) hence "real_of_int c * real_of_int x + Inum (y # bs) e \ 0"using rcpos by simp - thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" + thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] by simp qed thus ?case by blast next - case (4 c e) + case (4 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 4 have nbe: "numbound0 e" by simp fix y have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (NEq (CN 0 c e))) = ?I x (NEq (CN 0 c e))" - proof (simp add: less_floor_iff , rule allI, rule impI) + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) hence "real_of_int c * real_of_int x + Inum (y # bs) e \ 0"using rcpos by simp - thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" + thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] by simp qed thus ?case by blast next - case (5 c e) + case (5 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 5 have nbe: "numbound0 e" by simp fix y have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Lt (CN 0 c e))) = ?I x (Lt (CN 0 c e))" - proof (simp add: less_floor_iff , rule allI, rule impI) + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e < 0" + thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e < 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next - case (6 c e) + case (6 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 6 have nbe: "numbound0 e" by simp fix y have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Le (CN 0 c e))) = ?I x (Le (CN 0 c e))" - proof (simp add: less_floor_iff , rule allI, rule impI) + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" + thus "real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next - case (7 c e) + case (7 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 7 have nbe: "numbound0 e" by simp fix y have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Gt (CN 0 c e))) = ?I x (Gt (CN 0 c e))" - proof (simp add: less_floor_iff , rule allI, rule impI) + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - thus "\ (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e>0)" + thus "\ (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e>0)" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast next - case (8 c e) + case (8 c e) then have "c > 0" by simp hence rcpos: "real_of_int c > 0" by simp from 8 have nbe: "numbound0 e" by simp fix y have "\ x < (floor (- (Inum (y#bs) e) / (real_of_int c))). ?I x (?M (Ge (CN 0 c e))) = ?I x (Ge (CN 0 c e))" - proof (simp add: less_floor_iff , rule allI, rule impI) + proof (simp add: less_floor_iff , rule allI, rule impI) fix x :: int assume A: "real_of_int x + 1 \ - (Inum (y # bs) e / real_of_int c)" hence th1:"real_of_int x < - (Inum (y # bs) e / real_of_int c)" by simp with rcpos have "(real_of_int c)*(real_of_int x) < (real_of_int c)*(- (Inum (y # bs) e / real_of_int c))" by (simp only: mult_strict_left_mono [OF th1 rcpos]) - thus "\ real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" + thus "\ real_of_int c * real_of_int x + Inum (real_of_int x # bs) e \ 0" using numbound0_I[OF nbe, where b="y" and bs="bs" and b'="real_of_int x"] rcpos by simp qed thus ?case by blast @@ -2145,24 +2145,24 @@ assumes d: "d_\ p d" and linp: "iszlfm p (a # bs)" shows "Ifm ((real_of_int(x - k*d))#bs) (minusinf p) = Ifm (real_of_int x #bs) (minusinf p)" using linp d -proof(induct p rule: iszlfm.induct) +proof(induct p rule: iszlfm.induct) case (9 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ hence "\ k. d=i*k" by (simp add: dvd_def) then obtain "di" where di_def: "d=i*di" by blast - show ?case + show ?case proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real_of_int x - real_of_int k * real_of_int d" and b'="real_of_int x"] right_diff_distrib, rule iffI) - assume + assume "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") hence "\ (l::int). ?rt = ?ri * (real_of_int l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))" + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))" by (simp add: algebra_simps di_def) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int (l + c*k*di))" by (simp add: algebra_simps) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real_of_int l)" by blast thus "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" using rdvd_def by simp next - assume + assume "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx+?e") hence "\ (l::int). ?rc*?rx+?e = ?ri * (real_of_int l)" by (simp add: rdvd_def) hence "\ (l::int). ?rc*?rx - real_of_int c * (real_of_int k * real_of_int d) +?e = ?ri * (real_of_int l) - real_of_int c * (real_of_int k * real_of_int d)" by simp @@ -2176,20 +2176,20 @@ case (10 i c e) hence nbe: "numbound0 e" and id: "i dvd d" by simp+ hence "\ k. d=i*k" by (simp add: dvd_def) then obtain "di" where di_def: "d=i*di" by blast - show ?case + show ?case proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="real_of_int x - real_of_int k * real_of_int d" and b'="real_of_int x"] right_diff_distrib, rule iffI) - assume + assume "real_of_int i rdvd real_of_int c * real_of_int x - real_of_int c * (real_of_int k * real_of_int d) + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri rdvd ?rt") hence "\ (l::int). ?rt = ?ri * (real_of_int l)" by (simp add: rdvd_def) - hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))" + hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int l)+?rc*(?rk * (real_of_int i) * (real_of_int di))" by (simp add: algebra_simps di_def) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri*(real_of_int (l + c*k*di))" by (simp add: algebra_simps) hence "\ (l::int). ?rc*?rx+ ?I x e = ?ri* (real_of_int l)" by blast thus "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" using rdvd_def by simp next - assume + assume "real_of_int i rdvd real_of_int c * real_of_int x + Inum (real_of_int x # bs) e" (is "?ri rdvd ?rc*?rx+?e") hence "\ (l::int). ?rc*?rx+?e = ?ri * (real_of_int l)" by (simp add: rdvd_def) @@ -2221,7 +2221,7 @@ lemma minusinf_bex: assumes lin: "iszlfm p (real_of_int (a::int) #bs)" - shows "(\ (x::int). Ifm (real_of_int x#bs) (minusinf p)) = + shows "(\ (x::int). Ifm (real_of_int x#bs) (minusinf p)) = (\ (x::int)\ {1..\ p}. Ifm (real_of_int x#bs) (minusinf p))" (is "(\ x. ?P x) = _") proof- @@ -2234,7 +2234,7 @@ lemma dvd1_eq1: "x >0 \ (x::int) dvd 1 = (x = 1)" by auto -consts +consts a_\ :: "fm \ int \ fm" (* adjusts the coeffitients of a formula *) d_\ :: "fm \ int \ bool" (* tests if all coeffs c of c divide a given l*) \ :: "fm \ int" (* computes the lcm of all coefficients of x*) @@ -2242,8 +2242,8 @@ \ :: "fm \ num list" recdef a_\ "measure size" - "a_\ (And p q) = (\ k. And (a_\ p k) (a_\ q k))" - "a_\ (Or p q) = (\ k. Or (a_\ p k) (a_\ q k))" + "a_\ (And p q) = (\ k. And (a_\ p k) (a_\ q k))" + "a_\ (Or p q) = (\ k. Or (a_\ p k) (a_\ q k))" "a_\ (Eq (CN 0 c e)) = (\ k. Eq (CN 0 1 (Mul (k div c) e)))" "a_\ (NEq (CN 0 c e)) = (\ k. NEq (CN 0 1 (Mul (k div c) e)))" "a_\ (Lt (CN 0 c e)) = (\ k. Lt (CN 0 1 (Mul (k div c) e)))" @@ -2255,8 +2255,8 @@ "a_\ p = (\ k. p)" recdef d_\ "measure size" - "d_\ (And p q) = (\ k. (d_\ p k) \ (d_\ q k))" - "d_\ (Or p q) = (\ k. (d_\ p k) \ (d_\ q k))" + "d_\ (And p q) = (\ k. (d_\ p k) \ (d_\ q k))" + "d_\ (Or p q) = (\ k. (d_\ p k) \ (d_\ q k))" "d_\ (Eq (CN 0 c e)) = (\ k. c dvd k)" "d_\ (NEq (CN 0 c e)) = (\ k. c dvd k)" "d_\ (Lt (CN 0 c e)) = (\ k. c dvd k)" @@ -2268,8 +2268,8 @@ "d_\ p = (\ k. True)" recdef \ "measure size" - "\ (And p q) = lcm (\ p) (\ q)" - "\ (Or p q) = lcm (\ p) (\ q)" + "\ (And p q) = lcm (\ p) (\ q)" + "\ (Or p q) = lcm (\ p) (\ q)" "\ (Eq (CN 0 c e)) = c" "\ (NEq (CN 0 c e)) = c" "\ (Lt (CN 0 c e)) = c" @@ -2281,8 +2281,8 @@ "\ p = 1" recdef \ "measure size" - "\ (And p q) = (\ p @ \ q)" - "\ (Or p q) = (\ p @ \ q)" + "\ (And p q) = (\ p @ \ q)" + "\ (Or p q) = (\ p @ \ q)" "\ (Eq (CN 0 c e)) = [Sub (C (- 1)) e]" "\ (NEq (CN 0 c e)) = [Neg e]" "\ (Lt (CN 0 c e)) = []" @@ -2292,8 +2292,8 @@ "\ p = []" recdef \ "measure size" - "\ (And p q) = (\ p @ \ q)" - "\ (Or p q) = (\ p @ \ q)" + "\ (And p q) = (\ p @ \ q)" + "\ (Or p q) = (\ p @ \ q)" "\ (Eq (CN 0 c e)) = [Add (C (- 1)) e]" "\ (NEq (CN 0 c e)) = [e]" "\ (Lt (CN 0 c e)) = [e]" @@ -2303,8 +2303,8 @@ "\ p = []" consts mirror :: "fm \ fm" recdef mirror "measure size" - "mirror (And p q) = And (mirror p) (mirror q)" - "mirror (Or p q) = Or (mirror p) (mirror q)" + "mirror (And p q) = And (mirror p) (mirror q)" + "mirror (Or p q) = Or (mirror p) (mirror q)" "mirror (Eq (CN 0 c e)) = Eq (CN 0 c (Neg e))" "mirror (NEq (CN 0 c e)) = NEq (CN 0 c (Neg e))" "mirror (Lt (CN 0 c e)) = Gt (CN 0 c (Neg e))" @@ -2320,9 +2320,9 @@ shows "(Inum (real_of_int (i::int)#bs)) ` set (\ p) = (Inum (real_of_int i#bs)) ` set (\ (mirror p))" using lp by (induct p rule: mirror.induct) auto -lemma mirror: +lemma mirror: assumes lp: "iszlfm p (a#bs)" - shows "Ifm (real_of_int (x::int)#bs) (mirror p) = Ifm (real_of_int (- x)#bs) p" + shows "Ifm (real_of_int (x::int)#bs) (mirror p) = Ifm (real_of_int (- x)#bs) p" using lp proof(induct p rule: iszlfm.induct) case (9 j c e) @@ -2345,7 +2345,7 @@ lemma mirror_l: "iszlfm p (a#bs) \ iszlfm (mirror p) (a#bs)" by (induct p rule: mirror.induct) (auto simp add: isint_neg) -lemma mirror_d_\: "iszlfm p (a#bs) \ d_\ p 1 +lemma mirror_d_\: "iszlfm p (a#bs) \ d_\ p 1 \ iszlfm (mirror p) (a#bs) \ d_\ (mirror p) 1" by (induct p rule: mirror.induct) (auto simp add: isint_neg) @@ -2353,7 +2353,7 @@ by (induct p rule: mirror.induct) auto -lemma mirror_ex: +lemma mirror_ex: assumes lp: "iszlfm p (real_of_int (i::int)#bs)" shows "(\ (x::int). Ifm (real_of_int x#bs) (mirror p)) = (\ (x::int). Ifm (real_of_int x#bs) p)" (is "(\ x. ?I x ?mp) = (\ x. ?I x p)") @@ -2361,7 +2361,7 @@ fix x assume "?I x ?mp" hence "?I (- x) p" using mirror[OF lp] by blast thus "\ x. ?I x p" by blast next - fix x assume "?I x p" hence "?I (- x) ?mp" + fix x assume "?I x p" hence "?I (- x) ?mp" using mirror[OF lp, where x="- x", symmetric] by auto thus "\ x. ?I x ?mp" by blast qed @@ -2370,7 +2370,7 @@ shows "\ b\ set (\ p). numbound0 b" using lp by (induct p rule: \.induct,auto) -lemma d_\_mono: +lemma d_\_mono: assumes linp: "iszlfm p (a #bs)" and dr: "d_\ p l" and d: "l dvd l'" @@ -2383,7 +2383,7 @@ using lp by(induct p rule: \.induct, auto simp add: isint_add isint_c) -lemma \: +lemma \: assumes linp: "iszlfm p (a #bs)" shows "\ p > 0 \ d_\ p (\ p)" using linp @@ -2391,15 +2391,15 @@ case (1 p q) then have dl1: "\ p dvd lcm (\ p) (\ q)" by simp from 1 have dl2: "\ q dvd lcm (\ p) (\ q)" by simp - from 1 d_\_mono[where p = "p" and l="\ p" and l'="lcm (\ p) (\ q)"] - d_\_mono[where p = "q" and l="\ q" and l'="lcm (\ p) (\ q)"] + from 1 d_\_mono[where p = "p" and l="\ p" and l'="lcm (\ p) (\ q)"] + d_\_mono[where p = "q" and l="\ q" and l'="lcm (\ p) (\ q)"] dl1 dl2 show ?case by (auto simp add: lcm_pos_int) next case (2 p q) then have dl1: "\ p dvd lcm (\ p) (\ q)" by simp from 2 have dl2: "\ q dvd lcm (\ p) (\ q)" by simp - from 2 d_\_mono[where p = "p" and l="\ p" and l'="lcm (\ p) (\ q)"] - d_\_mono[where p = "q" and l="\ q" and l'="lcm (\ p) (\ q)"] + from 2 d_\_mono[where p = "p" and l="\ p" and l'="lcm (\ p) (\ q)"] + d_\_mono[where p = "q" and l="\ q" and l'="lcm (\ p) (\ q)"] dl1 dl2 show ?case by (auto simp add: lcm_pos_int) qed (auto simp add: lcm_pos_int) @@ -2412,10 +2412,10 @@ from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" + then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e < (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e < 0)" @@ -2430,10 +2430,10 @@ from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" + then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" @@ -2448,10 +2448,10 @@ from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" + then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e > (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e > 0)" @@ -2466,10 +2466,10 @@ from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" + then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" @@ -2484,10 +2484,10 @@ from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" + then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = 0)" @@ -2502,10 +2502,10 @@ from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" + then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ (0::real)) = (real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e \ 0)" @@ -2520,27 +2520,27 @@ from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" + then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(\ (k::int). real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k) = (\ (k::int). real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k)" by simp also have "\ = (\ (k::int). real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k) = real_of_int (l div c)*0)" by (simp add: algebra_simps) also fix k have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k = 0)" using zero_le_mult_iff [where a="real_of_int (l div c)" and b="real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k"] ldcp by simp also have "\ = (\ (k::int). real_of_int c * real_of_int x + Inum (real_of_int x # bs) e = real_of_int j * real_of_int k)" by simp - finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp + finally show ?case using numbound0_I[OF be,where b="real_of_int (l * x)" and b'="real_of_int x" and bs="bs"] rdvd_def be isint_Mul[OF ei] mult_strict_mono[OF ldcp jp ldcp ] by simp next 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+ from lp cp have clel: "c\l" by (simp add: zdvd_imp_le [OF d' lp]) from cp have cnz: "c \ 0" by simp have "c div c\ l div c" by (simp add: zdiv_mono1[OF clel cp]) - then have ldcp:"0 < l div c" + then have ldcp:"0 < l div c" by (simp add: div_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' dvd_eq_mod_eq_0[of "c" "l"] by simp - hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] + hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(\ (k::int). real_of_int l * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k) = (\ (k::int). real_of_int (c * (l div c)) * real_of_int x + real_of_int (l div c) * Inum (real_of_int x # bs) e = (real_of_int (l div c) * real_of_int j) * real_of_int k)" by simp also have "\ = (\ (k::int). real_of_int (l div c) * (real_of_int c * real_of_int x + Inum (real_of_int x # bs) e - real_of_int j * real_of_int k) = real_of_int (l div c)*0)" by (simp add: algebra_simps) @@ -2557,7 +2557,7 @@ have "(\ x. l dvd x \ ?P x) = (\ (x::int). ?P (l*x))" using unity_coeff_ex[where l="l" and P="?P", simplified] by simp also have "\ = (\ (x::int). ?P' x)" using a_\[OF linp d lp] by simp - finally show ?thesis . + finally show ?thesis . qed lemma \: @@ -2584,40 +2584,40 @@ from ie1 have ie: "real_of_int (floor ?e) = ?e" using isint_iff[where n="e" and bs="a#bs"] numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"] by (simp add: isint_iff) - {assume "real_of_int (x-d) +?e > 0" hence ?case using c1 + {assume "real_of_int (x-d) +?e > 0" hence ?case using c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] by (simp del: of_int_minus)} moreover - {assume H: "\ real_of_int (x-d) + ?e > 0" + {assume H: "\ real_of_int (x-d) + ?e > 0" let ?v="Neg e" have vb: "?v \ set (\ (Gt (CN 0 c e)))" by simp - from 7(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]] - have nob: "\ (\ j\ {1 ..d}. real_of_int x = - ?e + real_of_int j)" by auto + from 7(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]] + have nob: "\ (\ j\ {1 ..d}. real_of_int x = - ?e + real_of_int j)" by auto from H p have "real_of_int x + ?e > 0 \ real_of_int x + ?e \ real_of_int d" by (simp add: c1) hence "real_of_int (x + floor ?e) > real_of_int (0::int) \ real_of_int (x + floor ?e) \ real_of_int d" using ie by simp hence "x + floor ?e \ 1 \ x + floor ?e \ d" by simp hence "\ (j::int) \ {1 .. d}. j = x + floor ?e" by simp - hence "\ (j::int) \ {1 .. d}. real_of_int x = real_of_int (- floor ?e + j)" by force - hence "\ (j::int) \ {1 .. d}. real_of_int x = - ?e + real_of_int j" + hence "\ (j::int) \ {1 .. d}. real_of_int x = real_of_int (- floor ?e + j)" by force + hence "\ (j::int) \ {1 .. d}. real_of_int x = - ?e + real_of_int j" by (simp add: ie[simplified isint_iff]) with nob have ?case by auto} ultimately show ?case by blast next - case (8 c e) hence p: "Ifm (real_of_int x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" + case (8 c e) hence p: "Ifm (real_of_int x #bs) (Ge (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+ let ?e = "Inum (real_of_int x # bs) e" from ie1 have ie: "real_of_int (floor ?e) = ?e" using numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real_of_int x)#bs"] by (simp add: isint_iff) - {assume "real_of_int (x-d) +?e \ 0" hence ?case using c1 + {assume "real_of_int (x-d) +?e \ 0" hence ?case using c1 numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] by (simp del: of_int_minus)} moreover - {assume H: "\ real_of_int (x-d) + ?e \ 0" + {assume H: "\ real_of_int (x-d) + ?e \ 0" let ?v="Sub (C (- 1)) e" have vb: "?v \ set (\ (Ge (CN 0 c e)))" by simp - from 8(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]] - have nob: "\ (\ j\ {1 ..d}. real_of_int x = - ?e - 1 + real_of_int j)" by auto + from 8(5)[simplified simp_thms Inum.simps \.simps list.set bex_simps numbound0_I[OF bn,where b="a" and b'="real_of_int x" and bs="bs"]] + have nob: "\ (\ j\ {1 ..d}. real_of_int x = - ?e - 1 + real_of_int j)" by auto from H p have "real_of_int x + ?e \ 0 \ real_of_int x + ?e < real_of_int d" by (simp add: c1) hence "real_of_int (x + floor ?e) \ real_of_int (0::int) \ real_of_int (x + floor ?e) < real_of_int d" using ie by simp @@ -2625,12 +2625,12 @@ hence "\ (j::int) \ {1 .. d}. j = x + floor ?e + 1" by simp hence "\ (j::int) \ {1 .. d}. x= - floor ?e - 1 + j" by (simp add: algebra_simps) hence "\ (j::int) \ {1 .. d}. real_of_int x= real_of_int (- floor ?e - 1 + j)" by presburger - hence "\ (j::int) \ {1 .. d}. real_of_int x= - ?e - 1 + real_of_int j" + hence "\ (j::int) \ {1 .. d}. real_of_int x= - ?e - 1 + real_of_int j" by (simp add: ie[simplified isint_iff]) with nob have ?case by simp } ultimately show ?case by blast next - case (3 c e) hence p: "Ifm (real_of_int x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" + case (3 c e) hence p: "Ifm (real_of_int x #bs) (Eq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (real_of_int x # bs) e" let ?v="(Sub (C (- 1)) e)" @@ -2639,12 +2639,12 @@ by simp (erule ballE[where x="1"], simp_all add:algebra_simps numbound0_I[OF bn,where b="real_of_int x"and b'="a"and bs="bs"]) next - case (4 c e)hence p: "Ifm (real_of_int x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" + case (4 c e)hence p: "Ifm (real_of_int x #bs) (NEq (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" and ie1: "isint e (a #bs)" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (real_of_int x # bs) e" let ?v="Neg e" have vb: "?v \ set (\ (NEq (CN 0 c e)))" by simp - {assume "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) e \ 0" + {assume "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) e \ 0" hence ?case by (simp add: c1)} moreover {assume H: "real_of_int x - real_of_int d + Inum ((real_of_int (x -d)) # bs) e = 0" @@ -2653,57 +2653,57 @@ by (simp add: numbound0_I[OF bn,where b="real_of_int x - real_of_int d"and b'="a"and bs="bs"]) with 4(5) have ?case using dp by simp} ultimately show ?case by blast -next - case (9 j c e) hence p: "Ifm (real_of_int x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" +next + case (9 j c e) hence p: "Ifm (real_of_int x #bs) (Dvd j (CN 0 c e))" (is "?p x") and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (real_of_int x # bs) e" - from 9 have "isint e (a #bs)" by simp + from 9 have "isint e (a #bs)" by simp hence ie: "real_of_int (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real_of_int x)#bs"] numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] by (simp add: isint_iff) from 9 have id: "j dvd d" by simp from c1 ie[symmetric] have "?p x = (real_of_int j rdvd real_of_int (x+ floor ?e))" by simp - also have "\ = (j dvd x + floor ?e)" + also have "\ = (j dvd x + floor ?e)" using int_rdvd_real[where i="j" and x="real_of_int (x+ floor ?e)"] by simp - also have "\ = (j dvd x - d + floor ?e)" + also have "\ = (j dvd x - d + floor ?e)" using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp - also have "\ = (real_of_int j rdvd real_of_int (x - d + floor ?e))" + also have "\ = (real_of_int j rdvd real_of_int (x - d + floor ?e))" using int_rdvd_real[where i="j" and x="real_of_int (x-d + floor ?e)",symmetric, simplified] ie by simp - also have "\ = (real_of_int j rdvd real_of_int x - real_of_int d + ?e)" + also have "\ = (real_of_int j rdvd real_of_int x - real_of_int d + ?e)" using ie by simp - finally show ?case + finally show ?case using numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] c1 p by simp next case (10 j c e) hence p: "Ifm (real_of_int 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+ let ?e = "Inum (real_of_int x # bs) e" - from 10 have "isint e (a#bs)" by simp + from 10 have "isint e (a#bs)" by simp hence ie: "real_of_int (floor ?e) = ?e" using numbound0_I[OF bn,where b="real_of_int x" and b'="a" and bs="bs"] isint_iff[where n="e" and bs="(real_of_int x)#bs"] by (simp add: isint_iff) from 10 have id: "j dvd d" by simp from c1 ie[symmetric] have "?p x = (\ real_of_int j rdvd real_of_int (x+ floor ?e))" by simp - also have "\ = (\ j dvd x + floor ?e)" + also have "\ = (\ j dvd x + floor ?e)" using int_rdvd_real[where i="j" and x="real_of_int (x+ floor ?e)"] by simp - also have "\ = (\ j dvd x - d + floor ?e)" + also have "\ = (\ j dvd x - d + floor ?e)" using dvd_period[OF id, where x="x" and c="-1" and t="floor ?e"] by simp - also have "\ = (\ real_of_int j rdvd real_of_int (x - d + floor ?e))" + also have "\ = (\ real_of_int j rdvd real_of_int (x - d + floor ?e))" using int_rdvd_real[where i="j" and x="real_of_int (x-d + floor ?e)",symmetric, simplified] ie by simp - also have "\ = (\ real_of_int j rdvd real_of_int x - real_of_int d + ?e)" + also have "\ = (\ real_of_int j rdvd real_of_int x - real_of_int d + ?e)" using ie by simp finally show ?case using numbound0_I[OF bn,where b="real_of_int (x-d)" and b'="real_of_int x" and bs="bs"] c1 p by simp qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int (x - d)" and b'="real_of_int x"] simp del: of_int_diff) -lemma \': +lemma \': assumes lp: "iszlfm p (a #bs)" and u: "d_\ p 1" and d: "d_\ p d" and dp: "d > 0" shows "\ x. \(\(j::int) \ {1 .. d}. \ b\ set(\ p). Ifm ((Inum (a#bs) b + real_of_int j) #bs) p) \ Ifm (real_of_int x#bs) p \ Ifm (real_of_int (x - d)#bs) p" (is "\ x. ?b \ ?P x \ ?P (x - d)") proof(clarify) - fix x - assume nb:"?b" and px: "?P x" + fix x + assume nb:"?b" and px: "?P x" hence nb2: "\(\(j::int) \ {1 .. d}. \ b\ (Inum (a#bs)) ` set(\ p). real_of_int x = b + real_of_int j)" by auto from \[OF lp u d dp nb2 px] show "?P (x -d )" . @@ -2714,7 +2714,7 @@ using lp by (induct p rule: iszlfm.induct) (auto simp add: isint_neg isint_sub) lemma cpmi_eq: "0 < D \ (EX z::int. ALL x. x < z --> (P x = P1 x)) -==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) +==> ALL x.~(EX (j::int) : {1..D}. EX (b::int) : B. P(b+j)) --> P (x) --> P (x - D) ==> (ALL (x::int). ALL (k::int). ((P1 x)= (P1 (x-k*D)))) ==> (EX (x::int). P(x)) = ((EX (j::int) : {1..D} . (P1(j))) | (EX (j::int) : {1..D}. EX (b::int) : B. P (b+j)))" apply(rule iffI) @@ -2744,18 +2744,18 @@ shows "(\ (x::int). Ifm (real_of_int x #bs) p) = (\ j\ {1.. d}. Ifm (real_of_int j #bs) (minusinf p) \ (\ b \ set (\ p). Ifm ((Inum (a#bs) b + real_of_int j) #bs) p))" (is "(\ (x::int). ?P (real_of_int x)) = (\ j\ ?D. ?M j \ (\ b\ ?B. ?P (?I b + real_of_int j)))") proof- - from minusinf_inf[OF lp] + from minusinf_inf[OF lp] have th: "\(z::int). \x_int[OF lp] isint_iff[where bs="a # bs"] have B: "\ b\ ?B. real_of_int (floor (?I b)) = ?I b" by simp - from B[rule_format] - have "(\j\?D. \b\ ?B. ?P (?I b + real_of_int j)) = (\j\?D. \b\ ?B. ?P (real_of_int (floor (?I b)) + real_of_int j))" + from B[rule_format] + have "(\j\?D. \b\ ?B. ?P (?I b + real_of_int j)) = (\j\?D. \b\ ?B. ?P (real_of_int (floor (?I b)) + real_of_int j))" by simp also have "\ = (\j\?D. \b\ ?B. ?P (real_of_int (floor (?I b) + j)))" by simp also have"\ = (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j)))" by blast - finally have BB': - "(\j\?D. \b\ ?B. ?P (?I b + real_of_int j)) = (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j)))" - by blast + finally have BB': + "(\j\?D. \b\ ?B. ?P (?I b + real_of_int j)) = (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j)))" + by blast hence th2: "\ x. \ (\ j \ ?D. \ b \ ?B'. ?P (real_of_int (b + j))) \ ?P (real_of_int x) \ ?P (real_of_int (x - d))" using \'[OF lp u d dp] by blast from minusinf_repeats[OF d lp] have th3: "\ x k. ?M x = ?M (x-k*d)" by simp @@ -2765,14 +2765,14 @@ (* Reddy and Loveland *) -consts +consts \ :: "fm \ (num \ int) list" (* Compute the Reddy and Loveland Bset*) \_\:: "fm \ num \ int \ fm" (* Performs the modified substitution of Reddy and Loveland*) \_\ :: "fm \ (num\int) list" a_\ :: "fm \ int \ fm" recdef \ "measure size" - "\ (And p q) = (\ p @ \ q)" - "\ (Or p q) = (\ p @ \ q)" + "\ (And p q) = (\ p @ \ q)" + "\ (Or p q) = (\ p @ \ q)" "\ (Eq (CN 0 c e)) = [(Sub (C (- 1)) e,c)]" "\ (NEq (CN 0 c e)) = [(Neg e,c)]" "\ (Lt (CN 0 c e)) = []" @@ -2782,29 +2782,29 @@ "\ p = []" recdef \_\ "measure size" - "\_\ (And p q) = (\ (t,k). And (\_\ p (t,k)) (\_\ q (t,k)))" - "\_\ (Or p q) = (\ (t,k). Or (\_\ p (t,k)) (\_\ q (t,k)))" - "\_\ (Eq (CN 0 c e)) = (\ (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) + "\_\ (And p q) = (\ (t,k). And (\_\ p (t,k)) (\_\ q (t,k)))" + "\_\ (Or p q) = (\ (t,k). Or (\_\ p (t,k)) (\_\ q (t,k)))" + "\_\ (Eq (CN 0 c e)) = (\ (t,k). if k dvd c then (Eq (Add (Mul (c div k) t) e)) else (Eq (Add (Mul c t) (Mul k e))))" - "\_\ (NEq (CN 0 c e)) = (\ (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) + "\_\ (NEq (CN 0 c e)) = (\ (t,k). if k dvd c then (NEq (Add (Mul (c div k) t) e)) else (NEq (Add (Mul c t) (Mul k e))))" - "\_\ (Lt (CN 0 c e)) = (\ (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) + "\_\ (Lt (CN 0 c e)) = (\ (t,k). if k dvd c then (Lt (Add (Mul (c div k) t) e)) else (Lt (Add (Mul c t) (Mul k e))))" - "\_\ (Le (CN 0 c e)) = (\ (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) + "\_\ (Le (CN 0 c e)) = (\ (t,k). if k dvd c then (Le (Add (Mul (c div k) t) e)) else (Le (Add (Mul c t) (Mul k e))))" - "\_\ (Gt (CN 0 c e)) = (\ (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) + "\_\ (Gt (CN 0 c e)) = (\ (t,k). if k dvd c then (Gt (Add (Mul (c div k) t) e)) else (Gt (Add (Mul c t) (Mul k e))))" - "\_\ (Ge (CN 0 c e)) = (\ (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) + "\_\ (Ge (CN 0 c e)) = (\ (t,k). if k dvd c then (Ge (Add (Mul (c div k) t) e)) else (Ge (Add (Mul c t) (Mul k e))))" - "\_\ (Dvd i (CN 0 c e)) =(\ (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e)) + "\_\ (Dvd i (CN 0 c e)) =(\ (t,k). if k dvd c then (Dvd i (Add (Mul (c div k) t) e)) else (Dvd (i*k) (Add (Mul c t) (Mul k e))))" - "\_\ (NDvd i (CN 0 c e))=(\ (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e)) + "\_\ (NDvd i (CN 0 c e))=(\ (t,k). if k dvd c then (NDvd i (Add (Mul (c div k) t) e)) else (NDvd (i*k) (Add (Mul c t) (Mul k e))))" "\_\ p = (\ (t,k). p)" recdef \_\ "measure size" - "\_\ (And p q) = (\_\ p @ \_\ q)" - "\_\ (Or p q) = (\_\ p @ \_\ q)" + "\_\ (And p q) = (\_\ p @ \_\ q)" + "\_\ (Or p q) = (\_\ p @ \_\ q)" "\_\ (Eq (CN 0 c e)) = [(Add (C (- 1)) e,c)]" "\_\ (NEq (CN 0 c e)) = [(e,c)]" "\_\ (Lt (CN 0 c e)) = [(e,c)]" @@ -2822,19 +2822,19 @@ and tnb: "numbound0 t" and tint: "isint t (real_of_int x#bs)" and kdt: "k dvd floor (Inum (b'#bs) t)" - shows "Ifm (real_of_int x#bs) (\_\ p (t,k)) = - (Ifm ((real_of_int ((floor (Inum (b'#bs) t)) div k))#bs) p)" + shows "Ifm (real_of_int x#bs) (\_\ p (t,k)) = + (Ifm ((real_of_int ((floor (Inum (b'#bs) t)) div k))#bs) p)" (is "?I (real_of_int x) (?s p) = (?I (real_of_int ((floor (?N b' t)) div k)) p)" is "_ = (?I ?tk p)") using linp kpos tnb proof(induct p rule: \_\.induct) - case (3 c e) + case (3 c e) from 3 have cp: "c > 0" and nb: "numbound0 e" by auto - { assume kdc: "k dvd c" + { assume kdc: "k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } - moreover + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } + moreover { assume *: "\ k dvd c" from kpos have knz': "real_of_int k \ 0" by simp from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" @@ -2851,16 +2851,16 @@ numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } - ultimately show ?case by blast + ultimately show ?case by blast next - case (4 c e) + case (4 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto - { assume kdc: "k dvd c" + { assume kdc: "k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } - moreover + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } + moreover { assume *: "\ k dvd c" from kpos have knz': "real_of_int k \ 0" by simp from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp @@ -2876,16 +2876,16 @@ numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } - ultimately show ?case by blast + ultimately show ?case by blast next - case (5 c e) + case (5 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto - { assume kdc: "k dvd c" + { assume kdc: "k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } - moreover + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } + moreover { assume *: "\ k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Lt (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k < 0)" @@ -2900,16 +2900,16 @@ numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } - ultimately show ?case by blast + ultimately show ?case by blast next case (6 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto - { assume kdc: "k dvd c" + { assume kdc: "k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } - moreover + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } + moreover { assume *: "\ k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Le (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \ 0)" @@ -2924,16 +2924,16 @@ numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } - ultimately show ?case by blast + ultimately show ?case by blast next - case (7 c e) + case (7 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto - { assume kdc: "k dvd c" + { assume kdc: "k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } - moreover + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } + moreover { assume *: "\ k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Gt (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k > 0)" @@ -2948,16 +2948,16 @@ numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } - ultimately show ?case by blast + ultimately show ?case by blast next - case (8 c e) + case (8 c e) then have cp: "c > 0" and nb: "numbound0 e" by auto - { assume kdc: "k dvd c" + { assume kdc: "k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } - moreover + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } + moreover { assume *: "\ k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from assms * have "?I (real_of_int x) (?s (Ge (CN 0 c e))) = ((real_of_int c * (?N (real_of_int x) t / real_of_int k) + ?N (real_of_int x) e)* real_of_int k \ 0)" @@ -2972,16 +2972,16 @@ numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } - ultimately show ?case by blast + ultimately show ?case by blast next case (9 i c e) then have cp: "c > 0" and nb: "numbound0 e" by auto - { assume kdc: "k dvd c" + { assume kdc: "k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } - moreover + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } + moreover { assume *: "\ k dvd c" from kpos have knz: "k\0" by simp hence knz': "real_of_int k \ 0" by simp from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp @@ -2996,16 +2996,16 @@ numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } - ultimately show ?case by blast + ultimately show ?case by blast next case (10 i c e) then have cp: "c > 0" and nb: "numbound0 e" by auto - { assume kdc: "k dvd c" + { assume kdc: "k dvd c" from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp from kdc have ?case using real_of_int_div[OF kdc] real_of_int_div[OF kdt] numbound0_I[OF tnb, where bs="bs" and b="b'" and b'="real_of_int x"] - numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } - moreover + numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) } + moreover { assume *: "\ k dvd c" from kpos have knz: "k\0" by simp hence knz': "real_of_int k \ 0" by simp from tint have ti: "real_of_int (floor (?N (real_of_int x) t)) = ?N (real_of_int x) t" using isint_def by simp @@ -3020,7 +3020,7 @@ numbound0_I[OF nb, where bs="bs" and b="?tk" and b'="real_of_int x"] by (simp add: ti) finally have ?case . } - ultimately show ?case by blast + ultimately show ?case by blast qed (simp_all add: bound0_I[where bs="bs" and b="real_of_int ((floor (?N b' t)) div k)" and b'="real_of_int x"] numbound0_I[where bs="bs" and b="real_of_int ((floor (?N b' t)) div k)" and b'="real_of_int x"]) @@ -3056,7 +3056,7 @@ from mult_strict_left_mono[OF dp cp] have one:"1 \ {1 .. c*d}" by auto from nob[rule_format, where j="1", OF one] pi show ?case by simp next - case (4 c e) + case (4 c e) hence cp: "c >0" and nb: "numbound0 e" and ei: "isint e (real_of_int i#bs)" and nob: "\ j\ {1 .. c*d}. real_of_int (c*i) \ - ?N i e + real_of_int j" by simp+ @@ -3070,18 +3070,18 @@ ultimately show ?case by blast next case (5 c e) hence cp: "c > 0" by simp - from 5 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric] + from 5 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric] of_int_mult] - show ?case using 5 dp - apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] + show ?case using 5 dp + apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] algebra_simps del: mult_pos_pos) by (metis add.right_neutral of_int_0_less_iff of_int_mult pos_add_strict) next case (6 c e) hence cp: "c > 0" by simp - from 6 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric] + from 6 mult_strict_left_mono[OF dp cp, simplified of_int_less_iff[symmetric] of_int_mult] - show ?case using 6 dp - apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] + show ?case using 6 dp + apply (simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"] algebra_simps del: mult_pos_pos) using order_trans by fastforce next @@ -3096,9 +3096,9 @@ have "real_of_int (c*i) + ?N i e > real_of_int (c*d) \ real_of_int (c*i) + ?N i e \ real_of_int (c*d)" by auto moreover {assume "real_of_int (c*i) + ?N i e > real_of_int (c*d)" hence ?case - by (simp add: algebra_simps - numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])} - moreover + by (simp add: algebra_simps + numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])} + moreover {assume H:"real_of_int (c*i) + ?N i e \ real_of_int (c*d)" with ei[simplified isint_iff] have "real_of_int (c*i + ?fe) \ real_of_int (c*d)" by simp hence pid: "c*i + ?fe \ c*d" by (simp only: of_int_le_iff) @@ -3119,9 +3119,9 @@ have "real_of_int (c*i) + ?N i e \ real_of_int (c*d) \ real_of_int (c*i) + ?N i e < real_of_int (c*d)" by auto moreover {assume "real_of_int (c*i) + ?N i e \ real_of_int (c*d)" hence ?case - by (simp add: algebra_simps - numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])} - moreover + by (simp add: algebra_simps + numbound0_I[OF nb,where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"])} + moreover {assume H:"real_of_int (c*i) + ?N i e < real_of_int (c*d)" with ei[simplified isint_iff] have "real_of_int (c*i + ?fe) < real_of_int (c*d)" by simp hence pid: "c*i + 1 + ?fe \ c*d" by (simp only: of_int_le_iff) @@ -3137,61 +3137,61 @@ next case (9 j c e) hence p: "real_of_int j rdvd real_of_int (c*i) + ?N i e" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ let ?e = "Inum (real_of_int i # bs) e" - from 9 have "isint e (real_of_int i #bs)" by simp + from 9 have "isint e (real_of_int i #bs)" by simp hence ie: "real_of_int (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real_of_int i)#bs"] numbound0_I[OF bn,where b="real_of_int i" and b'="real_of_int i" and bs="bs"] by (simp add: isint_iff) from 9 have id: "j dvd d" by simp from ie[symmetric] have "?p i = (real_of_int j rdvd real_of_int (c*i+ floor ?e))" by simp - also have "\ = (j dvd c*i + floor ?e)" + also have "\ = (j dvd c*i + floor ?e)" using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp - also have "\ = (j dvd c*i - c*d + floor ?e)" + also have "\ = (j dvd c*i - c*d + floor ?e)" using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp - also have "\ = (real_of_int j rdvd real_of_int (c*i - c*d + floor ?e))" + also have "\ = (real_of_int j rdvd real_of_int (c*i - c*d + floor ?e))" using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] ie by simp - also have "\ = (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)" + also have "\ = (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)" using ie by (simp add:algebra_simps) - finally show ?case - using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p + finally show ?case + using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p by (simp add: algebra_simps) next case (10 j c e) hence p: "\ (real_of_int j rdvd real_of_int (c*i) + ?N i e)" (is "?p x") and cp: "c > 0" and bn:"numbound0 e" by simp+ let ?e = "Inum (real_of_int i # bs) e" - from 10 have "isint e (real_of_int i #bs)" by simp + from 10 have "isint e (real_of_int i #bs)" by simp hence ie: "real_of_int (floor ?e) = ?e" using isint_iff[where n="e" and bs="(real_of_int i)#bs"] numbound0_I[OF bn,where b="real_of_int i" and b'="real_of_int i" and bs="bs"] by (simp add: isint_iff) from 10 have id: "j dvd d" by simp from ie[symmetric] have "?p i = (\ (real_of_int j rdvd real_of_int (c*i+ floor ?e)))" by simp - also have "\ = Not (j dvd c*i + floor ?e)" + also have "\ = Not (j dvd c*i + floor ?e)" using int_rdvd_iff [where i="j" and t="c*i+ floor ?e"] by simp - also have "\ = Not (j dvd c*i - c*d + floor ?e)" + also have "\ = Not (j dvd c*i - c*d + floor ?e)" using dvd_period[OF id, where x="c*i" and c="-c" and t="floor ?e"] by simp - also have "\ = Not (real_of_int j rdvd real_of_int (c*i - c*d + floor ?e))" + also have "\ = Not (real_of_int j rdvd real_of_int (c*i - c*d + floor ?e))" using int_rdvd_iff[where i="j" and t="(c*i - c*d + floor ?e)",symmetric, simplified] ie by simp - also have "\ = Not (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)" + also have "\ = Not (real_of_int j rdvd real_of_int (c*(i - d)) + ?e)" using ie by (simp add:algebra_simps) - finally show ?case - using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p + finally show ?case + using numbound0_I[OF bn,where b="real_of_int i - real_of_int d" and b'="real_of_int i" and bs="bs"] p by (simp add: algebra_simps) qed (auto simp add: numbound0_I[where bs="bs" and b="real_of_int i - real_of_int d" and b'="real_of_int i"]) lemma \_nb: assumes lp: "iszlfm p (a#bs)" and nb: "numbound0 t" shows "bound0 (\ p k t)" using \_\_nb[OF lp nb] nb by (simp add: \_def) - + lemma \': assumes lp: "iszlfm p (a #bs)" and d: "d_\ p d" and dp: "d > 0" shows "\ x. \(\ (e,c) \ set(\ p). \(j::int) \ {1 .. c*d}. Ifm (a #bs) (\ p c (Add e (C j)))) \ Ifm (real_of_int x#bs) p \ Ifm (real_of_int (x - d)#bs) p" (is "\ x. ?b x \ ?P x \ ?P (x - d)") proof(clarify) - fix x - assume nob1:"?b x" and px: "?P x" + fix x + assume nob1:"?b x" and px: "?P x" from iszlfm_gen[OF lp, rule_format, where y="real_of_int x"] have lp': "iszlfm p (real_of_int x#bs)". - have nob: "\(e, c)\set (\ p). \j\{1..c * d}. real_of_int (c * x) \ Inum (real_of_int x # bs) e + real_of_int j" + have nob: "\(e, c)\set (\ p). \j\{1..c * d}. real_of_int (c * x) \ Inum (real_of_int x # bs) e + real_of_int j" proof(clarify) fix e c j assume ecR: "(e,c) \ set (\ p)" and jD: "j\ {1 .. c*d}" and cx: "real_of_int (c*x) = Inum (real_of_int x#bs) e + real_of_int j" @@ -3220,14 +3220,14 @@ have "Ifm (a#bs) (\ p c (Add e (C j)))" by blast with ecR jD nob1 show "False" by blast qed - from \[OF lp' px d dp nob] show "?P (x -d )" . + from \[OF lp' px d dp nob] show "?P (x -d )" . qed -lemma rl_thm: +lemma rl_thm: assumes lp: "iszlfm p (real_of_int (i::int)#bs)" shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ (e,c) \ set (\ p). \ j\ {1 .. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" - (is "(\(x::int). ?P x) = ((\ j\ {1.. \ p}. ?MP j)\(\ (e,c) \ ?R. \ j\ _. ?SP c e j))" + (is "(\(x::int). ?P x) = ((\ j\ {1.. \ p}. ?MP j)\(\ (e,c) \ ?R. \ j\ _. ?SP c e j))" is "?lhs = (?MD \ ?RD)" is "?lhs = ?rhs") proof- let ?d= "\ p" @@ -3244,9 +3244,9 @@ from nb have nb': "numbound0 (Add e (C j))" by simp from spx bound0_I[OF \_nb[OF lp nb', where k="c"], where bs="bs" and b="a" and b'="real_of_int i"] have spx': "Ifm (real_of_int i # bs) (\ p c (Add e (C j)))" by blast - from spx' have rcdej:"real_of_int c rdvd (Inum (real_of_int i#bs) (Add e (C j)))" + from spx' have rcdej:"real_of_int c rdvd (Inum (real_of_int i#bs) (Add e (C j)))" and sr:"Ifm (real_of_int i#bs) (\_\ p (Add e (C j),c))" by (simp add: \_def)+ - from rcdej eji[simplified isint_iff] + from rcdej eji[simplified isint_iff] have "real_of_int c rdvd real_of_int (floor (Inum (real_of_int i#bs) (Add e (C j))))" by simp hence cdej:"c dvd floor (Inum (real_of_int i#bs) (Add e (C j)))" by (simp only: int_rdvd_iff) from cp have cp': "real_of_int c > 0" by simp @@ -3260,7 +3260,7 @@ from \'[OF lp' d dp, rule_format, OF nob] have th:"\ x. ?P x \ ?P (x - ?d)" by blast from minusinf_inf[OF lp] obtain z where z:"\ x 0" by arith - from decr_lemma[OF dp,where x="x" and z="z"] + from decr_lemma[OF dp,where x="x" and z="z"] decr_mult_lemma[OF dp th zp, rule_format, OF px] z have th:"\ x. ?MP x" by auto with minusinf_bex[OF lp] px nob have ?thesis by blast} ultimately show ?thesis by blast @@ -3270,15 +3270,15 @@ shows "(\ (t,k). (Inum (a#bs) t, k)) ` set (\_\ p) = (\ (t,k). (Inum (a#bs) t,k)) ` set (\ (mirror p))" using lp by (induct p rule: mirror.induct) (simp_all add: split_def image_Un) - + text \The \\\ part\ text\Linearity for fm where Bound 0 ranges over \\\\ consts isrlfm :: "fm \ bool" (* Linearity test for fm *) recdef isrlfm "measure size" - "isrlfm (And p q) = (isrlfm p \ isrlfm q)" - "isrlfm (Or p q) = (isrlfm p \ isrlfm q)" + "isrlfm (And p q) = (isrlfm p \ isrlfm q)" + "isrlfm (Or p q) = (isrlfm p \ isrlfm q)" "isrlfm (Eq (CN 0 c e)) = (c>0 \ numbound0 e)" "isrlfm (NEq (CN 0 c e)) = (c>0 \ numbound0 e)" "isrlfm (Lt (CN 0 c e)) = (c>0 \ numbound0 e)" @@ -3288,21 +3288,21 @@ "isrlfm p = (isatom p \ (bound0 p))" definition fp :: "fm \ int \ num \ int \ fm" where - "fp p n s j \ (if n > 0 then + "fp p n s j \ (if n > 0 then (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j))))) (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1)))))))) - else - (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) + else + (And p (And (Le (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C j))))) (Gt (CN 0 (-n) (Add (Neg s) (Add (Floor s) (C (j + 1)))))))))" (* splits the bounded from the unbounded part*) function (sequential) rsplit0 :: "num \ (fm \ int \ num) list" where "rsplit0 (Bound 0) = [(T,1,C 0)]" -| "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b +| "rsplit0 (Add a b) = (let acs = rsplit0 a ; bcs = rsplit0 b in map (\ ((p,n,t),(q,m,s)). (And p q, n+m, Add t s)) [(a,b). a\acs,b\bcs])" | "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))" | "rsplit0 (Neg a) = map (\ (p,n,s). (p,-n,Neg s)) (rsplit0 a)" -| "rsplit0 (Floor a) = concat (map +| "rsplit0 (Floor a) = concat (map (\ (p,n,s). if n=0 then [(p,0,Floor s)] else (map (\ j. (fp p n s j, 0, Add (Floor s) (C j))) (if n > 0 then [0 .. n] else [n .. 0]))) (rsplit0 a))" @@ -3321,22 +3321,22 @@ lemma rsplit0_cs: - shows "\ (p,n,s) \ set (rsplit0 t). - (Ifm (x#bs) p \ (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \ numbound0 s \ isrlfm p" + shows "\ (p,n,s) \ set (rsplit0 t). + (Ifm (x#bs) p \ (Inum (x#bs) t = Inum (x#bs) (CN 0 n s))) \ numbound0 s \ isrlfm p" (is "\ (p,n,s) \ ?SS t. (?I p \ ?N t = ?N (CN 0 n s)) \ _ \ _ ") proof(induct t rule: rsplit0.induct) - case (5 a) + case (5 a) let ?p = "\ (p,n,s) j. fp p n s j" let ?f = "(\ (p,n,s) j. (?p (p,n,s) j, (0::int),Add (Floor s) (C j)))" let ?J = "\ n. if n>0 then [0..n] else [n..0]" let ?ff=" (\ (p,n,s). if n= 0 then [(p,0,Floor s)] else map (?f (p,n,s)) (?J n))" have int_cases: "\ (i::int). i= 0 \ i < 0 \ i > 0" by arith - have U1: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) = + have U1: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)]))" by auto - have U2': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n>0}. + have U2': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n>0}. ?ff (p,n,s) = map (?f(p,n,s)) [0..n]" by auto - hence U2: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) = - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). + hence U2: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) = + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) [0..n])))" proof- fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g @@ -3346,7 +3346,7 @@ qed have U3': "\ (p,n,s) \ {(p,n,s). (p,n,s) \ ?SS a \ n<0}. ?ff (p,n,s) = map (?f(p,n,s)) [n..0]" by auto - hence U3: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))) = + hence U3: "(UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))) = (UNION {(p,n,s). (p,n,s)\ ?SS a\n<0} (\(p,n,s). set (map (?f(p,n,s)) [n..0])))" proof- fix M :: "('a\'b\'c) set" and f :: "('a\'b\'c) \ 'd list" and g @@ -3357,29 +3357,29 @@ have "?SS (Floor a) = UNION (?SS a) (\x. set (?ff x))" by auto also have "\ = UNION (?SS a) (\ (p,n,s). set (?ff (p,n,s)))" by blast - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" using int_cases[rule_format] by blast - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un - (UNION {(p,n,s). (p,n,s)\ ?SS a\n>0} (\(p,n,s). set(map(?f(p,n,s)) [0..n]))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un + (UNION {(p,n,s). (p,n,s)\ ?SS a\n>0} (\(p,n,s). set(map(?f(p,n,s)) [0..n]))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) [n..0]))))" by (simp only: U1 U2 U3) - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). (?f(p,n,s)) ` {n .. 0})))" by (simp only: set_map set_upto list.set) - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast - finally - have FS: "?SS (Floor a) = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un + finally + have FS: "?SS (Floor a) = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast show ?case proof(simp only: FS, clarsimp simp del: Ifm.simps Inum.simps, -) @@ -3396,17 +3396,17 @@ ac < 0 \ (\j. p = fp ab ac ba j \ n = 0 \ s = Add (Floor ba) (C j) \ ac \ j \ j \ 0))" - moreover + moreover { fix s' assume "(p, 0, s') \ ?SS a" and "n = 0" and "s = Floor s'" hence ?ths using 5(1) by auto } moreover { fix p' n' s' j - assume pns: "(p', n', s') \ ?SS a" - and np: "0 < n'" - and p_def: "p = ?p (p',n',s') j" - and n0: "n = 0" - and s_def: "s = (Add (Floor s') (C j))" + assume pns: "(p', n', s') \ ?SS a" + and np: "0 < n'" + and p_def: "p = ?p (p',n',s') j" + and n0: "n = 0" + and s_def: "s = (Add (Floor s') (C j))" and jp: "0 \ j" and jn: "j \ n'" from 5 pns have H:"(Ifm ((x::real) # (bs::real list)) p' \ Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ @@ -3415,9 +3415,9 @@ from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by simp let ?nxs = "CN 0 n' s'" let ?l = "floor (?N s') + j" - from H - have "?I (?p (p',n',s') j) \ - (((?N ?nxs \ real_of_int ?l) \ (?N ?nxs < real_of_int (?l + 1))) \ (?N a = ?N ?nxs ))" + from H + have "?I (?p (p',n',s') j) \ + (((?N ?nxs \ real_of_int ?l) \ (?N ?nxs < real_of_int (?l + 1))) \ (?N a = ?N ?nxs ))" by (simp add: fp_def np algebra_simps) also have "\ \ ((floor (?N ?nxs) = ?l) \ (?N a = ?N ?nxs ))" using floor_unique_iff[where x="?N ?nxs" and a="?l"] by simp @@ -3428,11 +3428,11 @@ with s_def n0 p_def nb nf have ?ths by auto} moreover { fix p' n' s' j - assume pns: "(p', n', s') \ ?SS a" - and np: "n' < 0" - and p_def: "p = ?p (p',n',s') j" - and n0: "n = 0" - and s_def: "s = (Add (Floor s') (C j))" + assume pns: "(p', n', s') \ ?SS a" + and np: "n' < 0" + and p_def: "p = ?p (p',n',s') j" + and n0: "n = 0" + and s_def: "s = (Add (Floor s') (C j))" and jp: "n' \ j" and jn: "j \ 0" from 5 pns have H:"(Ifm ((x::real) # (bs::real list)) p' \ Inum (x # bs) a = Inum (x # bs) (CN 0 n' s')) \ @@ -3441,9 +3441,9 @@ from H have nf: "isrlfm (?p (p',n',s') j)" using fp_def np by simp let ?nxs = "CN 0 n' s'" let ?l = "floor (?N s') + j" - from H - have "?I (?p (p',n',s') j) \ - (((?N ?nxs \ real_of_int ?l) \ (?N ?nxs < real_of_int (?l + 1))) \ (?N a = ?N ?nxs ))" + from H + have "?I (?p (p',n',s') j) \ + (((?N ?nxs \ real_of_int ?l) \ (?N ?nxs < real_of_int (?l + 1))) \ (?N a = ?N ?nxs ))" by (simp add: np fp_def algebra_simps) also have "\ \ ((floor (?N ?nxs) = ?l) \ (?N a = ?N ?nxs ))" using floor_unique_iff[where x="?N ?nxs" and a="?l"] by simp @@ -3459,18 +3459,18 @@ by auto qed (auto simp add: Let_def split_def algebra_simps) -lemma real_in_int_intervals: +lemma real_in_int_intervals: assumes xb: "real_of_int m \ x \ x < real_of_int ((n::int) + 1)" shows "\ j\ {m.. n}. real_of_int j \ x \ x < real_of_int (j+1)" (is "\ j\ ?N. ?P j") -by (rule bexI[where P="?P" and x="floor x" and A="?N"]) -(auto simp add: floor_less_iff[where x="x" and z="n+1", simplified] +by (rule bexI[where P="?P" and x="floor x" and A="?N"]) +(auto simp add: floor_less_iff[where x="x" and z="n+1", simplified] xb[simplified] floor_mono[where x="real_of_int m" and y="x", OF conjunct1[OF xb], simplified floor_of_int[where z="m"]]) lemma rsplit0_complete: assumes xp:"0 \ x" and x1:"x < 1" shows "\ (p,n,s) \ set (rsplit0 t). Ifm (x#bs) p" (is "\ (p,n,s) \ ?SS t. ?I p") proof(induct t rule: rsplit0.induct) - case (2 a b) + case (2 a b) then have "\ (pa,na,sa) \ ?SS a. ?I pa" by auto then obtain "pa" "na" "sa" where pa: "(pa,na,sa)\ ?SS a \ ?I pa" by blast with 2 have "\ (pb,nb,sb) \ ?SS b. ?I pb" by blast @@ -3484,7 +3484,7 @@ moreover from pa pb have "?I (And pa pb)" by simp ultimately show ?case by blast next - case (5 a) + case (5 a) let ?p = "\ (p,n,s) j. fp p n s j" let ?f = "(\ (p,n,s) j. (?p (p,n,s) j, (0::int),(Add (Floor s) (C j))))" let ?J = "\ n. if n>0 then [0..n] else [n..0]" @@ -3512,30 +3512,30 @@ have "?SS (Floor a) = UNION (?SS a) (\x. set (?ff x))" by auto also have "\ = UNION (?SS a) (\ (p,n,s). set (?ff (p,n,s)))" by blast - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set (?ff (p,n,s)))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (?ff (p,n,s)))) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (?ff (p,n,s)))))" using int_cases[rule_format] by blast - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) [0..n]))) Un + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). set [(p,0,Floor s)])) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). set (map (?f(p,n,s)) [0..n]))) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). set (map (?f(p,n,s)) [n..0]))))" by (simp only: U1 U2 U3) - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). (?f(p,n,s)) ` {0 .. n})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). (?f(p,n,s)) ` {n .. 0})))" by (simp only: set_map set_upto list.set) - also have "\ = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un + also have "\ = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast - finally - have FS: "?SS (Floor a) = - ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un - (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un + finally + have FS: "?SS (Floor a) = + ((UNION {(p,n,s). (p,n,s) \ ?SS a \ n=0} (\ (p,n,s). {(p,0,Floor s)})) Un + (UNION {(p,n,s). (p,n,s) \ ?SS a \ n>0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {0 .. n}})) Un (UNION {(p,n,s). (p,n,s) \ ?SS a \ n<0} (\ (p,n,s). {?f(p,n,s) j| j. j\ {n .. 0}})))" by blast from 5 have "\ (p,n,s) \ ?SS a. ?I p" by auto @@ -3543,7 +3543,7 @@ let ?N = "\ t. Inum (x#bs) t" from rsplit0_cs[rule_format] pns have ans:"(?N a = ?N (CN 0 n s)) \ numbound0 s \ isrlfm p" by auto - + have "n=0 \ n >0 \ n <0" by arith moreover {assume "n=0" hence ?case using pns by (simp only: FS) auto } moreover @@ -3554,21 +3554,21 @@ finally have "?N (Floor s) \ real_of_int n * x + ?N s" . moreover {from x1 np have "real_of_int n *x + ?N s < real_of_int n + ?N s" by simp - also from real_of_int_floor_add_one_gt[where r="?N s"] + also from real_of_int_floor_add_one_gt[where r="?N s"] have "\ < real_of_int n + ?N (Floor s) + 1" by simp finally have "real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (n+1)" by simp} ultimately have "?N (Floor s) \ real_of_int n *x + ?N s\ real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (n+1)" by simp hence th: "0 \ real_of_int n *x + ?N s - ?N (Floor s) \ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (n+1)" by simp from real_in_int_intervals th have "\ j\ {0 .. n}. real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s)\ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (j+1)" by simp - + hence "\ j\ {0 .. n}. 0 \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1) < 0" - by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"]) + by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"]) hence "\ j\ {0.. n}. ?I (?p (p,n,s) j)" using pns by (simp add: fp_def np algebra_simps) then obtain "j" where j_def: "j\ {0 .. n} \ ?I (?p (p,n,s) j)" by blast hence "\x \ {?p (p,n,s) j |j. 0\ j \ j \ n }. ?I x" by auto - hence ?case using pns - by (simp only: FS,simp add: bex_Un) + hence ?case using pns + by (simp only: FS,simp add: bex_Un) (rule disjI2, rule disjI1,rule exI [where x="p"], rule exI [where x="n"],rule exI [where x="s"],simp_all add: np) } @@ -3576,27 +3576,27 @@ { assume nn: "n < 0" hence np: "-n >0" by simp from of_int_floor_le[of "?N s"] have "?N (Floor s) + 1 > ?N s" by simp moreover from mult_left_mono_neg[OF xp] nn have "?N s \ real_of_int n * x + ?N s" by simp - ultimately have "?N (Floor s) + 1 > real_of_int n * x + ?N s" by arith + ultimately have "?N (Floor s) + 1 > real_of_int n * x + ?N s" by arith moreover {from x1 nn have "real_of_int n *x + ?N s \ real_of_int n + ?N s" by simp moreover from of_int_floor_le[of "?N s"] have "real_of_int n + ?N s \ real_of_int n + ?N (Floor s)" by simp - ultimately have "real_of_int n *x + ?N s \ ?N (Floor s) + real_of_int n" + ultimately have "real_of_int n *x + ?N s \ ?N (Floor s) + real_of_int n" by (simp only: algebra_simps)} ultimately have "?N (Floor s) + real_of_int n \ real_of_int n *x + ?N s\ real_of_int n *x + ?N s < ?N (Floor s) + real_of_int (1::int)" by simp hence th: "real_of_int n \ real_of_int n *x + ?N s - ?N (Floor s) \ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (1::int)" by simp have th1: "\ (a::real). (- a > 0) = (a < 0)" by auto have th2: "\ (a::real). (0 \ - a) = (a \ 0)" by auto from real_in_int_intervals th have "\ j\ {n .. 0}. real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s)\ real_of_int n *x + ?N s - ?N (Floor s) < real_of_int (j+1)" by simp - + hence "\ j\ {n .. 0}. 0 \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j \ real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1) < 0" - by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"]) + by(simp only: myle[of _ "real_of_int n * x + Inum (x # bs) s - Inum (x # bs) (Floor s)"] less_iff_diff_less_0[where a="real_of_int n *x + ?N s - ?N (Floor s)"]) hence "\ j\ {n .. 0}. 0 \ - (real_of_int n *x + ?N s - ?N (Floor s) - real_of_int j) \ - (real_of_int n *x + ?N s - ?N (Floor s) - real_of_int (j+1)) > 0" by (simp only: th1[rule_format] th2[rule_format]) hence "\ j\ {n.. 0}. ?I (?p (p,n,s) j)" using pns by (simp add: fp_def nn algebra_simps - del: diff_less_0_iff_less diff_le_0_iff_le) + del: diff_less_0_iff_less diff_le_0_iff_le) then obtain "j" where j_def: "j\ {n .. 0} \ ?I (?p (p,n,s) j)" by blast hence "\x \ {?p (p,n,s) j |j. n\ j \ j \ 0 }. ?I x" by auto - hence ?case using pns + hence ?case using pns by (simp only: FS,simp add: bex_Un) (rule disjI2, rule disjI2,rule exI [where x="p"], rule exI [where x="n"],rule exI [where x="s"],simp_all add: nn) @@ -3615,7 +3615,7 @@ lemma foldr_conj_map: "Ifm bs (foldr conj (map f xs) T) = (\ x \ set xs. Ifm bs (f x))" by(induct xs, simp_all) -lemma foldr_disj_map_rlfm: +lemma foldr_disj_map_rlfm: assumes lf: "\ n s. numbound0 s \ isrlfm (f n s)" and \: "\ (\,n,s) \ set xs. numbound0 s \ isrlfm \" shows "isrlfm (foldr disj (map (\ (\, n, s). conj \ (f n s)) xs) F)" @@ -3631,7 +3631,7 @@ from foldr_disj_map_rlfm[OF lf th] rsplit_def show ?thesis by simp qed -lemma rsplit: +lemma rsplit: assumes xp: "x \ 0" and x1: "x < 1" and f: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ (Ifm (x#bs) (f n s) = Ifm (x#bs) (g a))" shows "Ifm (x#bs) (rsplit f a) = Ifm (x#bs) (g a)" @@ -3642,14 +3642,14 @@ hence "\ (\,n,s) \ set (rsplit0 a). ?I x (And \ (f n s))" using rsplit_ex by simp then obtain "\" "n" "s" where fnsS:"(\,n,s) \ set (rsplit0 a)" and "?I x (And \ (f n s))" by blast hence \: "?I x \" and fns: "?I x (f n s)" by auto - from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \ + from rsplit0_cs[where t="a" and bs="bs" and x="x", rule_format, OF fnsS] \ have th: "(?N x a = ?N x (CN 0 n s)) \ numbound0 s" by auto from f[rule_format, OF th] fns show "?I x (g a)" by simp next let ?I = "\x p. Ifm (x#bs) p" let ?N = "\ x t. Inum (x#bs) t" assume ga: "?I x (g a)" - from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] + from rsplit0_complete[OF xp x1, where bs="bs" and t="a"] obtain "\" "n" "s" where fnsS:"(\,n,s) \ set (rsplit0 a)" and fx: "?I x \" by blast from rsplit0_cs[where t="a" and x="x" and bs="bs"] fnsS fx have ans: "?N x a = ?N x (CN 0 n s)" and nb: "numbound0 s" by auto @@ -3658,27 +3658,27 @@ qed definition lt :: "int \ num \ fm" where - lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) + lt_def: "lt c t = (if c = 0 then (Lt t) else if c > 0 then (Lt (CN 0 c t)) else (Gt (CN 0 (-c) (Neg t))))" definition le :: "int \ num \ fm" where - le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) + le_def: "le c t = (if c = 0 then (Le t) else if c > 0 then (Le (CN 0 c t)) else (Ge (CN 0 (-c) (Neg t))))" definition gt :: "int \ num \ fm" where - gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) + gt_def: "gt c t = (if c = 0 then (Gt t) else if c > 0 then (Gt (CN 0 c t)) else (Lt (CN 0 (-c) (Neg t))))" definition ge :: "int \ num \ fm" where - ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) + ge_def: "ge c t = (if c = 0 then (Ge t) else if c > 0 then (Ge (CN 0 c t)) else (Le (CN 0 (-c) (Neg t))))" definition eq :: "int \ num \ fm" where - eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) + eq_def: "eq c t = (if c = 0 then (Eq t) else if c > 0 then (Eq (CN 0 c t)) else (Eq (CN 0 (-c) (Neg t))))" definition neq :: "int \ num \ fm" where - neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) + neq_def: "neq c t = (if c = 0 then (NEq t) else if c > 0 then (NEq (CN 0 c t)) else (NEq (CN 0 (-c) (Neg t))))" lemma lt_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (lt n s) = Ifm (x#bs) (Lt a)" @@ -3703,7 +3703,7 @@ qed lemma le_l: "isrlfm (rsplit le a)" - by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) + by (rule rsplit_l[where f="le" and a="a"], auto simp add: le_def) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat",simp_all) lemma gt_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (gt n s) = Ifm (x#bs) (Gt a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (gt n s) = ?I (Gt a)") @@ -3714,28 +3714,28 @@ (cases "n > 0", simp_all add: gt_def algebra_simps myless[of _ "0"]) qed lemma gt_l: "isrlfm (rsplit gt a)" - by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) + by (rule rsplit_l[where f="gt" and a="a"], auto simp add: gt_def) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) lemma ge_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (ge n s) = Ifm (x#bs) (Ge a)" (is "\ a n s . ?N a = ?N (CN 0 n s) \ _ \ ?I (ge n s) = ?I (Ge a)") proof(clarify) - fix a n s + fix a n s assume H: "?N a = ?N (CN 0 n s)" show "?I (ge n s) = ?I (Ge a)" using H by (cases "n=0", (simp add: ge_def)) (cases "n > 0", simp_all add: ge_def algebra_simps myle[of _ "0"]) qed lemma ge_l: "isrlfm (rsplit ge a)" - by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) + by (rule rsplit_l[where f="ge" and a="a"], auto simp add: ge_def) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) lemma eq_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (eq n s) = Ifm (x#bs) (Eq a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (eq n s) = ?I (Eq a)") proof(clarify) - fix a n s + fix a n s assume H: "?N a = ?N (CN 0 n s)" show "?I (eq n s) = ?I (Eq a)" using H by (auto simp add: eq_def algebra_simps) qed lemma eq_l: "isrlfm (rsplit eq a)" - by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) + by (rule rsplit_l[where f="eq" and a="a"], auto simp add: eq_def) (case_tac s, simp_all, rename_tac nat a b, case_tac"nat", simp_all) lemma neq_mono: "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (neq n s) = Ifm (x#bs) (NEq a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (neq n s) = ?I (NEq a)") @@ -3746,35 +3746,35 @@ qed lemma neq_l: "isrlfm (rsplit neq a)" - by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) + by (rule rsplit_l[where f="neq" and a="a"], auto simp add: neq_def) (case_tac s, simp_all, rename_tac nat a b, case_tac"nat", simp_all) -lemma small_le: +lemma small_le: assumes u0:"0 \ u" and u1: "u < 1" shows "(-u \ real_of_int (n::int)) = (0 \ n)" using u0 u1 by auto -lemma small_lt: +lemma small_lt: assumes u0:"0 \ u" and u1: "u < 1" shows "(real_of_int (n::int) < real_of_int (m::int) - u) = (n < m)" using u0 u1 by auto -lemma rdvd01_cs: +lemma rdvd01_cs: assumes up: "u \ 0" and u1: "u<1" and np: "real_of_int n > 0" shows "(real_of_int (i::int) rdvd real_of_int (n::int) * u - s) = (\ j\ {0 .. n - 1}. real_of_int n * u = s - real_of_int (floor s) + real_of_int j \ real_of_int i rdvd real_of_int (j - floor s))" (is "?lhs = ?rhs") proof- let ?ss = "s - real_of_int (floor s)" - from real_of_int_floor_add_one_gt[where r="s", simplified myless[of "s"]] - of_int_floor_le have ss0:"?ss \ 0" and ss1:"?ss < 1" + from real_of_int_floor_add_one_gt[where r="s", simplified myless[of "s"]] + of_int_floor_le have ss0:"?ss \ 0" and ss1:"?ss < 1" by (auto simp add: myle[of _ "s", symmetric] myless[of "?ss"]) from np have n0: "real_of_int n \ 0" by simp - from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] - have nu0:"real_of_int n * u - s \ -s" and nun:"real_of_int n * u -s < real_of_int n - s" by auto - from int_rdvd_real[where i="i" and x="real_of_int (n::int) * u - s"] - have "real_of_int i rdvd real_of_int n * u - s = - (i dvd floor (real_of_int n * u -s) \ (real_of_int (floor (real_of_int n * u - s)) = real_of_int n * u - s ))" + from mult_left_mono[OF up n0] mult_strict_left_mono[OF u1 np] + have nu0:"real_of_int n * u - s \ -s" and nun:"real_of_int n * u -s < real_of_int n - s" by auto + from int_rdvd_real[where i="i" and x="real_of_int (n::int) * u - s"] + have "real_of_int i rdvd real_of_int n * u - s = + (i dvd floor (real_of_int n * u -s) \ (real_of_int (floor (real_of_int n * u - s)) = real_of_int n * u - s ))" (is "_ = (?DE)" is "_ = (?D \ ?E)") by simp - also have "\ = (?DE \ real_of_int(floor (real_of_int n * u - s) + floor s)\ -?ss + also have "\ = (?DE \ real_of_int(floor (real_of_int n * u - s) + floor s)\ -?ss \ real_of_int(floor (real_of_int n * u - s) + floor s)< real_of_int n - ?ss)" (is "_=(?DE \real_of_int ?a \ _ \ real_of_int ?a < _)") using nu0 nun by auto also have "\ = (?DE \ ?a \ 0 \ ?a < n)" by(simp only: small_le[OF ss0 ss1] small_lt[OF ss0 ss1]) @@ -3797,43 +3797,43 @@ where NDVDJ_def: "NDVDJ i n s = (foldr conj (map (\ j. disj (NEq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor (Neg s))))) [0..n - 1]) T)" -lemma DVDJ_DVD: +lemma DVDJ_DVD: assumes xp:"x\ 0" and x1: "x < 1" and np:"real_of_int n > 0" shows "Ifm (x#bs) (DVDJ i n s) = Ifm (x#bs) (Dvd i (CN 0 n s))" proof- let ?f = "\ j. conj (Eq(CN 0 n (Add s (Sub(Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))" let ?s= "Inum (x#bs) s" from foldr_disj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"] - have "Ifm (x#bs) (DVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" + have "Ifm (x#bs) (DVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" by (simp add: np DVDJ_def) also have "\ = (\ j\ {0 .. (n - 1)}. real_of_int n * x = (- ?s) - real_of_int (floor (- ?s)) + real_of_int j \ real_of_int i rdvd real_of_int (j - floor (- ?s)))" by (simp add: algebra_simps) - also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] + also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] have "\ = (real_of_int i rdvd real_of_int n * x - (-?s))" by simp finally show ?thesis by simp qed -lemma NDVDJ_NDVD: +lemma NDVDJ_NDVD: assumes xp:"x\ 0" and x1: "x < 1" and np:"real_of_int n > 0" shows "Ifm (x#bs) (NDVDJ i n s) = Ifm (x#bs) (NDvd i (CN 0 n s))" proof- let ?f = "\ j. disj(NEq(CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (NDvd i (Sub (C j) (Floor(Neg s))))" let ?s= "Inum (x#bs) s" from foldr_conj_map[where xs="[0..n - 1]" and bs="x#bs" and f="?f"] - have "Ifm (x#bs) (NDVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" + have "Ifm (x#bs) (NDVDJ i n s) = (\ j\ {0 .. (n - 1)}. Ifm (x#bs) (?f j))" by (simp add: np NDVDJ_def) also have "\ = (\ (\ j\ {0 .. (n - 1)}. real_of_int n * x = (- ?s) - real_of_int (floor (- ?s)) + real_of_int j \ real_of_int i rdvd real_of_int (j - floor (- ?s))))" by (simp add: algebra_simps) - also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] + also from rdvd01_cs[OF xp x1 np, where i="i" and s="-?s"] have "\ = (\ (real_of_int i rdvd real_of_int n * x - (-?s)))" by simp finally show ?thesis by simp -qed - -lemma foldr_disj_map_rlfm2: +qed + +lemma foldr_disj_map_rlfm2: assumes lf: "\ n . isrlfm (f n)" shows "isrlfm (foldr disj (map f xs) F)" using lf by (induct xs, auto) -lemma foldr_And_map_rlfm2: +lemma foldr_And_map_rlfm2: assumes lf: "\ n . isrlfm (f n)" shows "isrlfm (foldr conj (map f xs) T)" using lf by (induct xs, auto) @@ -3844,7 +3844,7 @@ let ?f="\j. conj (Eq (CN 0 n (Add s (Sub (Floor (Neg s)) (C j))))) (Dvd i (Sub (C j) (Floor (Neg s))))" have th: "\ j. isrlfm (?f j)" using nb np by auto - from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp + from DVDJ_def foldr_disj_map_rlfm2[OF th] show ?thesis by simp qed lemma NDVDJ_l: assumes ip: "i >0" and np: "n>0" and nb: "numbound0 s" @@ -3858,58 +3858,58 @@ definition DVD :: "int \ int \ num \ fm" where DVD_def: "DVD i c t = - (if i=0 then eq c t else + (if i=0 then eq c t else if c = 0 then (Dvd i t) else if c >0 then DVDJ (abs i) c t else DVDJ (abs i) (-c) (Neg t))" definition NDVD :: "int \ int \ num \ fm" where "NDVD i c t = - (if i=0 then neq c t else + (if i=0 then neq c t else if c = 0 then (NDvd i t) else if c >0 then NDVDJ (abs i) c t else NDVDJ (abs i) (-c) (Neg t))" -lemma DVD_mono: - assumes xp: "0\ x" and x1: "x < 1" +lemma DVD_mono: + assumes xp: "0\ x" and x1: "x < 1" shows "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (DVD i n s) = Ifm (x#bs) (Dvd i a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (DVD i n s) = ?I (Dvd i a)") proof(clarify) - fix a n s + fix a n s assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" let ?th = "?I (DVD i n s) = ?I (Dvd i a)" have "i=0 \ (i\0 \ n=0) \ (i\0 \ n < 0) \ (i\0 \ n > 0)" by arith - moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] + moreover {assume iz: "i=0" hence ?th using eq_mono[rule_format, OF conjI[OF H nb]] by (simp add: DVD_def rdvd_left_0_eq)} - moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H DVD_def) } - moreover {assume inz: "i\0" and "n<0" hence ?th - by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 - rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) } + moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H DVD_def) } + moreover {assume inz: "i\0" and "n<0" hence ?th + by (simp add: DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1 + rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) } moreover {assume inz: "i\0" and "n>0" hence ?th by (simp add:DVD_def H DVDJ_DVD[OF xp x1] rdvd_abs1)} ultimately show ?th by blast qed -lemma NDVD_mono: assumes xp: "0\ x" and x1: "x < 1" +lemma NDVD_mono: assumes xp: "0\ x" and x1: "x < 1" shows "\ a n s. Inum (x#bs) a = Inum (x#bs) (CN 0 n s) \ numbound0 s \ Ifm (x#bs) (NDVD i n s) = Ifm (x#bs) (NDvd i a)" (is "\ a n s. ?N a = ?N (CN 0 n s) \ _ \ ?I (NDVD i n s) = ?I (NDvd i a)") proof(clarify) - fix a n s + fix a n s assume H: "?N a = ?N (CN 0 n s)" and nb: "numbound0 s" let ?th = "?I (NDVD i n s) = ?I (NDvd i a)" have "i=0 \ (i\0 \ n=0) \ (i\0 \ n < 0) \ (i\0 \ n > 0)" by arith - moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] + moreover {assume iz: "i=0" hence ?th using neq_mono[rule_format, OF conjI[OF H nb]] by (simp add: NDVD_def rdvd_left_0_eq)} - moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H NDVD_def) } - moreover {assume inz: "i\0" and "n<0" hence ?th - by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 - rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) } - moreover {assume inz: "i\0" and "n>0" hence ?th + moreover {assume inz: "i\0" and "n=0" hence ?th by (simp add: H NDVD_def) } + moreover {assume inz: "i\0" and "n<0" hence ?th + by (simp add: NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1 + rdvd_minus[where d="i" and t="real_of_int n * x + Inum (x # bs) s"]) } + moreover {assume inz: "i\0" and "n>0" hence ?th by (simp add:NDVD_def H NDVDJ_NDVD[OF xp x1] rdvd_abs1)} ultimately show ?th by blast qed lemma DVD_l: "isrlfm (rsplit (DVD i) a)" - by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) + by (rule rsplit_l[where f="DVD i" and a="a"], auto simp add: DVD_def eq_def DVDJ_l) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) lemma NDVD_l: "isrlfm (rsplit (NDVD i) a)" - by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) + by (rule rsplit_l[where f="NDVD i" and a="a"], auto simp add: NDVD_def neq_def NDVDJ_l) (case_tac s, simp_all, rename_tac nat a b, case_tac "nat", simp_all) consts rlfm :: "fm \ fm" @@ -3948,20 +3948,20 @@ lemma simpfm_rl: "isrlfm p \ isrlfm (simpfm p)" proof (induct p) - case (Lt a) + case (Lt a) hence "bound0 (Lt a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover - {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))" + {assume "bound0 (Lt a)" hence bn:"bound0 (simpfm (Lt a))" using simpfm_bound0 by blast have "isatom (simpfm (Lt a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} - moreover + moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" + from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp @@ -3972,20 +3972,20 @@ by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next - case (Le a) + case (Le a) hence "bound0 (Le a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover - { assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))" + { assume "bound0 (Le a)" hence bn:"bound0 (simpfm (Le a))" using simpfm_bound0 by blast have "isatom (simpfm (Le a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} - moreover + moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" + from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp @@ -3996,20 +3996,20 @@ by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next - case (Gt a) + case (Gt a) hence "bound0 (Gt a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a, simp_all, rename_tac nat a b,case_tac "nat", simp_all) moreover - {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))" + {assume "bound0 (Gt a)" hence bn:"bound0 (simpfm (Gt a))" using simpfm_bound0 by blast have "isatom (simpfm (Gt a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} - moreover + moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1: "numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" + from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp @@ -4020,20 +4020,20 @@ by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next - case (Ge a) + case (Ge a) hence "bound0 (Ge a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover - { assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))" + { assume "bound0 (Ge a)" hence bn:"bound0 (simpfm (Ge a))" using simpfm_bound0 by blast have "isatom (simpfm (Ge a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} - moreover + moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" + from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp @@ -4044,20 +4044,20 @@ by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next - case (Eq a) + case (Eq a) hence "bound0 (Eq a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover - { assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))" + { assume "bound0 (Eq a)" hence bn:"bound0 (simpfm (Eq a))" using simpfm_bound0 by blast have "isatom (simpfm (Eq a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} - moreover + moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" + from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp @@ -4068,20 +4068,20 @@ by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next - case (NEq a) + case (NEq a) hence "bound0 (NEq a) \ (\ c e. a = CN 0 c e \ c > 0 \ numbound0 e)" by (cases a,simp_all, rename_tac nat a b, case_tac "nat", simp_all) moreover - {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))" + {assume "bound0 (NEq a)" hence bn:"bound0 (simpfm (NEq a))" using simpfm_bound0 by blast have "isatom (simpfm (NEq a))" by (cases "simpnum a", auto simp add: Let_def) with bn bound0at_l have ?case by blast} - moreover + moreover { fix c e assume a: "a = CN 0 c e" and "c>0" and "numbound0 e" { assume cn1:"numgcd (CN 0 c (simpnum e)) \ 1" and cnz:"numgcd (CN 0 c (simpnum e)) \ 0" with numgcd_pos[where t="CN 0 c (simpnum e)"] have th1:"numgcd (CN 0 c (simpnum e)) > 0" by simp - from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" + from \c > 0\ have th:"numgcd (CN 0 c (simpnum e)) \ c" by (simp add: numgcd_def) from \c > 0\ have th': "c\0" by auto from \c > 0\ have cp: "c \ 0" by simp @@ -4092,12 +4092,12 @@ by (simp add: Let_def reducecoeff_def reducecoeffh_numbound0)} ultimately show ?case by blast next - case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))" + case (Dvd i a) hence "bound0 (Dvd i a)" by auto hence bn:"bound0 (simpfm (Dvd i a))" using simpfm_bound0 by blast have "isatom (simpfm (Dvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) with bn bound0at_l show ?case by blast next - case (NDvd i a) hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))" + case (NDvd i a) hence "bound0 (NDvd i a)" by auto hence bn:"bound0 (simpfm (NDvd i a))" using simpfm_bound0 by blast have "isatom (simpfm (NDvd i a))" by (cases "simpnum a", auto simp add: Let_def split_def) with bn bound0at_l show ?case by blast @@ -4107,15 +4107,15 @@ assumes qfp: "qfree p" and xp: "0 \ x" and x1: "x < 1" shows "(Ifm (x#bs) (rlfm p) = Ifm (x# bs) p) \ isrlfm (rlfm p)" - using qfp -by (induct p rule: rlfm.induct) + using qfp +by (induct p rule: rlfm.induct) (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 rsplit[OF xp x1 ge_mono] ge_l rsplit[OF xp x1 eq_mono] eq_l rsplit[OF xp x1 neq_mono] neq_l rsplit[OF xp x1 DVD_mono[OF xp x1]] DVD_l rsplit[OF xp x1 NDVD_mono[OF xp x1]] NDVD_l simpfm_rl) lemma rlfm_l: assumes qfp: "qfree p" shows "isrlfm (rlfm p)" - using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l + using qfp lt_l gt_l ge_l le_l eq_l neq_l DVD_l NDVD_l by (induct p rule: rlfm.induct) (auto simp add: simpfm_rl) (* Operations needed for Ferrante and Rackoff *) @@ -4128,7 +4128,7 @@ next case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto next - case (3 c e) + case (3 c e) from 3 have nb: "numbound0 e" by simp from 3 have cp: "real_of_int c > 0" by simp fix a @@ -4136,8 +4136,8 @@ let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real_of_int c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) + hence "(real_of_int c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (Eq (CN 0 c e))" @@ -4145,7 +4145,7 @@ hence "\ x < ?z. ?P ?z x (Eq (CN 0 c e))" by simp thus ?case by blast next - case (4 c e) + case (4 c e) from 4 have nb: "numbound0 e" by simp from 4 have cp: "real_of_int c > 0" by simp fix a @@ -4153,8 +4153,8 @@ let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real_of_int c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) + hence "(real_of_int c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith hence "real_of_int c * x + ?e \ 0" by simp with xz have "?P ?z x (NEq (CN 0 c e))" @@ -4162,7 +4162,7 @@ hence "\ x < ?z. ?P ?z x (NEq (CN 0 c e))" by simp thus ?case by blast next - case (5 c e) + case (5 c e) from 5 have nb: "numbound0 e" by simp from 5 have cp: "real_of_int c > 0" by simp fix a @@ -4170,15 +4170,15 @@ let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real_of_int c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) + hence "(real_of_int c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Lt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Lt (CN 0 c e))" by simp thus ?case by blast next - case (6 c e) + case (6 c e) from 6 have nb: "numbound0 e" by simp from 6 have cp: "real_of_int c > 0" by simp fix a @@ -4186,15 +4186,15 @@ let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real_of_int c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) + hence "(real_of_int c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Le (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Le (CN 0 c e))" by simp thus ?case by blast next - case (7 c e) + case (7 c e) from 7 have nb: "numbound0 e" by simp from 7 have cp: "real_of_int c > 0" by simp fix a @@ -4202,15 +4202,15 @@ let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real_of_int c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) + hence "(real_of_int c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Gt (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } hence "\ x < ?z. ?P ?z x (Gt (CN 0 c e))" by simp thus ?case by blast next - case (8 c e) + case (8 c e) from 8 have nb: "numbound0 e" by simp from 8 have cp: "real_of_int c > 0" by simp fix a @@ -4218,8 +4218,8 @@ let ?z = "(- ?e) / real_of_int c" {fix x assume xz: "x < ?z" - hence "(real_of_int c * x < - ?e)" - by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) + hence "(real_of_int c * x < - ?e)" + by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] ac_simps) hence "real_of_int c * x + ?e < 0" by arith with xz have "?P ?z x (Ge (CN 0 c e))" using numbound0_I[OF nb, where b="x" and bs="bs" and b'="a"] by simp } @@ -4236,7 +4236,7 @@ next case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto next - case (3 c e) + case (3 c e) from 3 have nb: "numbound0 e" by simp from 3 have cp: "real_of_int c > 0" by simp fix a @@ -4253,7 +4253,7 @@ hence "\ x > ?z. ?P ?z x (Eq (CN 0 c e))" by simp thus ?case by blast next - case (4 c e) + case (4 c e) from 4 have nb: "numbound0 e" by simp from 4 have cp: "real_of_int c > 0" by simp fix a @@ -4270,7 +4270,7 @@ hence "\ x > ?z. ?P ?z x (NEq (CN 0 c e))" by simp thus ?case by blast next - case (5 c e) + case (5 c e) from 5 have nb: "numbound0 e" by simp from 5 have cp: "real_of_int c > 0" by simp fix a @@ -4286,7 +4286,7 @@ hence "\ x > ?z. ?P ?z x (Lt (CN 0 c e))" by simp thus ?case by blast next - case (6 c e) + case (6 c e) from 6 have nb: "numbound0 e" by simp from 6 have cp: "real_of_int c > 0" by simp fix a @@ -4302,7 +4302,7 @@ hence "\ x > ?z. ?P ?z x (Le (CN 0 c e))" by simp thus ?case by blast next - case (7 c e) + case (7 c e) from 7 have nb: "numbound0 e" by simp from 7 have cp: "real_of_int c > 0" by simp fix a @@ -4318,7 +4318,7 @@ hence "\ x > ?z. ?P ?z x (Gt (CN 0 c e))" by simp thus ?case by blast next - case (8 c e) + case (8 c e) from 8 have nb: "numbound0 e" by simp from 8 have cp: "real_of_int c > 0" by simp fix a @@ -4354,7 +4354,7 @@ proof- from bound0_I [OF rminusinf_bound0[OF lp], where b="a" and bs ="bs"] ex have th: "\ x. Ifm (x#bs) (minusinf p)" by auto - from rminusinf_inf[OF lp, where bs="bs"] + from rminusinf_inf[OF lp, where bs="bs"] obtain z where z_def: "\x x. Ifm (x#bs) (plusinf p)" by auto - from rplusinf_inf[OF lp, where bs="bs"] + from rplusinf_inf[OF lp, where bs="bs"] obtain z where z_def: "\x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p" by blast from th have "Ifm ((z + 1)#bs) (plusinf p)" by simp moreover have "z + 1 > z" by simp ultimately show ?thesis using z_def by auto qed -consts +consts \:: "fm \ (num \ int) list" \ :: "fm \ (num \ int) \ fm " recdef \ "measure size" - "\ (And p q) = (\ p @ \ q)" - "\ (Or p q) = (\ p @ \ q)" + "\ (And p q) = (\ p @ \ q)" + "\ (Or p q) = (\ p @ \ q)" "\ (Eq (CN 0 c e)) = [(Neg e,c)]" "\ (NEq (CN 0 c e)) = [(Neg e,c)]" "\ (Lt (CN 0 c e)) = [(Neg e,c)]" @@ -4410,10 +4410,10 @@ have "?I ?u (Lt (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) < 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) < 0)" - by (simp only: pos_less_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" + by (simp only: pos_less_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) < 0)" - using np by simp + using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (6 c e) @@ -4421,10 +4421,10 @@ have "?I ?u (Le (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_le_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" + by (simp only: pos_le_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" - using np by simp + using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (7 c e) @@ -4432,10 +4432,10 @@ have "?I ?u (Gt (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) > 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) > 0)" - by (simp only: pos_divide_less_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" + by (simp only: pos_divide_less_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) > 0)" - using np by simp + using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (8 c e) @@ -4443,10 +4443,10 @@ have "?I ?u (Ge (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: pos_divide_le_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" + by (simp only: pos_divide_le_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" - using np by simp + using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (3 c e) @@ -4455,10 +4455,10 @@ have "?I ?u (Eq (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) = 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) = 0)" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) = 0)" - using np by simp + using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) next case (4 c e) @@ -4467,10 +4467,10 @@ have "?I ?u (NEq (CN 0 c e)) = (real_of_int c *(?t/?n) + (?N x e) \ 0)" using numbound0_I[OF nb, where bs="bs" and b="?u" and b'="x"] by simp also have "\ = (?n*(real_of_int c *(?t/?n)) + ?n*(?N x e) \ 0)" - by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" + by (simp only: nonzero_eq_divide_eq[OF np, where a="real_of_int c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: algebra_simps) also have "\ = (real_of_int c *?t + ?n* (?N x e) \ 0)" - using np by simp + using np by simp finally show ?case using nbt nb by (simp add: algebra_simps) qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real_of_int n" and b'="x"]) @@ -4491,7 +4491,7 @@ by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"]) then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real_of_int m * x \ ?N a s" by blast from \_l[OF lp] smU have mp: "real_of_int m > 0" by auto - from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" + from pos_divide_le_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" by (auto simp add: mult.commute) thus ?thesis using smU by auto qed @@ -4507,14 +4507,14 @@ by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"]) then obtain s m where smU: "(s,m) \ set (\ p)" and mx: "real_of_int m * x \ ?N a s" by blast from \_l[OF lp] smU have mp: "real_of_int m > 0" by auto - from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" + from pos_le_divide_eq[OF mp, where a="x" and b="?N a s", symmetric] mx have "x \ ?N a s / real_of_int m" by (auto simp add: mult.commute) thus ?thesis using smU by auto qed -lemma lin_dense: +lemma lin_dense: assumes lp: "isrlfm p" - and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real_of_int n) ` set (\ p)" + and noS: "\ t. l < t \ t< u \ t \ (\ (t,n). Inum (x#bs) t / real_of_int n) ` set (\ p)" (is "\ t. _ \ _ \ t \ (\ (t,n). ?N x t / real_of_int n ) ` (?U p)") and lx: "l < x" and xu:"x < u" and px:" Ifm (x#bs) p" and ly: "l < y" and yu: "y < u" @@ -4523,7 +4523,7 @@ proof (induct p rule: isrlfm.induct) case (5 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from 5 have "x * real_of_int c + ?N x e < 0" by (simp add: algebra_simps) - hence pxc: "x < (- ?N x e) / real_of_int c" + hence pxc: "x < (- ?N x e) / real_of_int c" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) from 5 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto @@ -4533,7 +4533,7 @@ by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y > (- ?N x e) / real_of_int c" + moreover {assume y: "y > (- ?N x e) / real_of_int c" with yu have eu: "u > (- ?N x e) / real_of_int c" by auto with noSc ly yu have "(- ?N x e) / real_of_int c \ l" by (cases "(- ?N x e) / real_of_int c > l", auto) with lx pxc have "False" by auto @@ -4542,7 +4542,7 @@ next case (6 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from 6 have "x * real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) - hence pxc: "x \ (- ?N x e) / real_of_int c" + hence pxc: "x \ (- ?N x e) / real_of_int c" by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) from 6 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto @@ -4552,7 +4552,7 @@ by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real_of_int c * y + ?N x e < 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y > (- ?N x e) / real_of_int c" + moreover {assume y: "y > (- ?N x e) / real_of_int c" with yu have eu: "u > (- ?N x e) / real_of_int c" by auto with noSc ly yu have "(- ?N x e) / real_of_int c \ l" by (cases "(- ?N x e) / real_of_int c > l", auto) with lx pxc have "False" by auto @@ -4561,7 +4561,7 @@ next case (7 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from 7 have "x * real_of_int c + ?N x e > 0" by (simp add: algebra_simps) - hence pxc: "x > (- ?N x e) / real_of_int c" + hence pxc: "x > (- ?N x e) / real_of_int c" by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) from 7 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto @@ -4571,7 +4571,7 @@ by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y < (- ?N x e) / real_of_int c" + moreover {assume y: "y < (- ?N x e) / real_of_int c" with ly have eu: "l < (- ?N x e) / real_of_int c" by auto with noSc ly yu have "(- ?N x e) / real_of_int c \ u" by (cases "(- ?N x e) / real_of_int c > l", auto) with xu pxc have "False" by auto @@ -4580,7 +4580,7 @@ next case (8 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from 8 have "x * real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) - hence pxc: "x \ (- ?N x e) / real_of_int c" + hence pxc: "x \ (- ?N x e) / real_of_int c" by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) from 8 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto @@ -4590,7 +4590,7 @@ by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real_of_int c * y + ?N x e > 0" by (simp add: algebra_simps) hence ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by simp} - moreover {assume y: "y < (- ?N x e) / real_of_int c" + moreover {assume y: "y < (- ?N x e) / real_of_int c" with ly have eu: "l < (- ?N x e) / real_of_int c" by auto with noSc ly yu have "(- ?N x e) / real_of_int c \ u" by (cases "(- ?N x e) / real_of_int c > l", auto) with xu pxc have "False" by auto @@ -4600,7 +4600,7 @@ case (3 c e) hence cp: "real_of_int c > 0" and nb: "numbound0 e" by simp_all from cp have cnz: "real_of_int c \ 0" by simp from 3 have "x * real_of_int c + ?N x e = 0" by (simp add: algebra_simps) - hence pxc: "x = (- ?N x e) / real_of_int c" + hence pxc: "x = (- ?N x e) / real_of_int c" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) from 3 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with lx xu have yne: "x \ - ?N x e / real_of_int c" by auto @@ -4610,10 +4610,10 @@ from cp have cnz: "real_of_int c \ 0" by simp from 4 have noSc:"\ t. l < t \ t < u \ t \ (- ?N x e) / real_of_int c" by auto with ly yu have yne: "y \ - ?N x e / real_of_int c" by auto - hence "y* real_of_int c \ -?N x e" + hence "y* real_of_int c \ -?N x e" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="y" and b="-?N x e"]) simp hence "y* real_of_int c + ?N x e \ 0" by (simp add: algebra_simps) - thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] + thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by (simp add: algebra_simps) qed (auto simp add: numbound0_I[where bs="bs" and b="y" and b'="x"]) @@ -4623,7 +4623,7 @@ and npi: "\ (Ifm (x#bs) (plusinf p))" (is "\ (Ifm (x#bs) (?P p))") and ex: "\ x. Ifm (x#bs) p" (is "\ x. ?I x p") shows "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). - ?I ((Inum (x#bs) l / real_of_int n + Inum (x#bs) s / real_of_int m) / 2) p" + ?I ((Inum (x#bs) l / real_of_int n + Inum (x#bs) s / real_of_int m) / 2) p" proof- let ?N = "\ x t. Inum (x#bs) t" let ?U = "set (\ p)" @@ -4636,10 +4636,10 @@ proof- let ?M = "(\ (t,c). ?N a t / real_of_int c) ` ?U" have fM: "finite ?M" by auto - from rminusinf_\[OF lp nmi pa] rplusinf_\[OF lp npi pa] + from rminusinf_\[OF lp nmi pa] rplusinf_\[OF lp npi pa] have "\ (l,n) \ set (\ p). \ (s,m) \ set (\ p). a \ ?N x l / real_of_int n \ a \ ?N x s / real_of_int m" by blast - then obtain "t" "n" "s" "m" where - tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" + then obtain "t" "n" "s" "m" where + tnU: "(t,n) \ ?U" and smU: "(s,m) \ ?U" and xs1: "a \ ?N x s / real_of_int m" and tx1: "a \ ?N x t / real_of_int n" by blast from \_l[OF lp] tnU smU numbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 have xs: "a \ ?N a s / real_of_int m" and tx: "a \ ?N a t / real_of_int n" by auto from tnU have Mne: "?M \ {}" by auto @@ -4649,23 +4649,23 @@ have linM: "?l \ ?M" using fM Mne by simp have uinM: "?u \ ?M" using fM Mne by simp have tnM: "?N a t / real_of_int n \ ?M" using tnU by auto - have smM: "?N a s / real_of_int m \ ?M" using smU by auto + have smM: "?N a s / real_of_int m \ ?M" using smU by auto have lM: "\ t\ ?M. ?l \ t" using Mne fM by auto have Mu: "\ t\ ?M. t \ ?u" using Mne fM by auto have "?l \ ?N a t / real_of_int n" using tnM Mne by simp hence lx: "?l \ a" using tx by simp have "?N a s / real_of_int m \ ?u" using smM Mne by simp hence xu: "a \ ?u" using xs by simp from finite_set_intervals2[where P="\ x. ?I x p",OF pa lx xu linM uinM fM lM Mu] - have "(\ s\ ?M. ?I s p) \ + have "(\ s\ ?M. ?I s p) \ (\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p)" . moreover { fix u assume um: "u\ ?M" and pu: "?I u p" hence "\ (tu,nu) \ ?U. u = ?N a tu / real_of_int nu" by auto then obtain "tu" "nu" where tuU: "(tu,nu) \ ?U" and tuu:"u= ?N a tu / real_of_int nu" by blast - have "(u + u) / 2 = u" by auto with pu tuu + have "(u + u) / 2 = u" by auto with pu tuu have "?I (((?N a tu / real_of_int nu) + (?N a tu / real_of_int nu)) / 2) p" by simp with tuU have ?thesis by blast} moreover{ assume "\ t1\ ?M. \ t2 \ ?M. (\ y. t1 < y \ y < t2 \ y \ ?M) \ t1 < a \ a < t2 \ ?I a p" - then obtain t1 and t2 where t1M: "t1 \ ?M" and t2M: "t2\ ?M" + then obtain t1 and t2 where t1M: "t1 \ ?M" and t2M: "t2\ ?M" and noM: "\ y. t1 < y \ y < t2 \ y \ ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p" by blast from t1M have "\ (t1u,t1n) \ ?U. t1 = ?N a t1u / real_of_int t1n" by auto @@ -4679,10 +4679,10 @@ with t1uU t2uU t1u t2u have ?thesis by blast} ultimately show ?thesis by blast qed - then obtain "l" "n" "s" "m" where lnU: "(l,n) \ ?U" and smU:"(s,m) \ ?U" + then obtain "l" "n" "s" "m" where lnU: "(l,n) \ ?U" and smU:"(s,m) \ ?U" and pu: "?I ((?N a l / real_of_int n + ?N a s / real_of_int m) / 2) p" by blast from lnU smU \_l[OF lp] have nbl: "numbound0 l" and nbs: "numbound0 s" by auto - from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] + from numbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu have "?I ((?N x l / real_of_int n + ?N x s / real_of_int m) / 2) p" by simp with lnU smU @@ -4690,7 +4690,7 @@ qed (* The Ferrante - Rackoff Theorem *) -theorem fr_eq: +theorem fr_eq: assumes lp: "isrlfm p" shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,n) \ set (\ p). \ (s,m) \ set (\ p). Ifm ((((Inum (x#bs) t)/ real_of_int n + (Inum (x#bs) s) / real_of_int m) /2)#bs) p))" (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") @@ -4702,7 +4702,7 @@ from rinf_\[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast} ultimately show "?D" by blast next - assume "?D" + assume "?D" moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } moreover {assume f:"?F" hence "?E" by blast} @@ -4710,7 +4710,7 @@ qed -lemma fr_eq_\: +lemma fr_eq_\: assumes lp: "isrlfm p" shows "(\ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) \ (Ifm (x#bs) (plusinf p)) \ (\ (t,k) \ set (\ p). \ (s,l) \ set (\ p). Ifm (x#bs) (\ p (Add(Mul l t) (Mul k s) , 2*k*l))))" (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") @@ -4729,15 +4729,15 @@ from tnb snb have st_nb: "numbound0 ?st" by simp have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnp mp np by (simp add: algebra_simps add_divide_distrib) - from \_I[OF lp mnp st_nb, where x="x" and bs="bs"] + from \_I[OF lp mnp st_nb, where x="x" and bs="bs"] have "?I x (\ p (?st,2*n*m)) = ?I ((?N t / real_of_int n + ?N s / real_of_int m) /2) p" by (simp only: st[symmetric])} with rinf_\[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} ultimately show "?D" by blast next - assume "?D" + assume "?D" moreover {assume m:"?M" from rminusinf_ex[OF lp m] have "?E" .} moreover {assume p: "?P" from rplusinf_ex[OF lp p] have "?E" . } - moreover {fix t k s l assume "(t,k) \ set (\ p)" and "(s,l) \ set (\ p)" + moreover {fix t k s l assume "(t,k) \ set (\ p)" and "(s,l) \ set (\ p)" and px:"?I x (\ p (Add (Mul l t) (Mul k s), 2*k*l))" with \_l[OF lp] have tnb: "numbound0 t" and np:"real_of_int k > 0" and snb: "numbound0 s" and mp:"real_of_int l > 0" by auto let ?st = "Add (Mul l t) (Mul k s)" @@ -4759,7 +4759,7 @@ have "x = real_of_int ?i + ?u" by simp hence "P (real_of_int ?i + ?u)" using Px by simp moreover have "real_of_int ?i \ x" using of_int_floor_le by simp hence "0 \ ?u" by arith - moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith + moreover have "?u < 1" using real_of_int_floor_add_one_gt[where r="x"] by arith ultimately show "(\ (i::int) (u::real). 0\ u \ u< 1 \ P (real_of_int i + u))" by blast qed @@ -4792,11 +4792,11 @@ | "exsplit (NOT p) = NOT (exsplit p)" | "exsplit p = p" -lemma exsplitnum: +lemma exsplitnum: "Inum (x#y#bs) (exsplitnum t) = Inum ((x+y) #bs) t" by(induct t rule: exsplitnum.induct) (simp_all add: algebra_simps) -lemma exsplit: +lemma exsplit: assumes qfp: "qfree p" shows "Ifm (x#y#bs) (exsplit p) = Ifm ((x+y)#bs) p" using qfp exsplitnum[where x="x" and y="y" and bs="bs"] @@ -4810,7 +4810,7 @@ by (simp add: myless[of _ "1"] myless[of _ "0"] ac_simps) also have "\ = (\ (i::int). \ x. 0\ x \ x < 1 \ Ifm ((real_of_int i + x) #bs) p)" by (simp only: exsplit[OF qf] ac_simps) - also have "\ = (\ x. Ifm (x#bs) p)" + also have "\ = (\ x. Ifm (x#bs) p)" by (simp only: real_ex_int_real01[where P="\ x. Ifm (x#bs) p"]) finally show ?thesis by simp qed @@ -4819,12 +4819,12 @@ definition ferrack01 :: "fm \ fm" where "ferrack01 p \ (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p); - U = remdups(map simp_num_pair + U = remdups(map simp_num_pair (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) - (alluopairs (\ p')))) + (alluopairs (\ p')))) in decr (evaldjf (\ p') U ))" -lemma fr_eq_01: +lemma fr_eq_01: assumes qf: "qfree p" shows "(\ x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (\ (t,n) \ set (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). \ (s,m) \ set (\ (rlfm (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p))). Ifm (x#bs) (\ (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)))" (is "(\ x. ?I x ?q) = ?F") @@ -4837,18 +4837,18 @@ by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))", simp_all) have PF: "?P = False" apply (simp add: Let_def reducecoeff_def numgcd_def rsplit_def ge_def lt_def conj_def disj_def) by (cases "rlfm p = And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))", simp_all) - have "(\ x. ?I x ?q ) = + have "(\ x. ?I x ?q ) = ((?I x (minusinf ?rq)) \ (?I x (plusinf ?rq )) \ (\ (t,n) \ set (\ ?rq). \ (s,m) \ set (\ ?rq ). ?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))))" (is "(\ x. ?I x ?q) = (?M \ ?P \ ?F)" is "?E = ?D") proof - assume "\ x. ?I x ?q" + assume "\ x. ?I x ?q" then obtain x where qx: "?I x ?q" by blast - hence xp: "0\ x" and x1: "x< 1" and px: "?I x p" + hence xp: "0\ x" and x1: "x< 1" and px: "?I x p" by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf]) - from qx have "?I x ?rq " + from qx have "?I x ?rq " by (simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) hence lqx: "?I x ?rq " using simpfm[where p="?rq" and bs="x#bs"] by auto - from qf have qfq:"isrlfm ?rq" + from qf have qfq:"isrlfm ?rq" by (auto simp add: rsplit_def lt_def ge_def rlfm_I[OF qf xp x1]) with lqx fr_eq_\[OF qfq] show "?M \ ?P \ ?F" by blast next @@ -4856,7 +4856,7 @@ let ?U = "set (\ ?rq )" from MF PF D have "?F" by auto then obtain t n s m where aU:"(t,n) \ ?U" and bU:"(s,m)\ ?U" and rqx: "?I x (\ ?rq (Add (Mul m t) (Mul n s), 2*n*m))" by blast - from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] + from qf have lrq:"isrlfm ?rq"using rlfm_l[OF qf] by (auto simp add: rsplit_def lt_def ge_def) from aU bU \_l[OF lrq] have tnb: "numbound0 t" and np:"real_of_int n > 0" and snb: "numbound0 s" and mp:"real_of_int m > 0" by (auto simp add: split_def) let ?st = "Add (Mul m t) (Mul n s)" @@ -4864,7 +4864,7 @@ from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute) from conjunct1[OF \_I[OF lrq mnp stnb, where bs="bs" and x="x"], symmetric] rqx have "\ x. ?I x ?rq" by auto - thus "?E" + thus "?E" using rlfm_I[OF qf] by (auto simp add: rsplit_def lt_def ge_def) qed with MF PF show ?thesis by blast @@ -4882,25 +4882,25 @@ let ?N = "\ t. Inum (x#bs) t" let ?st= "Add (Mul m t) (Mul n s)" from Ul th have mnz: "m \ 0" by auto - from Ul th have nnz: "n \ 0" by auto + from Ul th have nnz: "n \ 0" by auto have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnz nnz by (simp add: algebra_simps add_divide_distrib) - + thus "(real_of_int m * Inum (x # bs) t + real_of_int n * Inum (x # bs) s) / (2 * real_of_int n * real_of_int m) \ (\((t, n), s, m). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2) ` - (set U \ set U)"using mnz nnz th + (set U \ set U)"using mnz nnz th apply (auto simp add: th add_divide_distrib algebra_simps split_def image_def) - by (rule_tac x="(s,m)" in bexI,simp_all) + by (rule_tac x="(s,m)" in bexI,simp_all) (rule_tac x="(t,n)" in bexI,simp_all add: mult.commute) next fix t n s m - assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" + assume tnU: "(t,n) \ set U" and smU:"(s,m) \ set U" let ?N = "\ t. Inum (x#bs) t" let ?st= "Add (Mul m t) (Mul n s)" from Ul smU have mnz: "m \ 0" by auto - from Ul tnU have nnz: "n \ 0" by auto + from Ul tnU have nnz: "n \ 0" by auto have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" using mnz nnz by (simp add: algebra_simps add_divide_distrib) let ?P = "\ (t',n') (s',m'). (Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m')/2" @@ -4910,13 +4910,13 @@ from alluopairs_ex[OF Pc, where xs="U"] tnU smU have th':"\ ((t',n'),(s',m')) \ set (alluopairs U). ?P (t',n') (s',m')" by blast - then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \ set (alluopairs U)" + then obtain t' n' s' m' where ts'_U: "((t',n'),(s',m')) \ set (alluopairs U)" and Pts': "?P (t',n') (s',m')" by blast from ts'_U Up have mnz': "m' \ 0" and nnz': "n'\ 0" by auto let ?st' = "Add (Mul m' t') (Mul n' s')" have st': "(?N t' / real_of_int n' + ?N s' / real_of_int m')/2 = ?N ?st' / real_of_int (2*n'*m')" using mnz' nnz' by (simp add: algebra_simps add_divide_distrib) - from Pts' have + from Pts' have "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m)/2 = (Inum (x # bs) t' / real_of_int n' + Inum (x # bs) s' / real_of_int m')/2" by simp also have "\ = ((\(t, n). Inum (x # bs) t / real_of_int n) ((\((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ((t',n'),(s',m'))))" by (simp add: st') finally show "(Inum (x # bs) t / real_of_int n + Inum (x # bs) s / real_of_int m) / 2 @@ -4935,13 +4935,13 @@ (is "?lhs = ?rhs") proof assume ?lhs - then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and + then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and Pst: "Ifm (x#bs) (\ p (Add (Mul m t) (Mul n s),2*n*m))" by blast let ?N = "\ t. Inum (x#bs) t" - from tnU smU U have tnb: "numbound0 t" and np: "n > 0" + from tnU smU U have tnb: "numbound0 t" and np: "n > 0" and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real_of_int (2*n*m) > 0" + from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" @@ -4951,25 +4951,25 @@ by auto (rule_tac x="(a,b)" in bexI, auto) then obtain t' n' where tnU': "(t',n') \ U'" and th: "?g ((t,n),(s,m)) = ?f (t',n')" by blast from U' tnU' have tnb': "numbound0 t'" and np': "real_of_int n' > 0" by auto - from \_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst + from \_I[OF lp mnp stnb, where bs="bs" and x="x"] Pst have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p" by simp from conjunct1[OF \_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]] - have "Ifm (x # bs) (\ p (t', n')) " by (simp only: st) - then show ?rhs using tnU' by auto + have "Ifm (x # bs) (\ p (t', n')) " by (simp only: st) + then show ?rhs using tnU' by auto next assume ?rhs - then obtain t' n' where tnU': "(t',n') \ U'" and Pt': "Ifm (x # bs) (\ p (t', n'))" + then obtain t' n' where tnU': "(t',n') \ U'" and Pt': "Ifm (x # bs) (\ p (t', n'))" by blast from tnU' UU' have "?f (t',n') \ ?g ` (U\U)" by blast - hence "\ ((t,n),(s,m)) \ (U\U). ?f (t',n') = ?g ((t,n),(s,m))" + hence "\ ((t,n),(s,m)) \ (U\U). ?f (t',n') = ?g ((t,n),(s,m))" by auto (rule_tac x="(a,b)" in bexI, auto) - then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and + then obtain t n s m where tnU: "(t,n) \ U" and smU:"(s,m) \ U" and th: "?f (t',n') = ?g((t,n),(s,m)) "by blast let ?N = "\ t. Inum (x#bs) t" - from tnU smU U have tnb: "numbound0 t" and np: "n > 0" + from tnU smU U have tnb: "numbound0 t" and np: "n > 0" and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" - from np mp have mnp: "real_of_int (2*n*m) > 0" + from np mp have mnp: "real_of_int (2*n*m) > 0" by (simp add: mult.commute of_int_mult[symmetric] del: of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real_of_int n + ?N s / real_of_int m)/2 = ?N ?st / real_of_int (2*n*m)" @@ -4979,8 +4979,8 @@ have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real_of_int (2 * n * m) # bs) p" by simp with \_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast qed - -lemma ferrack01: + +lemma ferrack01: assumes qf: "qfree p" shows "((\ x. Ifm (x#bs) (And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p)) = (Ifm bs (ferrack01 p))) \ qfree (ferrack01 p)" (is "(?lhs = ?rhs) \ _") proof- @@ -4998,17 +4998,17 @@ let ?h = "\ ((t,n),(s,m)). (?N t/real_of_int n + ?N s/ real_of_int m) /2" let ?F = "\ p. \ a \ set (\ p). \ b \ set (\ p). ?I x (\ p (?g(a,b)))" let ?ep = "evaldjf (\ ?q) ?Y" - from rlfm_l[OF qf] have lq: "isrlfm ?q" + from rlfm_l[OF qf] have lq: "isrlfm ?q" by (simp add: rsplit_def lt_def ge_def conj_def disj_def Let_def reducecoeff_def numgcd_def) from alluopairs_set1[where xs="?U"] have UpU: "set ?Up \ (set ?U \ set ?U)" by simp from \_l[OF lq] have U_l: "\ (t,n) \ set ?U. numbound0 t \ n > 0" . - from U_l UpU + from U_l UpU have "\ ((t,n),(s,m)) \ set ?Up. numbound0 t \ n> 0 \ numbound0 s \ m > 0" by auto hence Snb: "\ (t,n) \ set ?S. numbound0 t \ n > 0 " by (auto) - have Y_l: "\ (t,n) \ set ?Y. numbound0 t \ n > 0" + have Y_l: "\ (t,n) \ set ?Y. numbound0 t \ n > 0" proof- - { fix t n assume tnY: "(t,n) \ set ?Y" + { fix t n assume tnY: "(t,n) \ set ?Y" hence "(t,n) \ set ?SS" by simp hence "\ (t',n') \ set ?S. simp_num_pair (t',n') = (t,n)" by (auto simp add: split_def simp del: map_map) @@ -5022,12 +5022,12 @@ have YU: "(?f ` set ?Y) = (?h ` (set ?U \ set ?U))" proof- - from simp_num_pair_ci[where bs="x#bs"] have + from simp_num_pair_ci[where bs="x#bs"] have "\x. (?f o simp_num_pair) x = ?f x" by auto hence th: "?f o simp_num_pair = ?f" using ext by blast have "(?f ` set ?Y) = ((?f o simp_num_pair) ` set ?S)" by (simp add: image_comp comp_assoc) also have "\ = (?f ` set ?S)" by (simp add: th) - also have "\ = ((?f o ?g) ` set ?Up)" + also have "\ = ((?f o ?g) ` set ?Up)" by (simp only: set_map o_def image_comp) also have "\ = (?h ` (set ?U \ set ?U))" using \_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_comp] by blast @@ -5047,8 +5047,8 @@ from fr_eq_01[OF qf, where bs="bs" and x="x"] have "?lhs = ?F ?q" by (simp only: split_def fst_conv snd_conv) also have "\ = (\ (t,n) \ set ?Y. ?I x (\ ?q (t,n)))" using \_cong[OF lq YU U_l Y_l] - by (simp only: split_def fst_conv snd_conv) - also have "\ = (Ifm (x#bs) ?ep)" + by (simp only: split_def fst_conv snd_conv) + also have "\ = (Ifm (x#bs) ?ep)" using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="\ ?q",symmetric] by (simp only: split_def prod.collapse) also have "\ = (Ifm bs (decr ?ep))" using decr[OF ep_nb] by blast @@ -5057,7 +5057,7 @@ with lr show ?thesis by blast qed -lemma cp_thm': +lemma cp_thm': assumes lp: "iszlfm p (real_of_int (i::int)#bs)" and up: "d_\ p 1" and dd: "d_\ p d" and dp: "d > 0" shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. d}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ j\ {1.. d}. \ b\ (Inum (real_of_int i#bs)) ` set (\ p). Ifm ((b+real_of_int j)#bs) p))" @@ -5070,12 +5070,12 @@ lemma unit: assumes qf: "qfree p" shows "\ q B d. unit p = (q,B,d) \ - ((\ (x::int). Ifm (real_of_int x#bs) p) = - (\ (x::int). Ifm (real_of_int x#bs) q)) \ - (Inum (real_of_int i#bs)) ` set B = (Inum (real_of_int i#bs)) ` set (\ q) \ + ((\ (x::int). Ifm (real_of_int x#bs) p) = + (\ (x::int). Ifm (real_of_int x#bs) q)) \ + (Inum (real_of_int i#bs)) ` set B = (Inum (real_of_int i#bs)) ` set (\ q) \ d_\ q 1 \ d_\ q d \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ b\ set B. numbound0 b)" proof- - fix q B d + fix q B d assume qBd: "unit p = (q,B,d)" let ?thes = "((\ (x::int). Ifm (real_of_int x#bs) p) = (\ (x::int). Ifm (real_of_int x#bs) q)) \ Inum (real_of_int i#bs) ` set B = Inum (real_of_int i#bs) ` set (\ q) \ @@ -5089,22 +5089,22 @@ let ?B'= "remdups (map simpnum (\ ?q))" let ?A = "set (\ ?q)" let ?A'= "remdups (map simpnum (\ ?q))" - from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] + from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] have pp': "\ i. ?I i ?p' = ?I i p" by auto from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]]] - have lp': "\ (i::int). iszlfm ?p' (real_of_int i#bs)" by simp + have lp': "\ (i::int). iszlfm ?p' (real_of_int i#bs)" by simp hence lp'': "iszlfm ?p' (real_of_int (i::int)#bs)" by simp from lp' \[where p="?p'" and bs="bs"] have lp: "?l >0" and dl: "d_\ ?p' ?l" by auto from a_\_ex[where p="?p'" and l="?l" and bs="bs", OF lp'' dl lp] pp' - have pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by (simp add: int_rdvd_iff) - from lp'' lp a_\[OF lp'' dl lp] have lq:"iszlfm ?q (real_of_int i#bs)" and uq: "d_\ ?q 1" + have pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by (simp add: int_rdvd_iff) + from lp'' lp a_\[OF lp'' dl lp] have lq:"iszlfm ?q (real_of_int i#bs)" and uq: "d_\ ?q 1" by (auto simp add: isint_def) from \[OF lq] have dp:"?d >0" and dd: "d_\ ?q ?d" by blast+ let ?N = "\ t. Inum (real_of_int (i::int)#bs) t" - have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_comp) + have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by (simp add:image_comp) also have "\ = ?N ` ?B" using simpnum_ci[where bs="real_of_int i #bs"] by auto finally have BB': "?N ` set ?B' = ?N ` ?B" . - have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_comp) + have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by (simp add:image_comp) also have "\ = ?N ` ?A" using simpnum_ci[where bs="real_of_int i #bs"] by auto finally have AA': "?N ` set ?A' = ?N ` ?A" . from \_numbound0[OF lq] have B_nb:"\ b\ set ?B'. numbound0 b" @@ -5114,16 +5114,16 @@ { assume "length ?B' \ length ?A'" hence q:"q=?q" and "B = ?B'" and d:"d = ?d" using qBd by (auto simp add: Let_def unit_def) - with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" + with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" and bn: "\b\ set B. numbound0 b" by simp+ with pq_ex dp uq dd lq q d have ?thes by simp } - moreover + moreover { assume "\ (length ?B' \ length ?A')" hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" using qBd by (auto simp add: Let_def unit_def) - with AA' mirror_\_\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" + with AA' mirror_\_\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" and bn: "\b\ set B. numbound0 b" by simp+ - from mirror_ex[OF lq] pq_ex q + from mirror_ex[OF lq] pq_ex q have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp from lq uq q mirror_d_\ [where p="?q" and bs="bs" and a="real_of_int i"] have lq': "iszlfm q (real_of_int i#bs)" and uq: "d_\ q 1" by auto @@ -5135,17 +5135,17 @@ (* Cooper's Algorithm *) definition cooper :: "fm \ fm" where - "cooper p \ + "cooper p \ (let (q,B,d) = unit p; js = [1..d]; mq = simpfm (minusinf q); md = evaldjf (\ j. simpfm (subst0 (C j) mq)) js in if md = T then T else - (let qd = evaldjf (\ t. simpfm (subst0 t q)) - (remdups (map (\ (b,j). simpnum (Add b (C j))) + (let qd = evaldjf (\ t. simpfm (subst0 t q)) + (remdups (map (\ (b,j). simpnum (Add b (C j))) [(b,j). b\B,j\js])) in decr (disj md qd)))" lemma cooper: assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (cooper p))) \ qfree (cooper p)" + shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (cooper p))) \ qfree (cooper p)" (is "(?lhs = ?rhs) \ _") proof- @@ -5163,65 +5163,65 @@ let ?sbjs = "map (\ (b,j). simpnum (Add b (C j))) ?bjs" let ?qd = "evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs)" have qbf:"unit p = (?q,?B,?d)" by simp - from unit[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and - B:"?N ` set ?B = ?N ` set (\ ?q)" and - uq:"d_\ ?q 1" and dd: "d_\ ?q ?d" and dp: "?d > 0" and - lq: "iszlfm ?q (real_of_int i#bs)" and + from unit[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and + B:"?N ` set ?B = ?N ` set (\ ?q)" and + uq:"d_\ ?q 1" and dd: "d_\ ?q ?d" and dp: "?d > 0" and + lq: "iszlfm ?q (real_of_int i#bs)" and Bn: "\ b\ set ?B. numbound0 b" by auto from zlin_qfree[OF lq] have qfq: "qfree ?q" . from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". have jsnb: "\ j \ set ?js. numbound0 (C j)" by simp - hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" + hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" by (auto simp only: subst0_bound0[OF qfmq]) hence th: "\ j\ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" by auto - from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp + from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp from Bn jsnb have "\ (b,j) \ set ?bjs. numbound0 (Add b (C j))" by simp hence "\ (b,j) \ set ?bjs. numbound0 (simpnum (Add b (C j)))" using simpnum_numbound0 by blast hence "\ t \ set ?sbjs. numbound0 t" by simp hence "\ t \ set (remdups ?sbjs). bound0 (subst0 t ?q)" - using subst0_bound0[OF qfq] by auto + using subst0_bound0[OF qfq] by auto hence th': "\ t \ set (remdups ?sbjs). bound0 (simpfm (subst0 t ?q))" using simpfm_bound0 by blast from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp - from mdb qdb + from mdb qdb have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) from trans [OF pq_ex cp_thm'[OF lq uq dd dp]] B have "?lhs = (\ j\ {1.. ?d}. ?I j ?mq \ (\ b\ ?N ` set ?B. Ifm ((b+ real_of_int j)#bs) ?q))" by auto also have "\ = ((\ j\ set ?js. ?I j ?smq) \ (\ (b,j) \ (?N ` set ?B \ set ?js). Ifm ((b+ real_of_int j)#bs) ?q))" by auto also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (Add b (C j))) ` set ?bjs. Ifm (t #bs) ?q))" by simp also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ (\ (b,j). ?N (simpnum (Add b (C j)))) ` set ?bjs. Ifm (t #bs) ?q))" by (simp only: simpnum_ci) - also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ set ?sbjs. Ifm (?N t #bs) ?q))" - by (auto simp add: split_def) + also have "\= ((\ j\ set ?js. ?I j ?smq) \ (\ t \ set ?sbjs. Ifm (?N t #bs) ?q))" + by (auto simp add: split_def) also have "\ = ((\ j\ set ?js. (\ j. ?I i (simpfm (subst0 (C j) ?smq))) j) \ (\ t \ set (remdups ?sbjs). (\ t. ?I i (simpfm (subst0 t ?q))) t))" by (simp only: simpfm subst0_I[OF qfq] Inum.simps subst0_I[OF qfmq] set_remdups) also have "\ = ((?I i (evaldjf (\ j. simpfm (subst0 (C j) ?smq)) ?js)) \ (?I i (evaldjf (\ t. simpfm (subst0 t ?q)) (remdups ?sbjs))))" by (simp only: evaldjf_ex) finally have mdqd: "?lhs = (?I i (disj ?md ?qd))" by simp hence mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" using decr [OF mdqdb] by simp {assume mdT: "?md = T" - hence cT:"cooper p = T" + hence cT:"cooper p = T" by (simp only: cooper_def unit_def split_def Let_def if_True) simp from mdT mdqd have lhs:"?lhs" by auto from mdT have "?rhs" by (simp add: cooper_def unit_def split_def) with lhs cT have ?thesis by simp } moreover - {assume mdT: "?md \ T" hence "cooper p = decr (disj ?md ?qd)" - by (simp only: cooper_def unit_def split_def Let_def if_False) + {assume mdT: "?md \ T" hence "cooper p = decr (disj ?md ?qd)" + by (simp only: cooper_def unit_def split_def Let_def if_False) with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } ultimately show ?thesis by blast qed -lemma DJcooper: +lemma DJcooper: assumes qf: "qfree p" shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (DJ cooper p))) \ qfree (DJ cooper p)" proof- from cooper have cqf: "\ p. qfree p \ qfree (cooper p)" by blast from DJ_qf[OF cqf] qf have thqf:"qfree (DJ cooper p)" by blast - have "Ifm bs (DJ cooper p) = (\ q\ set (disjuncts p). Ifm bs (cooper q))" + have "Ifm bs (DJ cooper p) = (\ q\ set (disjuncts p). Ifm bs (cooper q))" by (simp add: DJ_def evaldjf_ex) - also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real_of_int x#bs) q)" + also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real_of_int x#bs) q)" using cooper disjuncts_qf[OF qf] by blast also have "\ = (\ (x::int). Ifm (real_of_int x#bs) p)" by (induct p rule: disjuncts.induct, auto) finally show ?thesis using thqf by blast @@ -5231,20 +5231,20 @@ lemma \_\_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" shows "Ifm (a#bs) (\_\ p (t,c)) = Ifm (a#bs) (\_\ p (t',c))" - using lp + using lp by (induct p rule: iszlfm.induct, auto simp add: tt') lemma \_cong: assumes lp: "iszlfm p (a#bs)" and tt': "Inum (a#bs) t = Inum (a#bs) t'" shows "Ifm (a#bs) (\ p c t) = Ifm (a#bs) (\ p c t')" by (simp add: \_def tt' \_\_cong[OF lp tt']) -lemma \_cong: assumes lp: "iszlfm p (a#bs)" +lemma \_cong: assumes lp: "iszlfm p (a#bs)" and RR: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" shows "(\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))) = (\ (e,c) \ set (\ p). \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j))))" (is "?lhs = ?rhs") proof let ?d = "\ p" - assume ?lhs then obtain e c j where ecR: "(e,c) \ R" and jD:"j \ {1 .. c*?d}" + assume ?lhs then obtain e c j where ecR: "(e,c) \ R" and jD:"j \ {1 .. c*?d}" and px: "Ifm (a#bs) (\ p c (Add e (C j)))" (is "?sp c e j") by blast from ecR have "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` R" by auto hence "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" using RR by simp @@ -5252,14 +5252,14 @@ then obtain e' c' where ecRo:"(e',c') \ set (\ p)" and ee':"Inum (a#bs) e = Inum (a#bs) e'" and cc':"c = c'" by blast from ee' have tt': "Inum (a#bs) (Add e (C j)) = Inum (a#bs) (Add e' (C j))" by simp - + from \_cong[OF lp tt', where c="c"] px have px':"?sp c e' j" by simp from ecRo jD px' show ?rhs apply (auto simp: cc') by (rule_tac x="(e', c')" in bexI,simp_all) (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) next let ?d = "\ p" - assume ?rhs then obtain e c j where ecR: "(e,c) \ set (\ p)" and jD:"j \ {1 .. c*?d}" + assume ?rhs then obtain e c j where ecR: "(e,c) \ set (\ p)" and jD:"j \ {1 .. c*?d}" and px: "Ifm (a#bs) (\ p c (Add e (C j)))" (is "?sp c e j") by blast from ecR have "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" by auto hence "(Inum (a#bs) e,c) \ (\(b,k). (Inum (a#bs) b,k)) ` R" using RR by simp @@ -5273,30 +5273,30 @@ (rule_tac x="j" in bexI, simp_all add: cc'[symmetric]) qed -lemma rl_thm': - assumes lp: "iszlfm p (real_of_int (i::int)#bs)" +lemma rl_thm': + assumes lp: "iszlfm p (real_of_int (i::int)#bs)" and R: "(\(b,k). (Inum (a#bs) b,k)) ` R = (\(b,k). (Inum (a#bs) b,k)) ` set (\ p)" shows "(\ (x::int). Ifm (real_of_int x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real_of_int j#bs) (minusinf p)) \ (\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" - using rl_thm[OF lp] \_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp + using rl_thm[OF lp] \_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp definition chooset :: "fm \ fm \ ((num\int) list) \ int" where "chooset p \ (let q = zlfm p ; d = \ q; - B = remdups (map (\ (t,k). (simpnum t,k)) (\ q)) ; + B = remdups (map (\ (t,k). (simpnum t,k)) (\ q)) ; a = remdups (map (\ (t,k). (simpnum t,k)) (\_\ q)) in if length B \ length a then (q,B,d) else (mirror q, a,d))" lemma chooset: assumes qf: "qfree p" - shows "\ q B d. chooset p = (q,B,d) \ - ((\ (x::int). Ifm (real_of_int x#bs) p) = - (\ (x::int). Ifm (real_of_int x#bs) q)) \ + shows "\ q B d. chooset p = (q,B,d) \ + ((\ (x::int). Ifm (real_of_int x#bs) p) = + (\ (x::int). Ifm (real_of_int x#bs) q)) \ ((\(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\ q)) \ (\ q = d) \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" proof- - fix q B d + fix q B d assume qBd: "chooset p = (q,B,d)" - let ?thes = "((\ (x::int). Ifm (real_of_int x#bs) p) = - (\ (x::int). Ifm (real_of_int x#bs) q)) \ ((\(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\ q)) \ - (\ q = d) \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" + let ?thes = "((\ (x::int). Ifm (real_of_int x#bs) p) = + (\ (x::int). Ifm (real_of_int x#bs) q)) \ ((\(t,k). (Inum (real_of_int i#bs) t,k)) ` set B = (\(t,k). (Inum (real_of_int i#bs) t,k)) ` set (\ q)) \ + (\ q = d) \ d >0 \ iszlfm q (real_of_int (i::int)#bs) \ (\ (e,c)\ set B. numbound0 e \ c>0)" let ?I = "\ (x::int) p. Ifm (real_of_int x#bs) p" let ?q = "zlfm p" let ?d = "\ ?q" @@ -5305,20 +5305,20 @@ let ?B'= "remdups (map ?f (\ ?q))" let ?A = "set (\_\ ?q)" let ?A'= "remdups (map ?f (\_\ ?q))" - from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] + from conjunct1[OF zlfm_I[OF qf, where bs="bs"]] have pp': "\ i. ?I i ?q = ?I i p" by auto - hence pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by simp + hence pq_ex:"(\ (x::int). ?I x p) = (\ x. ?I x ?q)" by simp from iszlfm_gen[OF conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]], rule_format, where y="real_of_int i"] - have lq: "iszlfm ?q (real_of_int (i::int)#bs)" . + have lq: "iszlfm ?q (real_of_int (i::int)#bs)" . from \[OF lq] have dp:"?d >0" by blast let ?N = "\ (t,c). (Inum (real_of_int (i::int)#bs) t,c)" have "?N ` set ?B' = ((?N o ?f) ` ?B)" by (simp add: split_def image_comp) also have "\ = ?N ` ?B" by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int i #bs"] image_def) finally have BB': "?N ` set ?B' = ?N ` ?B" . - have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_comp) + have "?N ` set ?A' = ((?N o ?f) ` ?A)" by (simp add: split_def image_comp) also have "\ = ?N ` ?A" using simpnum_ci[where bs="real_of_int i #bs"] - by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int i #bs"] image_def) + by(simp add: split_def image_comp simpnum_ci[where bs="real_of_int i #bs"] image_def) finally have AA': "?N ` set ?A' = ?N ` ?A" . from \_l[OF lq] have B_nb:"\ (e,c)\ set ?B'. numbound0 e \ c > 0" by (simp add: split_def) @@ -5327,16 +5327,16 @@ {assume "length ?B' \ length ?A'" hence q:"q=?q" and "B = ?B'" and d:"d = ?d" using qBd by (auto simp add: Let_def chooset_def) - with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" + with BB' B_nb have b: "?N ` (set B) = ?N ` set (\ q)" and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto with pq_ex dp lq q d have ?thes by simp} - moreover + moreover {assume "\ (length ?B' \ length ?A')" hence q:"q=mirror ?q" and "B = ?A'" and d:"d = ?d" using qBd by (auto simp add: Let_def chooset_def) - with AA' mirror_\_\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" - and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto - from mirror_ex[OF lq] pq_ex q + with AA' mirror_\_\[OF lq] A_nb have b:"?N ` (set B) = ?N ` set (\ q)" + and bn: "\(e,c)\ set B. numbound0 e \ c > 0" by auto + from mirror_ex[OF lq] pq_ex q have pqm_eq:"(\ (x::int). ?I x p) = (\ (x::int). ?I x q)" by simp from lq q mirror_l [where p="?q" and bs="bs" and a="real_of_int i"] have lq': "iszlfm q (real_of_int i#bs)" by auto @@ -5366,7 +5366,7 @@ qed definition redlove :: "fm \ fm" where - "redlove p \ + "redlove p \ (let (q,B,d) = chooset p; mq = simpfm (minusinf q); md = evaldjf (\ j. simpfm (subst0 (C j) mq)) [1..d] @@ -5375,7 +5375,7 @@ in decr (disj md qd)))" lemma redlove: assumes qf: "qfree p" - shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (redlove p))) \ qfree (redlove p)" + shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (redlove p))) \ qfree (redlove p)" (is "(?lhs = ?rhs) \ _") proof- @@ -5391,53 +5391,53 @@ let ?N = "\ (t,k). (Inum (real_of_int (i::int)#bs) t,k)" let ?qd = "evaldjf (stage ?q ?d) ?B" have qbf:"chooset p = (?q,?B,?d)" by simp - from chooset[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and - B:"?N ` set ?B = ?N ` set (\ ?q)" and dd: "\ ?q = ?d" and dp: "?d > 0" and - lq: "iszlfm ?q (real_of_int i#bs)" and + from chooset[OF qf qbf] have pq_ex: "(\(x::int). ?I x p) = (\ (x::int). ?I x ?q)" and + B:"?N ` set ?B = ?N ` set (\ ?q)" and dd: "\ ?q = ?d" and dp: "?d > 0" and + lq: "iszlfm ?q (real_of_int i#bs)" and Bn: "\ (e,c)\ set ?B. numbound0 e \ c > 0" by auto from zlin_qfree[OF lq] have qfq: "qfree ?q" . from simpfm_qf[OF minusinf_qfree[OF qfq]] have qfmq: "qfree ?smq". have jsnb: "\ j \ set ?js. numbound0 (C j)" by simp - hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" + hence "\ j\ set ?js. bound0 (subst0 (C j) ?smq)" by (auto simp only: subst0_bound0[OF qfmq]) hence th: "\ j\ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" by auto - from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp + from evaldjf_bound0[OF th] have mdb: "bound0 ?md" by simp from Bn stage_nb[OF lq] have th:"\ x \ set ?B. bound0 (stage ?q ?d x)" by auto from evaldjf_bound0[OF th] have qdb: "bound0 ?qd" . - from mdb qdb + from mdb qdb have mdqdb: "bound0 (disj ?md ?qd)" by (simp only: disj_def, cases "?md=T \ ?qd=T", simp_all) from trans [OF pq_ex rl_thm'[OF lq B]] dd have "?lhs = ((\ j\ {1.. ?d}. ?I j ?mq) \ (\ (e,c)\ set ?B. \ j\ {1 .. c*?d}. Ifm (real_of_int i#bs) (\ ?q c (Add e (C j)))))" by auto - also have "\ = ((\ j\ {1.. ?d}. ?I j ?smq) \ (\ (e,c)\ set ?B. ?I i (stage ?q ?d (e,c) )))" + also have "\ = ((\ j\ {1.. ?d}. ?I j ?smq) \ (\ (e,c)\ set ?B. ?I i (stage ?q ?d (e,c) )))" by (simp add: stage split_def) also have "\ = ((\ j\ {1 .. ?d}. ?I i (subst0 (C j) ?smq)) \ ?I i ?qd)" by (simp add: evaldjf_ex subst0_I[OF qfmq]) - finally have mdqd:"?lhs = (?I i ?md \ ?I i ?qd)" by (simp only: evaldjf_ex set_upto simpfm) + finally have mdqd:"?lhs = (?I i ?md \ ?I i ?qd)" by (simp only: evaldjf_ex set_upto simpfm) also have "\ = (?I i (disj ?md ?qd))" by simp - also have "\ = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) - finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . + also have "\ = (Ifm bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) + finally have mdqd2: "?lhs = (Ifm bs (decr (disj ?md ?qd)))" . {assume mdT: "?md = T" hence cT:"redlove p = T" by (simp add: redlove_def Let_def chooset_def split_def) - from mdT have lhs:"?lhs" using mdqd by simp + from mdT have lhs:"?lhs" using mdqd by simp from mdT have "?rhs" by (simp add: redlove_def chooset_def split_def) with lhs cT have ?thesis by simp } moreover - {assume mdT: "?md \ T" hence "redlove p = decr (disj ?md ?qd)" + {assume mdT: "?md \ T" hence "redlove p = decr (disj ?md ?qd)" by (simp add: redlove_def chooset_def split_def Let_def) with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } ultimately show ?thesis by blast qed -lemma DJredlove: +lemma DJredlove: assumes qf: "qfree p" shows "((\ (x::int). Ifm (real_of_int x#bs) p) = (Ifm bs (DJ redlove p))) \ qfree (DJ redlove p)" proof- from redlove have cqf: "\ p. qfree p \ qfree (redlove p)" by blast from DJ_qf[OF cqf] qf have thqf:"qfree (DJ redlove p)" by blast - have "Ifm bs (DJ redlove p) = (\ q\ set (disjuncts p). Ifm bs (redlove q))" + have "Ifm bs (DJ redlove p) = (\ q\ set (disjuncts p). Ifm bs (redlove q))" by (simp add: DJ_def evaldjf_ex) - also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real_of_int x#bs) q)" + also have "\ = (\ q \ set(disjuncts p). \ (x::int). Ifm (real_of_int x#bs) q)" using redlove disjuncts_qf[OF qf] by blast also have "\ = (\ (x::int). Ifm (real_of_int x#bs) p)" by (induct p rule: disjuncts.induct, auto) finally show ?thesis using thqf by blast @@ -5461,13 +5461,13 @@ show "qfree (mircfr p)\(Ifm bs (mircfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") proof- let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" - have "?rhs = (\ (i::int). \ x. Ifm (x#real_of_int i#bs) ?es)" + have "?rhs = (\ (i::int). \ x. Ifm (x#real_of_int i#bs) ?es)" using splitex[OF qf] by simp with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real_of_int i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ with DJcooper[OF qf'] show ?thesis by (simp add: mircfr_def) qed qed - + lemma mirlfr: "\ bs p. qfree p \ qfree(mirlfr p) \ Ifm bs (mirlfr p) = Ifm bs (E p)" proof(clarsimp simp del: Ifm.simps) fix bs p @@ -5475,13 +5475,13 @@ show "qfree (mirlfr p)\(Ifm bs (mirlfr p) = Ifm bs (E p))" (is "_ \ (?lhs = ?rhs)") proof- let ?es = "(And (And (Ge (CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) (simpfm (exsplit p)))" - have "?rhs = (\ (i::int). \ x. Ifm (x#real_of_int i#bs) ?es)" + have "?rhs = (\ (i::int). \ x. Ifm (x#real_of_int i#bs) ?es)" using splitex[OF qf] by simp with ferrack01[OF simpfm_qf[OF exsplit_qf[OF qf]]] have th1: "?rhs = (\ (i::int). Ifm (real_of_int i#bs) (ferrack01 (simpfm (exsplit p))))" and qf':"qfree (ferrack01 (simpfm (exsplit p)))" by simp+ with DJredlove[OF qf'] show ?thesis by (simp add: mirlfr_def) qed qed - + definition mircfrqe:: "fm \ fm" where "mircfrqe p = qelim (prep p) mircfr" @@ -5566,7 +5566,7 @@ | fm_of_term vs (@{term "op \ :: real \ real \ bool"} $ t1 $ t2) = @{code Le} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) | fm_of_term vs (@{term "op = :: real \ real \ bool"} $ t1 $ t2) = - @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) + @{code Eq} (@{code Sub} (num_of_term vs t1, num_of_term vs t2)) | fm_of_term vs (@{term "op rdvd"} $ (@{term "of_int :: int \ real"} $ (@{term "numeral :: _ \ int"} $ t1)) $ t2) = mk_Dvd (HOLogic.dest_num t1, num_of_term vs t2) | fm_of_term vs (@{term "op rdvd"} $ (@{term "of_int :: int \ real"} $ (@{term "- numeral :: _ \ int"} $ t1)) $ t2) = @@ -5606,7 +5606,7 @@ | term_of_num vs (@{code CN} (n, i, t)) = term_of_num vs (@{code Add} (@{code Mul} (i, @{code Bound} n), t)) | term_of_num vs (@{code CF} (c, t, s)) = term_of_num vs (@{code Add} (@{code Mul} (c, @{code Floor} t), s)); -fun term_of_fm vs @{code T} = @{term True} +fun term_of_fm vs @{code T} = @{term True} | term_of_fm vs @{code F} = @{term False} | term_of_fm vs (@{code Lt} t) = @{term "op < :: real \ real \ bool"} $ term_of_num vs t $ @{term "0::real"} @@ -5637,7 +5637,7 @@ in fn (ctxt, t) => - let + let val fs = Misc_Legacy.term_frees t; val vs = map_index swap fs; (*If quick_and_dirty then run without proof generation as oracle*) @@ -5647,8 +5647,8 @@ end; \ -lemmas iff_real_of_int = of_int_eq_iff [where 'a = real, symmetric] - of_int_less_iff [where 'a = real, symmetric] +lemmas iff_real_of_int = of_int_eq_iff [where 'a = real, symmetric] + of_int_less_iff [where 'a = real, symmetric] of_int_le_iff [where 'a = real, symmetric] ML_file "mir_tac.ML" @@ -5665,7 +5665,7 @@ by mir lemma "\x::real. 2*\x\ \ \2*x\ \ \2*x\ \ 2*\x+1\" - by mir + by mir lemma "\x::real. \y \ x. (\x\ = \y\)" by mir