Up to index of Isabelle/HOL/HOL-Complex/ex
theory ReflectedFerrack(* Title: Complex/ex/ReflectedFerrack.thy Author: Amine Chaieb *) header {* Quatifier elimination for R(0,1,+,<) *} theory ReflectedFerrack imports GCD Real Efficient_Nat uses ("linreif.ML") ("linrtac.ML") begin (*********************************************************************************) (* SOME GENERAL STUFF< HAS TO BE MOVED IN SOME LIB *) (*********************************************************************************) consts alluopairs:: "'a list => ('a × 'a) list" primrec "alluopairs [] = []" "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)" lemma alluopairs_set1: "set (alluopairs xs) ≤ {(x,y). x∈ set xs ∧ y∈ set xs}" by (induct xs, auto) lemma alluopairs_set: "[|x∈ set xs ; y ∈ set xs|] ==> (x,y) ∈ set (alluopairs xs) ∨ (y,x) ∈ set (alluopairs xs) " by (induct xs, auto) lemma alluopairs_ex: assumes Pc: "∀ x y. P x y = P y x" shows "(∃ x ∈ set xs. ∃ y ∈ set xs. P x y) = (∃ (x,y) ∈ set (alluopairs xs). P x y)" proof assume "∃x∈set xs. ∃y∈set xs. P x y" then obtain x y where x: "x ∈ set xs" and y:"y ∈ set xs" and P: "P x y" by blast from alluopairs_set[OF x y] P Pc show"∃(x, y)∈set (alluopairs xs). P x y" by auto next assume "∃(x, y)∈set (alluopairs xs). P x y" then obtain "x" and "y" where xy:"(x,y) ∈ set (alluopairs xs)" and P: "P x y" by blast+ from xy have "x ∈ set xs ∧ y∈ set xs" using alluopairs_set1 by blast with P show "∃x∈set xs. ∃y∈set xs. P x y" by blast qed lemma nth_pos2: "0 < n ==> (x#xs) ! n = xs ! (n - 1)" using Nat.gr0_conv_Suc by clarsimp lemma filter_length: "length (List.filter P xs) < Suc (length xs)" apply (induct xs, auto) done consts remdps:: "'a list => 'a list" recdef remdps "measure size" "remdps [] = []" "remdps (x#xs) = (x#(remdps (List.filter (λ y. y ≠ x) xs)))" (hints simp add: filter_length[rule_format]) lemma remdps_set[simp]: "set (remdps xs) = set xs" by (induct xs rule: remdps.induct, auto) (*********************************************************************************) (**** SHADOW SYNTAX AND SEMANTICS ****) (*********************************************************************************) datatype num = C int | Bound nat | CN nat int num | Neg num | Add num num| Sub num num | Mul int num (* A size for num to make inductive proofs simpler*) consts num_size :: "num => nat" primrec "num_size (C c) = 1" "num_size (Bound n) = 1" "num_size (Neg a) = 1 + num_size a" "num_size (Add a b) = 1 + num_size a + num_size b" "num_size (Sub a b) = 3 + num_size a + num_size b" "num_size (Mul c a) = 1 + num_size a" "num_size (CN n c a) = 3 + num_size a " (* Semantics of numeral terms (num) *) consts Inum :: "real list => num => real" primrec "Inum bs (C c) = (real c)" "Inum bs (Bound n) = bs!n" "Inum bs (CN n c a) = (real c) * (bs!n) + (Inum bs a)" "Inum bs (Neg a) = -(Inum bs a)" "Inum bs (Add a b) = Inum bs a + Inum bs b" "Inum bs (Sub a b) = Inum bs a - Inum bs b" "Inum bs (Mul c a) = (real c) * Inum bs a" (* FORMULAE *) datatype fm = T| F| Lt num| Le num| Gt num| Ge num| Eq num| NEq num| NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm (* A size for fm *) consts fmsize :: "fm => nat" recdef fmsize "measure size" "fmsize (NOT p) = 1 + fmsize p" "fmsize (And p q) = 1 + fmsize p + fmsize q" "fmsize (Or p q) = 1 + fmsize p + fmsize q" "fmsize (Imp p q) = 3 + fmsize p + fmsize q" "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)" "fmsize (E p) = 1 + fmsize p" "fmsize (A p) = 4+ fmsize p" "fmsize p = 1" (* several lemmas about fmsize *) lemma fmsize_pos: "fmsize p > 0" by (induct p rule: fmsize.induct) simp_all (* Semantics of formulae (fm) *) consts Ifm ::"real list => fm => bool" primrec "Ifm bs T = True" "Ifm bs F = False" "Ifm bs (Lt a) = (Inum bs a < 0)" "Ifm bs (Gt a) = (Inum bs a > 0)" "Ifm bs (Le a) = (Inum bs a ≤ 0)" "Ifm bs (Ge a) = (Inum bs a ≥ 0)" "Ifm bs (Eq a) = (Inum bs a = 0)" "Ifm bs (NEq a) = (Inum bs a ≠ 0)" "Ifm bs (NOT p) = (¬ (Ifm bs p))" "Ifm bs (And p q) = (Ifm bs p ∧ Ifm bs q)" "Ifm bs (Or p q) = (Ifm bs p ∨ Ifm bs q)" "Ifm bs (Imp p q) = ((Ifm bs p) --> (Ifm bs q))" "Ifm bs (Iff p q) = (Ifm bs p = Ifm bs q)" "Ifm bs (E p) = (∃ x. Ifm (x#bs) p)" "Ifm bs (A p) = (∀ x. Ifm (x#bs) p)" lemma IfmLeSub: "[| Inum bs s = s' ; Inum bs t = t' |] ==> Ifm bs (Le (Sub s t)) = (s' ≤ t')" apply simp done lemma IfmLtSub: "[| Inum bs s = s' ; Inum bs t = t' |] ==> Ifm bs (Lt (Sub s t)) = (s' < t')" apply simp done lemma IfmEqSub: "[| Inum bs s = s' ; Inum bs t = t' |] ==> Ifm bs (Eq (Sub s t)) = (s' = t')" apply simp done lemma IfmNOT: " (Ifm bs p = P) ==> (Ifm bs (NOT p) = (¬P))" apply simp done lemma IfmAnd: " [| Ifm bs p = P ; Ifm bs q = Q|] ==> (Ifm bs (And p q) = (P ∧ Q))" apply simp done lemma IfmOr: " [| Ifm bs p = P ; Ifm bs q = Q|] ==> (Ifm bs (Or p q) = (P ∨ Q))" apply simp done lemma IfmImp: " [| Ifm bs p = P ; Ifm bs q = Q|] ==> (Ifm bs (Imp p q) = (P --> Q))" apply simp done lemma IfmIff: " [| Ifm bs p = P ; Ifm bs q = Q|] ==> (Ifm bs (Iff p q) = (P = Q))" apply simp done lemma IfmE: " (!! x. Ifm (x#bs) p = P x) ==> (Ifm bs (E p) = (∃x. P x))" apply simp done lemma IfmA: " (!! x. Ifm (x#bs) p = P x) ==> (Ifm bs (A p) = (∀x. P x))" apply simp done consts not:: "fm => fm" recdef not "measure size" "not (NOT p) = p" "not T = F" "not F = T" "not p = NOT p" lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)" by (cases p) auto constdefs conj :: "fm => fm => fm" "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) constdefs disj :: "fm => fm => fm" "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) constdefs imp :: "fm => fm => fm" "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) constdefs iff :: "fm => fm => fm" "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)" by (unfold iff_def,cases "p=q", simp,cases "p=NOT q", simp) (cases "NOT p= q", auto) lemma conj_simps: "conj F Q = F" "conj P F = F" "conj T Q = Q" "conj P T = P" "conj P P = P" "P ≠ T ==> P ≠ F ==> Q ≠ T ==> Q ≠ F ==> P ≠ Q ==> conj P Q = And P Q" by (simp_all add: conj_def) lemma disj_simps: "disj T Q = T" "disj P T = T" "disj F Q = Q" "disj P F = P" "disj P P = P" "P ≠ T ==> P ≠ F ==> Q ≠ T ==> Q ≠ F ==> P ≠ Q ==> disj P Q = Or P Q" by (simp_all add: disj_def) lemma imp_simps: "imp F Q = T" "imp P T = T" "imp T Q = Q" "imp P F = not P" "imp P P = T" "P ≠ T ==> P ≠ F ==> P ≠ Q ==> Q ≠ T ==> Q ≠ F ==> imp P Q = Imp P Q" by (simp_all add: imp_def) lemma trivNOT: "p ≠ NOT p" "NOT p ≠ p" apply (induct p, auto) done lemma iff_simps: "iff p p = T" "iff p (NOT p) = F" "iff (NOT p) p = F" "iff p F = not p" "iff F p = not p" "p ≠ NOT T ==> iff T p = p" "p≠ NOT T ==> iff p T = p" "p≠q ==> p≠ NOT q ==> q≠ NOT p ==> p≠ F ==> q≠ F ==> p ≠ T ==> q ≠ T ==> iff p q = Iff p q" using trivNOT by (simp_all add: iff_def, cases p, auto) (* Quantifier freeness *) consts qfree:: "fm => bool" recdef qfree "measure size" "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 (Iff p q) = (qfree p ∧ qfree q)" "qfree p = True" (* Boundedness and substitution *) consts numbound0:: "num => bool" (* a num is INDEPENDENT of Bound 0 *) bound0:: "fm => bool" (* A Formula is independent of Bound 0 *) primrec "numbound0 (C c) = True" "numbound0 (Bound n) = (n>0)" "numbound0 (CN n c 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 (Mul i a) = numbound0 a" lemma numbound0_I: assumes nb: "numbound0 a" shows "Inum (b#bs) a = Inum (b'#bs) a" using nb by (induct a rule: numbound0.induct,auto simp add: nth_pos2) primrec "bound0 T = True" "bound0 F = True" "bound0 (Lt a) = numbound0 a" "bound0 (Le a) = numbound0 a" "bound0 (Gt a) = numbound0 a" "bound0 (Ge a) = numbound0 a" "bound0 (Eq a) = numbound0 a" "bound0 (NEq a) = numbound0 a" "bound0 (NOT p) = bound0 p" "bound0 (And p q) = (bound0 p ∧ bound0 q)" "bound0 (Or p q) = (bound0 p ∧ bound0 q)" "bound0 (Imp p q) = ((bound0 p) ∧ (bound0 q))" "bound0 (Iff p q) = (bound0 p ∧ bound0 q)" "bound0 (E p) = False" "bound0 (A p) = False" lemma bound0_I: assumes bp: "bound0 p" shows "Ifm (b#bs) p = Ifm (b'#bs) p" using bp numbound0_I[where b="b" and bs="bs" and b'="b'"] by (induct p rule: bound0.induct) (auto simp add: nth_pos2) lemma not_qf[simp]: "qfree p ==> qfree (not p)" by (cases p, auto) lemma not_bn[simp]: "bound0 p ==> bound0 (not p)" by (cases p, auto) lemma conj_qf[simp]: "[|qfree p ; qfree q|] ==> qfree (conj p q)" using conj_def by auto lemma conj_nb[simp]: "[|bound0 p ; bound0 q|] ==> bound0 (conj p q)" using conj_def by auto lemma disj_qf[simp]: "[|qfree p ; qfree q|] ==> qfree (disj p q)" using disj_def by auto lemma disj_nb[simp]: "[|bound0 p ; bound0 q|] ==> bound0 (disj p q)" using disj_def by auto lemma imp_qf[simp]: "[|qfree p ; qfree q|] ==> qfree (imp p q)" using imp_def by (cases "p=F ∨ q=T",simp_all add: imp_def) lemma imp_nb[simp]: "[|bound0 p ; bound0 q|] ==> bound0 (imp p q)" using imp_def by (cases "p=F ∨ q=T ∨ p=q",simp_all add: imp_def) lemma iff_qf[simp]: "[|qfree p ; qfree q|] ==> qfree (iff p q)" by (unfold iff_def,cases "p=q", auto) lemma iff_nb[simp]: "[|bound0 p ; bound0 q|] ==> bound0 (iff p q)" using iff_def by (unfold iff_def,cases "p=q", auto) consts decrnum:: "num => num" decr :: "fm => fm" recdef decrnum "measure size" "decrnum (Bound n) = Bound (n - 1)" "decrnum (Neg a) = Neg (decrnum a)" "decrnum (Add a b) = Add (decrnum a) (decrnum b)" "decrnum (Sub a b) = Sub (decrnum a) (decrnum b)" "decrnum (Mul c a) = Mul c (decrnum a)" "decrnum (CN n c a) = CN (n - 1) c (decrnum a)" "decrnum a = a" recdef decr "measure size" "decr (Lt a) = Lt (decrnum a)" "decr (Le a) = Le (decrnum a)" "decr (Gt a) = Gt (decrnum a)" "decr (Ge a) = Ge (decrnum a)" "decr (Eq a) = Eq (decrnum a)" "decr (NEq a) = NEq (decrnum a)" "decr (NOT p) = NOT (decr p)" "decr (And p q) = conj (decr p) (decr q)" "decr (Or p q) = disj (decr p) (decr q)" "decr (Imp p q) = imp (decr p) (decr q)" "decr (Iff p q) = iff (decr p) (decr q)" "decr p = p" lemma decrnum: assumes nb: "numbound0 t" shows "Inum (x#bs) t = Inum bs (decrnum t)" using nb by (induct t rule: decrnum.induct, simp_all add: nth_pos2) lemma decr: assumes nb: "bound0 p" shows "Ifm (x#bs) p = Ifm bs (decr p)" using nb by (induct p rule: decr.induct, simp_all add: nth_pos2 decrnum) lemma decr_qf: "bound0 p ==> qfree (decr p)" by (induct p, simp_all) consts isatom :: "fm => bool" (* test for atomicity *) recdef isatom "measure size" "isatom T = True" "isatom F = True" "isatom (Lt a) = True" "isatom (Le a) = True" "isatom (Gt a) = True" "isatom (Ge a) = True" "isatom (Eq a) = True" "isatom (NEq a) = True" "isatom p = False" lemma bound0_qf: "bound0 p ==> qfree p" by (induct p, simp_all) constdefs djf:: "('a => fm) => 'a => fm => fm" "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 (f p) q))" constdefs evaldjf:: "('a => fm) => 'a list => fm" "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) lemma djf_simps: "djf f p T = T" "djf f p F = f p" "q≠T ==> q≠F ==> djf f p q = (let fp = f p in case fp of T => T | F => q | _ => Or (f p) q)" by (simp_all add: 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: assumes nb: "∀ x∈ set xs. bound0 (f x)" shows "bound0 (evaldjf f xs)" using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) lemma evaldjf_qf: assumes nb: "∀ x∈ set xs. qfree (f x)" shows "qfree (evaldjf f xs)" using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto) consts disjuncts :: "fm => fm list" recdef disjuncts "measure size" "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)" "disjuncts F = []" "disjuncts p = [p]" lemma disjuncts: "(∃ q∈ set (disjuncts p). Ifm bs q) = Ifm bs p" by(induct p rule: disjuncts.induct, auto) lemma disjuncts_nb: "bound0 p ==> ∀ q∈ set (disjuncts p). bound0 q" proof- assume nb: "bound0 p" hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto) thus ?thesis by (simp only: list_all_iff) qed lemma disjuncts_qf: "qfree p ==> ∀ q∈ set (disjuncts p). qfree q" proof- assume qf: "qfree p" hence "list_all qfree (disjuncts p)" by (induct p rule: disjuncts.induct, auto) thus ?thesis by (simp only: list_all_iff) qed constdefs DJ :: "(fm => fm) => fm => fm" "DJ f p ≡ evaldjf f (disjuncts p)" lemma DJ: assumes fdj: "∀ p q. Ifm bs (f (Or p q)) = Ifm bs (Or (f p) (f q))" and fF: "f F = F" 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) also have "… = Ifm bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto) finally show ?thesis . qed lemma DJ_qf: assumes fqf: "∀ p. qfree p --> qfree (f p)" shows "∀p. qfree p --> qfree (DJ f p) " proof(clarify) fix p assume qf: "qfree p" 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 lemma DJ_qe: assumes qe: "∀ bs p. qfree p --> qfree (qe p) ∧ (Ifm bs (qe p) = Ifm bs (E p))" shows "∀ bs p. qfree p --> qfree (DJ qe p) ∧ (Ifm bs ((DJ qe p)) = Ifm bs (E p))" proof(clarify) fix p::fm and bs assume qf: "qfree p" from qe have qth: "∀ p. qfree p --> qfree (qe p)" by blast from DJ_qf[OF qth] qf have qfth:"qfree (DJ qe p)" by auto have "Ifm bs (DJ qe p) = (∃ q∈ set (disjuncts p). Ifm bs (qe q))" by (simp add: DJ_def evaldjf_ex) also have "… = (∃ q ∈ set(disjuncts p). Ifm bs (E q))" using qe disjuncts_qf[OF qf] by auto also have "… = Ifm bs (E p)" by (induct p rule: disjuncts.induct, auto) finally show "qfree (DJ qe p) ∧ Ifm bs (DJ qe p) = Ifm bs (E p)" using qfth by blast qed (* Simplification *) consts numgcd :: "num => int" numgcdh:: "num => int => int" reducecoeffh:: "num => int => num" reducecoeff :: "num => num" dvdnumcoeff:: "num => int => bool" consts maxcoeff:: "num => int" recdef maxcoeff "measure size" "maxcoeff (C i) = abs i" "maxcoeff (CN n c t) = max (abs c) (maxcoeff t)" "maxcoeff t = 1" lemma maxcoeff_pos: "maxcoeff t ≥ 0" by (induct t rule: maxcoeff.induct, auto) recdef numgcdh "measure size" "numgcdh (C i) = (λg. igcd i g)" "numgcdh (CN n c t) = (λg. igcd c (numgcdh t g))" "numgcdh t = (λg. 1)" defs numgcd_def [code func]: "numgcd t ≡ numgcdh t (maxcoeff t)" recdef reducecoeffh "measure size" "reducecoeffh (C i) = (λ g. C (i div g))" "reducecoeffh (CN n c t) = (λ g. CN n (c div g) (reducecoeffh t g))" "reducecoeffh t = (λg. t)" defs reducecoeff_def: "reducecoeff t ≡ (let g = numgcd t in if g = 0 then C 0 else if g=1 then t else reducecoeffh t g)" recdef dvdnumcoeff "measure size" "dvdnumcoeff (C i) = (λ g. g dvd i)" "dvdnumcoeff (CN n c t) = (λ g. g dvd c ∧ (dvdnumcoeff t g))" "dvdnumcoeff t = (λg. False)" lemma dvdnumcoeff_trans: assumes gdg: "g dvd g'" and dgt':"dvdnumcoeff t g'" shows "dvdnumcoeff t g" using dgt' gdg by (induct t rule: dvdnumcoeff.induct, simp_all add: gdg zdvd_trans[OF gdg]) declare zdvd_trans [trans add] lemma natabs0: "(nat (abs x) = 0) = (x = 0)" by arith lemma numgcd0: assumes g0: "numgcd t = 0" shows "Inum bs t = 0" using g0[simplified numgcd_def] by (induct t rule: numgcdh.induct, auto simp add: igcd_def gcd_zero natabs0 max_def maxcoeff_pos) lemma numgcdh_pos: assumes gp: "g ≥ 0" shows "numgcdh t g ≥ 0" using gp by (induct t rule: numgcdh.induct, auto simp add: igcd_def) lemma numgcd_pos: "numgcd t ≥0" by (simp add: numgcd_def numgcdh_pos maxcoeff_pos) lemma reducecoeffh: assumes gt: "dvdnumcoeff t g" and gp: "g > 0" shows "real g *(Inum bs (reducecoeffh t g)) = Inum bs t" using gt proof(induct t rule: reducecoeffh.induct) case (1 i) hence gd: "g dvd i" by simp from gp have gnz: "g ≠ 0" by simp from prems show ?case by (simp add: real_of_int_div[OF gnz gd]) next case (2 n c t) hence gd: "g dvd c" by simp from gp have gnz: "g ≠ 0" by simp from prems show ?case by (simp add: real_of_int_div[OF gnz gd] ring_simps) qed (auto simp add: numgcd_def gp) consts ismaxcoeff:: "num => int => bool" recdef ismaxcoeff "measure size" "ismaxcoeff (C i) = (λ x. abs i ≤ x)" "ismaxcoeff (CN n c t) = (λx. abs c ≤ x ∧ (ismaxcoeff t x))" "ismaxcoeff t = (λx. True)" lemma ismaxcoeff_mono: "ismaxcoeff t c ==> c ≤ c' ==> ismaxcoeff t c'" by (induct t rule: ismaxcoeff.induct, auto) lemma maxcoeff_ismaxcoeff: "ismaxcoeff t (maxcoeff t)" proof (induct t rule: maxcoeff.induct) case (2 n c t) hence H:"ismaxcoeff t (maxcoeff t)" . have thh: "maxcoeff t ≤ max (abs c) (maxcoeff t)" by (simp add: le_maxI2) from ismaxcoeff_mono[OF H thh] show ?case by (simp add: le_maxI1) qed simp_all lemma igcd_gt1: "igcd i j > 1 ==> ((abs i > 1 ∧ abs j > 1) ∨ (abs i = 0 ∧ abs j > 1) ∨ (abs i > 1 ∧ abs j = 0))" apply (cases "abs i = 0", simp_all add: igcd_def) apply (cases "abs j = 0", simp_all) apply (cases "abs i = 1", simp_all) apply (cases "abs j = 1", simp_all) apply auto done lemma numgcdh0:"numgcdh t m = 0 ==> m =0" by (induct t rule: numgcdh.induct, auto simp add:igcd0) lemma dvdnumcoeff_aux: assumes "ismaxcoeff t m" and mp:"m ≥ 0" and "numgcdh t m > 1" shows "dvdnumcoeff t (numgcdh t m)" using prems proof(induct t rule: numgcdh.induct) case (2 n c t) let ?g = "numgcdh t m" from prems have th:"igcd c ?g > 1" by simp from igcd_gt1[OF th] numgcdh_pos[OF mp, where t="t"] have "(abs c > 1 ∧ ?g > 1) ∨ (abs c = 0 ∧ ?g > 1) ∨ (abs c > 1 ∧ ?g = 0)" by simp moreover {assume "abs c > 1" and gp: "?g > 1" with prems have th: "dvdnumcoeff t ?g" by simp have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2) from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1)} moreover {assume "abs c = 0 ∧ ?g > 1" with prems have th: "dvdnumcoeff t ?g" by simp have th': "igcd c ?g dvd ?g" by (simp add:igcd_dvd2) from dvdnumcoeff_trans[OF th' th] have ?case by (simp add: igcd_dvd1) hence ?case by simp } moreover {assume "abs c > 1" and g0:"?g = 0" from numgcdh0[OF g0] have "m=0". with prems have ?case by simp } ultimately show ?case by blast qed(auto simp add: igcd_dvd1) lemma dvdnumcoeff_aux2: assumes "numgcd t > 1" shows "dvdnumcoeff t (numgcd t) ∧ numgcd t > 0" using prems proof (simp add: numgcd_def) let ?mc = "maxcoeff t" let ?g = "numgcdh t ?mc" have th1: "ismaxcoeff t ?mc" by (rule maxcoeff_ismaxcoeff) have th2: "?mc ≥ 0" by (rule maxcoeff_pos) assume H: "numgcdh t ?mc > 1" from dvdnumcoeff_aux[OF th1 th2 H] show "dvdnumcoeff t ?g" . qed lemma reducecoeff: "real (numgcd t) * (Inum bs (reducecoeff t)) = Inum bs t" proof- 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 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)} ultimately show ?thesis by blast qed lemma reducecoeffh_numbound0: "numbound0 t ==> numbound0 (reducecoeffh t g)" by (induct t rule: reducecoeffh.induct, auto) lemma reducecoeff_numbound0: "numbound0 t ==> numbound0 (reducecoeff t)" using reducecoeffh_numbound0 by (simp add: reducecoeff_def Let_def) consts simpnum:: "num => num" numadd:: "num × num => num" nummul:: "num => int => num" recdef numadd "measure (λ (t,s). size t + size s)" "numadd (CN n1 c1 r1,CN n2 c2 r2) = (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 (C b1, C b2) = C (b1+b2)" "numadd (a,b) = Add a b" lemma numadd[simp]: "Inum bs (numadd (t,s)) = Inum bs (Add t s)" apply (induct t s rule: numadd.induct, simp_all add: Let_def) apply (case_tac "c1+c2 = 0",case_tac "n1 ≤ n2", simp_all) apply (case_tac "n1 = n2", simp_all add: ring_simps) by (simp only: left_distrib[symmetric],simp) lemma numadd_nb[simp]: "[| numbound0 t ; numbound0 s|] ==> numbound0 (numadd (t,s))" by (induct t s rule: numadd.induct, auto simp add: Let_def) recdef nummul "measure size" "nummul (C j) = (λ i. C (i*j))" "nummul (CN n c a) = (λ i. CN n (i*c) (nummul a i))" "nummul t = (λ i. Mul i t)" lemma nummul[simp]: "!! i. Inum bs (nummul t i) = Inum bs (Mul i t)" by (induct t rule: nummul.induct, auto simp add: ring_simps) lemma nummul_nb[simp]: "!! i. numbound0 t ==> numbound0 (nummul t i)" by (induct t rule: nummul.induct, auto ) constdefs numneg :: "num => num" "numneg t ≡ nummul t (- 1)" constdefs numsub :: "num => num => num" "numsub s t ≡ (if s = t then C 0 else numadd (s,numneg t))" lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" using numneg_def by simp lemma numneg_nb[simp]: "numbound0 t ==> numbound0 (numneg t)" using numneg_def by simp lemma numsub[simp]: "Inum bs (numsub a b) = Inum bs (Sub a b)" using numsub_def by simp lemma numsub_nb[simp]: "[| numbound0 t ; numbound0 s|] ==> numbound0 (numsub t s)" using numsub_def by simp recdef simpnum "measure size" "simpnum (C j) = C j" "simpnum (Bound n) = CN n 1 (C 0)" "simpnum (Neg t) = numneg (simpnum t)" "simpnum (Add t s) = numadd (simpnum t,simpnum s)" "simpnum (Sub t s) = numsub (simpnum t) (simpnum s)" "simpnum (Mul i t) = (if i = 0 then (C 0) else nummul (simpnum t) i)" "simpnum (CN n c t) = (if c = 0 then simpnum t else numadd (CN n c (C 0),simpnum t))" lemma simpnum_ci[simp]: "Inum bs (simpnum t) = Inum bs t" by (induct t rule: simpnum.induct, auto simp add: numneg numadd numsub nummul) lemma simpnum_numbound0[simp]: "numbound0 t ==> numbound0 (simpnum t)" by (induct t rule: simpnum.induct, auto) consts nozerocoeff:: "num => bool" recdef nozerocoeff "measure size" "nozerocoeff (C c) = True" "nozerocoeff (CN n c t) = (c≠0 ∧ nozerocoeff t)" "nozerocoeff t = True" lemma numadd_nz : "nozerocoeff a ==> nozerocoeff b ==> nozerocoeff (numadd (a,b))" by (induct a b rule: numadd.induct,auto simp add: Let_def) lemma nummul_nz : "!! i. i≠0 ==> nozerocoeff a ==> nozerocoeff (nummul a i)" by (induct a rule: nummul.induct,auto simp add: Let_def numadd_nz) lemma numneg_nz : "nozerocoeff a ==> nozerocoeff (numneg a)" by (simp add: numneg_def nummul_nz) lemma numsub_nz: "nozerocoeff a ==> nozerocoeff b ==> nozerocoeff (numsub a b)" by (simp add: numsub_def numneg_nz numadd_nz) lemma simpnum_nz: "nozerocoeff (simpnum t)" by(induct t rule: simpnum.induct, auto simp add: numadd_nz numneg_nz numsub_nz nummul_nz) lemma maxcoeff_nz: "nozerocoeff t ==> maxcoeff t = 0 ==> t = C 0" proof (induct t rule: maxcoeff.induct) case (2 n c 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 add: le_maxI1) with cnz have "max (abs c) (maxcoeff t) > 0" by arith with prems show ?case by simp qed auto lemma numgcd_nz: assumes nz: "nozerocoeff t" and g0: "numgcd t = 0" shows "t = C 0" proof- from g0 have th:"numgcdh t (maxcoeff t) = 0" by (simp add: numgcd_def) from numgcdh0[OF th] have th:"maxcoeff t = 0" . from maxcoeff_nz[OF nz th] show ?thesis . qed constdefs simp_num_pair:: "(num × int) => num × int" "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' = igcd n g in if g' = 1 then (t',n) else (reducecoeffh t' g', n div g')) else (t',n))))" lemma simp_num_pair_ci: shows "((λ (t,n). Inum bs t / real n) (simp_num_pair (t,n))) = ((λ (t,n). Inum bs t / real n) (t,n))" (is "?lhs = ?rhs") proof- let ?t' = "simpnum t" let ?g = "numgcd ?t'" let ?g' = "igcd n ?g" {assume nz: "n = 0" hence ?thesis by (simp add: Let_def simp_num_pair_def)} moreover { assume nnz: "n ≠ 0" {assume "¬ ?g > 1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)} moreover {assume g1:"?g>1" hence g0: "?g > 0" by simp from igcd0 g1 nnz have gp0: "?g' ≠ 0" by simp hence g'p: "?g' > 0" using igcd_pos[where i="n" and j="numgcd ?t'"] by arith hence "?g'= 1 ∨ ?g' > 1" by arith moreover {assume "?g'=1" hence ?thesis by (simp add: Let_def simp_num_pair_def simpnum_ci)} moreover {assume g'1:"?g'>1" from dvdnumcoeff_aux2[OF g1] have th1:"dvdnumcoeff ?t' ?g" .. let ?tt = "reducecoeffh ?t' ?g'" let ?t = "Inum bs ?tt" have gpdg: "?g' dvd ?g" by (simp add: igcd_dvd2) have gpdd: "?g' dvd n" by (simp add: igcd_dvd1) have gpdgp: "?g' dvd ?g'" by simp from reducecoeffh[OF dvdnumcoeff_trans[OF gpdg th1] g'p] have th2:"real ?g' * ?t = Inum bs ?t'" by simp from prems have "?lhs = ?t / real (n div ?g')" by (simp add: simp_num_pair_def Let_def) also have "… = (real ?g' * ?t) / (real ?g' * (real (n div ?g')))" by simp also have "… = (Inum bs ?t' / real n)" using real_of_int_div[OF gp0 gpdd] th2 gp0 by simp finally have "?lhs = Inum bs t / real n" by (simp add: simpnum_ci) then have ?thesis using prems by (simp add: simp_num_pair_def)} ultimately have ?thesis by blast} ultimately have ?thesis by blast} ultimately show ?thesis by blast qed lemma simp_num_pair_l: assumes tnb: "numbound0 t" and np: "n >0" and tn: "simp_num_pair (t,n) = (t',n')" shows "numbound0 t' ∧ n' >0" proof- let ?t' = "simpnum t" let ?g = "numgcd ?t'" let ?g' = "igcd n ?g" {assume nz: "n = 0" hence ?thesis using prems by (simp add: Let_def simp_num_pair_def)} moreover { assume nnz: "n ≠ 0" {assume "¬ ?g > 1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)} moreover {assume g1:"?g>1" hence g0: "?g > 0" by simp from igcd0 g1 nnz have gp0: "?g' ≠ 0" by simp hence g'p: "?g' > 0" using igcd_pos[where i="n" and j="numgcd ?t'"] by arith hence "?g'= 1 ∨ ?g' > 1" by arith moreover {assume "?g'=1" hence ?thesis using prems by (auto simp add: Let_def simp_num_pair_def simpnum_numbound0)} moreover {assume g'1:"?g'>1" have gpdg: "?g' dvd ?g" by (simp add: igcd_dvd2) have gpdd: "?g' dvd n" by (simp add: igcd_dvd1) have gpdgp: "?g' dvd ?g'" by simp from zdvd_imp_le[OF gpdd np] have g'n: "?g' ≤ n" . from zdiv_mono1[OF g'n g'p, simplified zdiv_self[OF gp0]] have "n div ?g' >0" by simp hence ?thesis using prems by(auto simp add: simp_num_pair_def Let_def reducecoeffh_numbound0 simpnum_numbound0)} ultimately have ?thesis by blast} ultimately have ?thesis by blast} ultimately show ?thesis by blast qed consts simpfm :: "fm => fm" recdef simpfm "measure fmsize" "simpfm (And p q) = conj (simpfm p) (simpfm q)" "simpfm (Or p q) = disj (simpfm p) (simpfm q)" "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 | _ => Lt a')" "simpfm (Le a) = (let a' = simpnum a in case a' of C v => if (v ≤ 0) then T else F | _ => Le a')" "simpfm (Gt a) = (let a' = simpnum a in case a' of C v => if (v > 0) then T else F | _ => Gt a')" "simpfm (Ge a) = (let a' = simpnum a in case a' of C v => if (v ≥ 0) then T else F | _ => Ge a')" "simpfm (Eq a) = (let a' = simpnum a in case a' of C v => if (v = 0) then T else F | _ => Eq a')" "simpfm (NEq a) = (let a' = simpnum a in case a' of C v => if (v ≠ 0) then T else F | _ => NEq a')" "simpfm p = p" lemma simpfm: "Ifm bs (simpfm p) = Ifm bs p" proof(induct p rule: simpfm.induct) case (6 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume "¬ (∃ v. ?sa = C v)" hence ?case using sa by (cases ?sa, simp_all add: Let_def)} ultimately show ?case by blast next case (7 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume "¬ (∃ v. ?sa = C v)" hence ?case using sa by (cases ?sa, simp_all add: Let_def)} ultimately show ?case by blast next case (8 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume "¬ (∃ v. ?sa = C v)" hence ?case using sa by (cases ?sa, simp_all add: Let_def)} ultimately show ?case by blast next case (9 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume "¬ (∃ v. ?sa = C v)" hence ?case using sa by (cases ?sa, simp_all add: Let_def)} ultimately show ?case by blast next case (10 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume "¬ (∃ v. ?sa = C v)" hence ?case using sa by (cases ?sa, simp_all add: Let_def)} ultimately show ?case by blast next case (11 a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp {fix v assume "?sa = C v" hence ?case using sa by simp } moreover {assume "¬ (∃ v. ?sa = C v)" hence ?case using sa by (cases ?sa, simp_all add: Let_def)} ultimately show ?case by blast qed (induct p rule: simpfm.induct, simp_all add: conj disj imp iff not) lemma simpfm_bound0: "bound0 p ==> bound0 (simpfm p)" proof(induct p rule: simpfm.induct) case (6 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def) next case (7 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def) next case (8 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def) next case (9 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def) next case (10 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def) next case (11 a) hence nb: "numbound0 a" by simp hence "numbound0 (simpnum a)" by (simp only: simpnum_numbound0[OF nb]) thus ?case by (cases "simpnum a", auto simp add: Let_def) qed(auto simp add: disj_def imp_def iff_def conj_def not_bn) lemma simpfm_qf: "qfree p ==> qfree (simpfm p)" by (induct p rule: simpfm.induct, auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def) (case_tac "simpnum a",auto)+ consts prep :: "fm => fm" recdef prep "measure fmsize" "prep (E T) = T" "prep (E F) = F" "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))" "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))" "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))" "prep (E (NOT (And p q))) = disj (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))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))" "prep (E p) = E (prep p)" "prep (A (And p q)) = conj (prep (A p)) (prep (A q))" "prep (A p) = prep (NOT (E (NOT p)))" "prep (NOT (NOT p)) = prep p" "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))" "prep (NOT (A p)) = prep (E (NOT p))" "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))" "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))" "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))" "prep (NOT p) = not (prep p)" "prep (Or p q) = disj (prep p) (prep q)" "prep (And p q) = conj (prep p) (prep q)" "prep (Imp p q) = prep (Or (NOT p) q)" "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))" "prep p = p" (hints simp add: fmsize_pos) lemma prep: "!! bs. Ifm bs (prep p) = Ifm bs p" by (induct p rule: prep.induct, auto) (* Generic quantifier elimination *) consts qelim :: "fm => (fm => fm) => fm" recdef qelim "measure fmsize" "qelim (E p) = (λ qe. DJ 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 (Imp p q) = (λ qe. imp (qelim p qe) (qelim q qe))" "qelim (Iff p q) = (λ qe. iff (qelim p qe) (qelim q qe))" "qelim p = (λ y. simpfm p)" 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 qe_inv] by(induct p rule: qelim.induct) (auto simp add: not disj conj iff imp not_qf disj_qf conj_qf imp_qf iff_qf simpfm simpfm_qf simp del: simpfm.simps) consts plusinf:: "fm => fm" (* Virtual substitution of +∞*) minusinf:: "fm => fm" (* Virtual substitution of -∞*) recdef minusinf "measure size" "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" "minusinf (Le (CN 0 c e)) = T" "minusinf (Gt (CN 0 c e)) = F" "minusinf (Ge (CN 0 c e)) = F" "minusinf p = p" recdef plusinf "measure size" "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" "plusinf (Le (CN 0 c e)) = F" "plusinf (Gt (CN 0 c e)) = T" "plusinf (Ge (CN 0 c e)) = T" "plusinf p = p" 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 (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)" "isrlfm (Le (CN 0 c e)) = (c>0 ∧ numbound0 e)" "isrlfm (Gt (CN 0 c e)) = (c>0 ∧ numbound0 e)" "isrlfm (Ge (CN 0 c e)) = (c>0 ∧ numbound0 e)" "isrlfm p = (isatom p ∧ (bound0 p))" (* splits the bounded from the unbounded part*) consts rsplit0 :: "num => int × num" recdef rsplit0 "measure num_size" "rsplit0 (Bound 0) = (1,C 0)" "rsplit0 (Add a b) = (let (ca,ta) = rsplit0 a ; (cb,tb) = rsplit0 b in (ca+cb, Add ta tb))" "rsplit0 (Sub a b) = rsplit0 (Add a (Neg b))" "rsplit0 (Neg a) = (let (c,t) = rsplit0 a in (-c,Neg t))" "rsplit0 (Mul c a) = (let (ca,ta) = rsplit0 a in (c*ca,Mul c ta))" "rsplit0 (CN 0 c a) = (let (ca,ta) = rsplit0 a in (c+ca,ta))" "rsplit0 (CN n c a) = (let (ca,ta) = rsplit0 a in (ca,CN n c ta))" "rsplit0 t = (0,t)" lemma rsplit0: shows "Inum bs ((split (CN 0)) (rsplit0 t)) = Inum bs t ∧ numbound0 (snd (rsplit0 t))" proof (induct t rule: rsplit0.induct) case (2 a b) let ?sa = "rsplit0 a" let ?sb = "rsplit0 b" let ?ca = "fst ?sa" let ?cb = "fst ?sb" let ?ta = "snd ?sa" let ?tb = "snd ?sb" from prems have nb: "numbound0 (snd(rsplit0 (Add a b)))" by(cases "rsplit0 a",auto simp add: Let_def split_def) have "Inum bs ((split (CN 0)) (rsplit0 (Add a b))) = Inum bs ((split (CN 0)) ?sa)+Inum bs ((split (CN 0)) ?sb)" by (simp add: Let_def split_def ring_simps) also have "… = Inum bs a + Inum bs b" using prems by (cases "rsplit0 a", simp_all) finally show ?case using nb by simp qed(auto simp add: Let_def split_def ring_simps , simp add: right_distrib[symmetric]) (* Linearize a formula*) definition lt :: "int => num => fm" where "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 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 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 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 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 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: "numnoabs t ==> Ifm bs (split lt (rsplit0 t)) = Ifm bs (Lt t) ∧ isrlfm (split lt (rsplit0 t))" using rsplit0[where bs = "bs" and t="t"] by (auto simp add: lt_def split_def,cases "snd(rsplit0 t)",auto,case_tac "nat",auto) lemma le: "numnoabs t ==> Ifm bs (split le (rsplit0 t)) = Ifm bs (Le t) ∧ isrlfm (split le (rsplit0 t))" using rsplit0[where bs = "bs" and t="t"] by (auto simp add: le_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) lemma gt: "numnoabs t ==> Ifm bs (split gt (rsplit0 t)) = Ifm bs (Gt t) ∧ isrlfm (split gt (rsplit0 t))" using rsplit0[where bs = "bs" and t="t"] by (auto simp add: gt_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) lemma ge: "numnoabs t ==> Ifm bs (split ge (rsplit0 t)) = Ifm bs (Ge t) ∧ isrlfm (split ge (rsplit0 t))" using rsplit0[where bs = "bs" and t="t"] by (auto simp add: ge_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) lemma eq: "numnoabs t ==> Ifm bs (split eq (rsplit0 t)) = Ifm bs (Eq t) ∧ isrlfm (split eq (rsplit0 t))" using rsplit0[where bs = "bs" and t="t"] by (auto simp add: eq_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) lemma neq: "numnoabs t ==> Ifm bs (split neq (rsplit0 t)) = Ifm bs (NEq t) ∧ isrlfm (split neq (rsplit0 t))" using rsplit0[where bs = "bs" and t="t"] by (auto simp add: neq_def split_def) (cases "snd(rsplit0 t)",auto,case_tac "nat",auto) lemma conj_lin: "isrlfm p ==> isrlfm q ==> isrlfm (conj p q)" by (auto simp add: conj_def) lemma disj_lin: "isrlfm p ==> isrlfm q ==> isrlfm (disj p q)" by (auto simp add: disj_def) consts rlfm :: "fm => fm" recdef rlfm "measure fmsize" "rlfm (And p q) = conj (rlfm p) (rlfm q)" "rlfm (Or p q) = disj (rlfm p) (rlfm q)" "rlfm (Imp p q) = disj (rlfm (NOT p)) (rlfm q)" "rlfm (Iff p q) = disj (conj (rlfm p) (rlfm q)) (conj (rlfm (NOT p)) (rlfm (NOT q)))" "rlfm (Lt a) = split lt (rsplit0 a)" "rlfm (Le a) = split le (rsplit0 a)" "rlfm (Gt a) = split gt (rsplit0 a)" "rlfm (Ge a) = split ge (rsplit0 a)" "rlfm (Eq a) = split eq (rsplit0 a)" "rlfm (NEq a) = split neq (rsplit0 a)" "rlfm (NOT (And p q)) = disj (rlfm (NOT p)) (rlfm (NOT q))" "rlfm (NOT (Or p q)) = conj (rlfm (NOT p)) (rlfm (NOT q))" "rlfm (NOT (Imp p q)) = conj (rlfm p) (rlfm (NOT q))" "rlfm (NOT (Iff p q)) = disj (conj(rlfm p) (rlfm(NOT q))) (conj(rlfm(NOT p)) (rlfm q))" "rlfm (NOT (NOT p)) = rlfm p" "rlfm (NOT T) = F" "rlfm (NOT F) = T" "rlfm (NOT (Lt a)) = rlfm (Ge a)" "rlfm (NOT (Le a)) = rlfm (Gt a)" "rlfm (NOT (Gt a)) = rlfm (Le a)" "rlfm (NOT (Ge a)) = rlfm (Lt a)" "rlfm (NOT (Eq a)) = rlfm (NEq a)" "rlfm (NOT (NEq a)) = rlfm (Eq a)" "rlfm p = p" (hints simp add: fmsize_pos) lemma rlfm_I: assumes qfp: "qfree p" shows "(Ifm bs (rlfm p) = Ifm bs p) ∧ isrlfm (rlfm p)" using qfp by (induct p rule: rlfm.induct, auto simp add: lt le gt ge eq neq conj disj conj_lin disj_lin) (* Operations needed for Ferrante and Rackoff *) lemma rminusinf_inf: assumes lp: "isrlfm p" shows "∃ z. ∀ x < z. Ifm (x#bs) (minusinf p) = Ifm (x#bs) p" (is "∃ z. ∀ x. ?P z x p") using lp proof (induct p rule: minusinf.induct) case (1 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto next case (2 p q) thus ?case by (auto,rule_tac x= "min z za" in exI) auto next case (3 c e) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x < ?z" hence "(real c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) hence "real c * x + ?e < 0" by arith hence "real c * x + ?e ≠ 0" by simp with xz have "?P ?z x (Eq (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 (Eq (CN 0 c e))" by simp thus ?case by blast next case (4 c e) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x < ?z" hence "(real c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) hence "real c * x + ?e < 0" by arith hence "real c * x + ?e ≠ 0" by simp with xz have "?P ?z x (NEq (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 (NEq (CN 0 c e))" by simp thus ?case by blast next case (5 c e) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x < ?z" hence "(real c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) hence "real 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) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x < ?z" hence "(real c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) hence "real 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) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x < ?z" hence "(real c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) hence "real 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) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x < ?z" hence "(real c * x < - ?e)" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="- ?e"] mult_ac) hence "real 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 } hence "∀ x < ?z. ?P ?z x (Ge (CN 0 c e))" by simp thus ?case by blast qed simp_all lemma rplusinf_inf: assumes lp: "isrlfm p" shows "∃ z. ∀ x > z. Ifm (x#bs) (plusinf p) = Ifm (x#bs) p" (is "∃ z. ∀ x. ?P z x p") using lp proof (induct p rule: isrlfm.induct) case (1 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto next case (2 p q) thus ?case by (auto,rule_tac x= "max z za" in exI) auto next case (3 c e) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real c * x > - ?e)" by (simp add: mult_ac) hence "real c * x + ?e > 0" by arith hence "real c * x + ?e ≠ 0" by simp with xz have "?P ?z x (Eq (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 (Eq (CN 0 c e))" by simp thus ?case by blast next case (4 c e) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real c * x > - ?e)" by (simp add: mult_ac) hence "real c * x + ?e > 0" by arith hence "real c * x + ?e ≠ 0" by simp with xz have "?P ?z x (NEq (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 (NEq (CN 0 c e))" by simp thus ?case by blast next case (5 c e) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real c * x > - ?e)" by (simp add: mult_ac) hence "real 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) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real c * x > - ?e)" by (simp add: mult_ac) hence "real 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) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real c * x > - ?e)" by (simp add: mult_ac) hence "real 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) from prems have nb: "numbound0 e" by simp from prems have cp: "real c > 0" by simp fix a let ?e="Inum (a#bs) e" let ?z = "(- ?e) / real c" {fix x assume xz: "x > ?z" with mult_strict_right_mono [OF xz cp] cp have "(real c * x > - ?e)" by (simp add: mult_ac) hence "real 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 } hence "∀ x > ?z. ?P ?z x (Ge (CN 0 c e))" by simp thus ?case by blast qed simp_all lemma rminusinf_bound0: assumes lp: "isrlfm p" shows "bound0 (minusinf p)" using lp by (induct p rule: minusinf.induct) simp_all lemma rplusinf_bound0: assumes lp: "isrlfm p" shows "bound0 (plusinf p)" using lp by (induct p rule: plusinf.induct) simp_all lemma rminusinf_ex: assumes lp: "isrlfm p" and ex: "Ifm (a#bs) (minusinf p)" shows "∃ x. Ifm (x#bs) p" 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"] obtain z where z_def: "∀x<z. Ifm (x # bs) (minusinf p) = Ifm (x # bs) p" by blast from th have "Ifm ((z - 1)#bs) (minusinf p)" by simp moreover have "z - 1 < z" by simp ultimately show ?thesis using z_def by auto qed lemma rplusinf_ex: assumes lp: "isrlfm p" and ex: "Ifm (a#bs) (plusinf p)" shows "∃ x. Ifm (x#bs) p" proof- from bound0_I [OF rplusinf_bound0[OF lp], where b="a" and bs ="bs"] ex have th: "∀ x. Ifm (x#bs) (plusinf p)" by auto 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 uset:: "fm => (num × int) list" usubst :: "fm => (num × int) => fm " recdef uset "measure size" "uset (And p q) = (uset p @ uset q)" "uset (Or p q) = (uset p @ uset q)" "uset (Eq (CN 0 c e)) = [(Neg e,c)]" "uset (NEq (CN 0 c e)) = [(Neg e,c)]" "uset (Lt (CN 0 c e)) = [(Neg e,c)]" "uset (Le (CN 0 c e)) = [(Neg e,c)]" "uset (Gt (CN 0 c e)) = [(Neg e,c)]" "uset (Ge (CN 0 c e)) = [(Neg e,c)]" "uset p = []" recdef usubst "measure size" "usubst (And p q) = (λ (t,n). And (usubst p (t,n)) (usubst q (t,n)))" "usubst (Or p q) = (λ (t,n). Or (usubst p (t,n)) (usubst q (t,n)))" "usubst (Eq (CN 0 c e)) = (λ (t,n). Eq (Add (Mul c t) (Mul n e)))" "usubst (NEq (CN 0 c e)) = (λ (t,n). NEq (Add (Mul c t) (Mul n e)))" "usubst (Lt (CN 0 c e)) = (λ (t,n). Lt (Add (Mul c t) (Mul n e)))" "usubst (Le (CN 0 c e)) = (λ (t,n). Le (Add (Mul c t) (Mul n e)))" "usubst (Gt (CN 0 c e)) = (λ (t,n). Gt (Add (Mul c t) (Mul n e)))" "usubst (Ge (CN 0 c e)) = (λ (t,n). Ge (Add (Mul c t) (Mul n e)))" "usubst p = (λ (t,n). p)" lemma usubst_I: assumes lp: "isrlfm p" and np: "real n > 0" and nbt: "numbound0 t" shows "(Ifm (x#bs) (usubst p (t,n)) = Ifm (((Inum (x#bs) t)/(real n))#bs) p) ∧ bound0 (usubst p (t,n))" (is "(?I x (usubst p (t,n)) = ?I ?u p) ∧ ?B p" is "(_ = ?I (?t/?n) p) ∧ _" is "(_ = ?I (?N x t /_) p) ∧ _") using lp proof(induct p rule: usubst.induct) case (5 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ have "?I ?u (Lt (CN 0 c e)) = (real 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 c *(?t/?n)) + ?n*(?N x e) < 0)" by (simp only: pos_less_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: ring_simps) also have "… = (real c *?t + ?n* (?N x e) < 0)" using np by simp finally show ?case using nbt nb by (simp add: ring_simps) next case (6 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ have "?I ?u (Le (CN 0 c e)) = (real 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 c *(?t/?n)) + ?n*(?N x e) ≤ 0)" by (simp only: pos_le_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: ring_simps) also have "… = (real c *?t + ?n* (?N x e) ≤ 0)" using np by simp finally show ?case using nbt nb by (simp add: ring_simps) next case (7 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ have "?I ?u (Gt (CN 0 c e)) = (real 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 c *(?t/?n)) + ?n*(?N x e) > 0)" by (simp only: pos_divide_less_eq[OF np, where a="real c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: ring_simps) also have "… = (real c *?t + ?n* (?N x e) > 0)" using np by simp finally show ?case using nbt nb by (simp add: ring_simps) next case (8 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ have "?I ?u (Ge (CN 0 c e)) = (real 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 c *(?t/?n)) + ?n*(?N x e) ≥ 0)" by (simp only: pos_divide_le_eq[OF np, where a="real c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: ring_simps) also have "… = (real c *?t + ?n* (?N x e) ≥ 0)" using np by simp finally show ?case using nbt nb by (simp add: ring_simps) next case (3 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ from np have np: "real n ≠ 0" by simp have "?I ?u (Eq (CN 0 c e)) = (real 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 c *(?t/?n)) + ?n*(?N x e) = 0)" by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: ring_simps) also have "… = (real c *?t + ?n* (?N x e) = 0)" using np by simp finally show ?case using nbt nb by (simp add: ring_simps) next case (4 c e) from prems have cp: "c >0" and nb: "numbound0 e" by simp+ from np have np: "real n ≠ 0" by simp have "?I ?u (NEq (CN 0 c e)) = (real 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 c *(?t/?n)) + ?n*(?N x e) ≠ 0)" by (simp only: nonzero_eq_divide_eq[OF np, where a="real c *(?t/?n) + (?N x e)" and b="0", simplified divide_zero_left]) (simp only: ring_simps) also have "… = (real c *?t + ?n* (?N x e) ≠ 0)" using np by simp finally show ?case using nbt nb by (simp add: ring_simps) qed(simp_all add: nbt numbound0_I[where bs ="bs" and b="(Inum (x#bs) t)/ real n" and b'="x"] nth_pos2) lemma uset_l: assumes lp: "isrlfm p" shows "∀ (t,k) ∈ set (uset p). numbound0 t ∧ k >0" using lp by(induct p rule: uset.induct,auto) lemma rminusinf_uset: assumes lp: "isrlfm p" and nmi: "¬ (Ifm (a#bs) (minusinf p))" (is "¬ (Ifm (a#bs) (?M p))") and ex: "Ifm (x#bs) p" (is "?I x p") shows "∃ (s,m) ∈ set (uset p). x ≥ Inum (a#bs) s / real m" (is "∃ (s,m) ∈ ?U p. x ≥ ?N a s / real m") proof- have "∃ (s,m) ∈ set (uset p). real m * x ≥ Inum (a#bs) s " (is "∃ (s,m) ∈ ?U p. real m *x ≥ ?N a s") using lp nmi ex by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) then obtain s m where smU: "(s,m) ∈ set (uset p)" and mx: "real m * x ≥ ?N a s" by blast from uset_l[OF lp] smU have mp: "real 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 m" by (auto simp add: mult_commute) thus ?thesis using smU by auto qed lemma rplusinf_uset: assumes lp: "isrlfm p" and nmi: "¬ (Ifm (a#bs) (plusinf p))" (is "¬ (Ifm (a#bs) (?M p))") and ex: "Ifm (x#bs) p" (is "?I x p") shows "∃ (s,m) ∈ set (uset p). x ≤ Inum (a#bs) s / real m" (is "∃ (s,m) ∈ ?U p. x ≤ ?N a s / real m") proof- have "∃ (s,m) ∈ set (uset p). real m * x ≤ Inum (a#bs) s " (is "∃ (s,m) ∈ ?U p. real m *x ≤ ?N a s") using lp nmi ex by (induct p rule: minusinf.induct, auto simp add:numbound0_I[where bs="bs" and b="a" and b'="x"] nth_pos2) then obtain s m where smU: "(s,m) ∈ set (uset p)" and mx: "real m * x ≤ ?N a s" by blast from uset_l[OF lp] smU have mp: "real 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 m" by (auto simp add: mult_commute) thus ?thesis using smU by auto qed lemma lin_dense: assumes lp: "isrlfm p" and noS: "∀ t. l < t ∧ t< u --> t ∉ (λ (t,n). Inum (x#bs) t / real n) ` set (uset p)" (is "∀ t. _ ∧ _ --> t ∉ (λ (t,n). ?N x t / real 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" shows "Ifm (y#bs) p" using lp px noS proof (induct p rule: isrlfm.induct) case (5 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ from prems have "x * real c + ?N x e < 0" by (simp add: ring_simps) hence pxc: "x < (- ?N x e) / real c" by (simp only: pos_less_divide_eq[OF cp, where a="x" and b="-?N x e"]) from prems have noSc:"∀ t. l < t ∧ t < u --> t ≠ (- ?N x e) / real c" by auto with ly yu have yne: "y ≠ - ?N x e / real c" by auto hence "y < (- ?N x e) / real c ∨ y > (-?N x e) / real c" by auto moreover {assume y: "y < (-?N x e)/ real c" hence "y * real c < - ?N x e" by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real c * y + ?N x e < 0" by (simp add: ring_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 c" with yu have eu: "u > (- ?N x e) / real c" by auto with noSc ly yu have "(- ?N x e) / real c ≤ l" by (cases "(- ?N x e) / real c > l", auto) with lx pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next case (6 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp + from prems have "x * real c + ?N x e ≤ 0" by (simp add: ring_simps) hence pxc: "x ≤ (- ?N x e) / real c" by (simp only: pos_le_divide_eq[OF cp, where a="x" and b="-?N x e"]) from prems have noSc:"∀ t. l < t ∧ t < u --> t ≠ (- ?N x e) / real c" by auto with ly yu have yne: "y ≠ - ?N x e / real c" by auto hence "y < (- ?N x e) / real c ∨ y > (-?N x e) / real c" by auto moreover {assume y: "y < (-?N x e)/ real c" hence "y * real c < - ?N x e" by (simp add: pos_less_divide_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real c * y + ?N x e < 0" by (simp add: ring_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 c" with yu have eu: "u > (- ?N x e) / real c" by auto with noSc ly yu have "(- ?N x e) / real c ≤ l" by (cases "(- ?N x e) / real c > l", auto) with lx pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next case (7 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ from prems have "x * real c + ?N x e > 0" by (simp add: ring_simps) hence pxc: "x > (- ?N x e) / real c" by (simp only: pos_divide_less_eq[OF cp, where a="x" and b="-?N x e"]) from prems have noSc:"∀ t. l < t ∧ t < u --> t ≠ (- ?N x e) / real c" by auto with ly yu have yne: "y ≠ - ?N x e / real c" by auto hence "y < (- ?N x e) / real c ∨ y > (-?N x e) / real c" by auto moreover {assume y: "y > (-?N x e)/ real c" hence "y * real c > - ?N x e" by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real c * y + ?N x e > 0" by (simp add: ring_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 c" with ly have eu: "l < (- ?N x e) / real c" by auto with noSc ly yu have "(- ?N x e) / real c ≥ u" by (cases "(- ?N x e) / real c > l", auto) with xu pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next case (8 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ from prems have "x * real c + ?N x e ≥ 0" by (simp add: ring_simps) hence pxc: "x ≥ (- ?N x e) / real c" by (simp only: pos_divide_le_eq[OF cp, where a="x" and b="-?N x e"]) from prems have noSc:"∀ t. l < t ∧ t < u --> t ≠ (- ?N x e) / real c" by auto with ly yu have yne: "y ≠ - ?N x e / real c" by auto hence "y < (- ?N x e) / real c ∨ y > (-?N x e) / real c" by auto moreover {assume y: "y > (-?N x e)/ real c" hence "y * real c > - ?N x e" by (simp add: pos_divide_less_eq[OF cp, where a="y" and b="-?N x e", symmetric]) hence "real c * y + ?N x e > 0" by (simp add: ring_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 c" with ly have eu: "l < (- ?N x e) / real c" by auto with noSc ly yu have "(- ?N x e) / real c ≥ u" by (cases "(- ?N x e) / real c > l", auto) with xu pxc have "False" by auto hence ?case by simp } ultimately show ?case by blast next case (3 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ from cp have cnz: "real c ≠ 0" by simp from prems have "x * real c + ?N x e = 0" by (simp add: ring_simps) hence pxc: "x = (- ?N x e) / real c" by (simp only: nonzero_eq_divide_eq[OF cnz, where a="x" and b="-?N x e"]) from prems have noSc:"∀ t. l < t ∧ t < u --> t ≠ (- ?N x e) / real c" by auto with lx xu have yne: "x ≠ - ?N x e / real c" by auto with pxc show ?case by simp next case (4 c e) hence cp: "real c > 0" and nb: "numbound0 e" by simp+ from cp have cnz: "real c ≠ 0" by simp from prems have noSc:"∀ t. l < t ∧ t < u --> t ≠ (- ?N x e) / real c" by auto with ly yu have yne: "y ≠ - ?N x e / real c" by auto hence "y* real 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 c + ?N x e ≠ 0" by (simp add: ring_simps) thus ?case using numbound0_I[OF nb, where bs="bs" and b="x" and b'="y"] by (simp add: ring_simps) qed (auto simp add: nth_pos2 numbound0_I[where bs="bs" and b="y" and b'="x"]) lemma finite_set_intervals: assumes px: "P (x::real)" and lx: "l ≤ x" and xu: "x ≤ u" and linS: "l∈ S" and uinS: "u ∈ S" and fS:"finite S" and lS: "∀ x∈ S. l ≤ x" and Su: "∀ x∈ S. x ≤ u" shows "∃ a ∈ S. ∃ b ∈ S. (∀ y. a < y ∧ y < b --> y ∉ S) ∧ a ≤ x ∧ x ≤ b ∧ P x" proof- let ?Mx = "{y. y∈ S ∧ y ≤ x}" let ?xM = "{y. y∈ S ∧ x ≤ y}" let ?a = "Max ?Mx" let ?b = "Min ?xM" have MxS: "?Mx ⊆ S" by blast hence fMx: "finite ?Mx" using fS finite_subset by auto from lx linS have linMx: "l ∈ ?Mx" by blast hence Mxne: "?Mx ≠ {}" by blast have xMS: "?xM ⊆ S" by blast hence fxM: "finite ?xM" using fS finite_subset by auto from xu uinS have linxM: "u ∈ ?xM" by blast hence xMne: "?xM ≠ {}" by blast have ax:"?a ≤ x" using Mxne fMx by auto have xb:"x ≤ ?b" using xMne fxM by auto have "?a ∈ ?Mx" using Max_in[OF fMx Mxne] by simp hence ainS: "?a ∈ S" using MxS by blast have "?b ∈ ?xM" using Min_in[OF fxM xMne] by simp hence binS: "?b ∈ S" using xMS by blast have noy:"∀ y. ?a < y ∧ y < ?b --> y ∉ S" proof(clarsimp) fix y assume ay: "?a < y" and yb: "y < ?b" and yS: "y ∈ S" from yS have "y∈ ?Mx ∨ y∈ ?xM" by auto moreover {assume "y ∈ ?Mx" hence "y ≤ ?a" using Mxne fMx by auto with ay have "False" by simp} moreover {assume "y ∈ ?xM" hence "y ≥ ?b" using xMne fxM by auto with yb have "False" by simp} ultimately show "False" by blast qed from ainS binS noy ax xb px show ?thesis by blast qed lemma finite_set_intervals2: assumes px: "P (x::real)" and lx: "l ≤ x" and xu: "x ≤ u" and linS: "l∈ S" and uinS: "u ∈ S" and fS:"finite S" and lS: "∀ x∈ S. l ≤ x" and Su: "∀ x∈ S. x ≤ u" shows "(∃ s∈ S. P s) ∨ (∃ a ∈ S. ∃ b ∈ S. (∀ y. a < y ∧ y < b --> y ∉ S) ∧ a < x ∧ x < b ∧ P x)" proof- from finite_set_intervals[where P="P", OF px lx xu linS uinS fS lS Su] obtain a and b where as: "a∈ S" and bs: "b∈ S" and noS:"∀y. a < y ∧ y < b --> y ∉ S" and axb: "a ≤ x ∧ x ≤ b ∧ P x" by auto from axb have "x= a ∨ x= b ∨ (a < x ∧ x < b)" by auto thus ?thesis using px as bs noS by blast qed lemma rinf_uset: assumes lp: "isrlfm p" and nmi: "¬ (Ifm (x#bs) (minusinf p))" (is "¬ (Ifm (x#bs) (?M p))") 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 (uset p). ∃ (s,m) ∈ set (uset p). ?I ((Inum (x#bs) l / real n + Inum (x#bs) s / real m) / 2) p" proof- let ?N = "λ x t. Inum (x#bs) t" let ?U = "set (uset p)" from ex obtain a where pa: "?I a p" by blast from bound0_I[OF rminusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] nmi have nmi': "¬ (?I a (?M p))" by simp from bound0_I[OF rplusinf_bound0[OF lp], where bs="bs" and b="x" and b'="a"] npi have npi': "¬ (?I a (?P p))" by simp have "∃ (l,n) ∈ set (uset p). ∃ (s,m) ∈ set (uset p). ?I ((?N a l/real n + ?N a s /real m) / 2) p" proof- let ?M = "(λ (t,c). ?N a t / real c) ` ?U" have fM: "finite ?M" by auto from rminusinf_uset[OF lp nmi pa] rplusinf_uset[OF lp npi pa] have "∃ (l,n) ∈ set (uset p). ∃ (s,m) ∈ set (uset p). a ≤ ?N x l / real n ∧ a ≥ ?N x s / real m" by blast then obtain "t" "n" "s" "m" where tnU: "(t,n) ∈ ?U" and smU: "(s,m) ∈ ?U" and xs1: "a ≤ ?N x s / real m" and tx1: "a ≥ ?N x t / real n" by blast from uset_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 m" and tx: "a ≥ ?N a t / real n" by auto from tnU have Mne: "?M ≠ {}" by auto hence Une: "?U ≠ {}" by simp let ?l = "Min ?M" let ?u = "Max ?M" have linM: "?l ∈ ?M" using fM Mne by simp have uinM: "?u ∈ ?M" using fM Mne by simp have tnM: "?N a t / real n ∈ ?M" using tnU by auto have smM: "?N a s / real 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 n" using tnM Mne by simp hence lx: "?l ≤ a" using tx by simp have "?N a s / real 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) ∨ (∃ 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 nu" by auto then obtain "tu" "nu" where tuU: "(tu,nu) ∈ ?U" and tuu:"u= ?N a tu / real nu" by blast have "(u + u) / 2 = u" by auto with pu tuu have "?I (((?N a tu / real nu) + (?N a tu / real 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" 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 t1n" by auto then obtain "t1u" "t1n" where t1uU: "(t1u,t1n) ∈ ?U" and t1u: "t1 = ?N a t1u / real t1n" by blast from t2M have "∃ (t2u,t2n) ∈ ?U. t2 = ?N a t2u / real t2n" by auto then obtain "t2u" "t2n" where t2uU: "(t2u,t2n) ∈ ?U" and t2u: "t2 = ?N a t2u / real t2n" by blast from t1x xt2 have t1t2: "t1 < t2" by simp let ?u = "(t1 + t2) / 2" from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" . 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" and pu: "?I ((?N a l / real n + ?N a s / real m) / 2) p" by blast from lnU smU uset_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"] numbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu have "?I ((?N x l / real n + ?N x s / real m) / 2) p" by simp with lnU smU show ?thesis by auto qed (* The Ferrante - Rackoff Theorem *) 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 (uset p). ∃ (s,m) ∈ set (uset p). Ifm ((((Inum (x#bs) t)/ real n + (Inum (x#bs) s) / real m) /2)#bs) p))" (is "(∃ x. ?I x p) = (?M ∨ ?P ∨ ?F)" is "?E = ?D") proof assume px: "∃ x. ?I x p" have "?M ∨ ?P ∨ (¬ ?M ∧ ¬ ?P)" by blast moreover {assume "?M ∨ ?P" hence "?D" by blast} moreover {assume nmi: "¬ ?M" and npi: "¬ ?P" from rinf_uset[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast} ultimately show "?D" by blast next 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} ultimately show "?E" by blast qed lemma fr_equsubst: assumes lp: "isrlfm p" shows "(∃ x. Ifm (x#bs) p) = ((Ifm (x#bs) (minusinf p)) ∨ (Ifm (x#bs) (plusinf p)) ∨ (∃ (t,k) ∈ set (uset p). ∃ (s,l) ∈ set (uset p). Ifm (x#bs) (usubst p (Add(Mul l t) (Mul k s) , 2*k*l))))" (is "(∃ x. ?I x p) = (?M ∨ ?P ∨ ?F)" is "?E = ?D") proof assume px: "∃ x. ?I x p" have "?M ∨ ?P ∨ (¬ ?M ∧ ¬ ?P)" by blast moreover {assume "?M ∨ ?P" hence "?D" by blast} moreover {assume nmi: "¬ ?M" and npi: "¬ ?P" let ?f ="λ (t,n). Inum (x#bs) t / real n" let ?N = "λ t. Inum (x#bs) t" {fix t n s m assume "(t,n)∈ set (uset p)" and "(s,m) ∈ set (uset p)" with uset_l[OF lp] have tnb: "numbound0 t" and np:"real n > 0" and snb: "numbound0 s" and mp:"real m > 0" by auto let ?st = "Add (Mul m t) (Mul n s)" from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" by (simp add: mult_commute) from tnb snb have st_nb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mnp mp np by (simp add: ring_simps add_divide_distrib) from usubst_I[OF lp mnp st_nb, where x="x" and bs="bs"] have "?I x (usubst p (?st,2*n*m)) = ?I ((?N t / real n + ?N s / real m) /2) p" by (simp only: st[symmetric])} with rinf_uset[OF lp nmi npi px] have "?F" by blast hence "?D" by blast} ultimately show "?D" by blast next 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 (uset p)" and "(s,l) ∈ set (uset p)" and px:"?I x (usubst p (Add (Mul l t) (Mul k s), 2*k*l))" with uset_l[OF lp] have tnb: "numbound0 t" and np:"real k > 0" and snb: "numbound0 s" and mp:"real l > 0" by auto let ?st = "Add (Mul l t) (Mul k s)" from mult_pos_pos[OF np mp] have mnp: "real (2*k*l) > 0" by (simp add: mult_commute) from tnb snb have st_nb: "numbound0 ?st" by simp from usubst_I[OF lp mnp st_nb, where bs="bs"] px have "?E" by auto} ultimately show "?E" by blast qed (* Implement the right hand side of Ferrante and Rackoff's Theorem. *) constdefs ferrack:: "fm => fm" "ferrack p ≡ (let p' = rlfm (simpfm p); mp = minusinf p'; pp = plusinf p' in if (mp = T ∨ pp = T) then T else (let U = remdps(map simp_num_pair (map (λ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs (uset p')))) in decr (disj mp (disj pp (evaldjf (simpfm o (usubst p')) U)))))" lemma uset_cong_aux: assumes Ul: "∀ (t,n) ∈ set U. numbound0 t ∧ n >0" shows "((λ (t,n). Inum (x#bs) t /real n) ` (set (map (λ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) (alluopairs U)))) = ((λ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (set U × set U))" (is "?lhs = ?rhs") proof(auto) fix t n s m assume "((t,n),(s,m)) ∈ set (alluopairs U)" hence th: "((t,n),(s,m)) ∈ (set U × set U)" using alluopairs_set1[where xs="U"] by blast 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 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mnz nnz by (simp add: ring_simps add_divide_distrib) thus "(real m * Inum (x # bs) t + real n * Inum (x # bs) s) / (2 * real n * real m) ∈ (λ((t, n), s, m). (Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) ` (set U × set U)"using mnz nnz th apply (auto simp add: th add_divide_distrib ring_simps split_def image_def) by (rule_tac x="(s,m)" in bexI,simp_all) (rule_tac x="(t,n)" in bexI,simp_all) next fix t n s m 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 have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mnz nnz by (simp add: ring_simps add_divide_distrib) let ?P = "λ (t',n') (s',m'). (Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" have Pc:"∀ a b. ?P a b = ?P b a" by auto from Ul alluopairs_set1 have Up:"∀ ((t,n),(s,m)) ∈ set (alluopairs U). n ≠ 0 ∧ m ≠ 0" by blast 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)" 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 n' + ?N s' / real m')/2 = ?N ?st' / real (2*n'*m')" using mnz' nnz' by (simp add: ring_simps add_divide_distrib) from Pts' have "(Inum (x # bs) t / real n + Inum (x # bs) s / real m)/2 = (Inum (x # bs) t' / real n' + Inum (x # bs) s' / real m')/2" by simp also have "… = ((λ(t, n). Inum (x # bs) t / real 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 n + Inum (x # bs) s / real m) / 2 ∈ (λ(t, n). Inum (x # bs) t / real n) ` (λ((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m)) ` set (alluopairs U)" using ts'_U by blast qed lemma uset_cong: assumes lp: "isrlfm p" and UU': "((λ (t,n). Inum (x#bs) t /real n) ` U') = ((λ ((t,n),(s,m)). (Inum (x#bs) t /real n + Inum (x#bs) s /real m)/2) ` (U × U))" (is "?f ` U' = ?g ` (U×U)") and U: "∀ (t,n) ∈ U. numbound0 t ∧ n > 0" and U': "∀ (t,n) ∈ U'. numbound0 t ∧ n > 0" shows "(∃ (t,n) ∈ U. ∃ (s,m) ∈ U. Ifm (x#bs) (usubst p (Add (Mul m t) (Mul n s),2*n*m))) = (∃ (t,n) ∈ U'. Ifm (x#bs) (usubst p (t,n)))" (is "?lhs = ?rhs") proof assume ?lhs then obtain t n s m where tnU: "(t,n) ∈ U" and smU:"(s,m) ∈ U" and Pst: "Ifm (x#bs) (usubst 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" and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mp np by (simp add: ring_simps add_divide_distrib) from tnU smU UU' have "?g ((t,n),(s,m)) ∈ ?f ` U'" by blast hence "∃ (t',n') ∈ U'. ?g ((t,n),(s,m)) = ?f (t',n')" 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 n' > 0" by auto from usubst_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 (2 * n * m) # bs) p" by simp from conjunct1[OF usubst_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) (usubst 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) (usubst 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))" 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 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" and snb: "numbound0 s" and mp:"m > 0" by auto let ?st= "Add (Mul m t) (Mul n s)" from mult_pos_pos[OF np mp] have mnp: "real (2*n*m) > 0" by (simp add: mult_commute real_of_int_mult[symmetric] del: real_of_int_mult) from tnb snb have stnb: "numbound0 ?st" by simp have st: "(?N t / real n + ?N s / real m)/2 = ?N ?st / real (2*n*m)" using mp np by (simp add: ring_simps add_divide_distrib) from U' tnU' have tnb': "numbound0 t'" and np': "real n' > 0" by auto from usubst_I[OF lp np' tnb', where bs="bs" and x="x",simplified th[simplified split_def fst_conv snd_conv] st] Pt' have Pst2: "Ifm (Inum (x # bs) (Add (Mul m t) (Mul n s)) / real (2 * n * m) # bs) p" by simp with usubst_I[OF lp mnp stnb, where x="x" and bs="bs"] tnU smU show ?lhs by blast qed lemma ferrack: assumes qf: "qfree p" shows "qfree (ferrack p) ∧ ((Ifm bs (ferrack p)) = (∃ x. Ifm (x#bs) p))" (is "_ ∧ (?rhs = ?lhs)") proof- let ?I = "λ x p. Ifm (x#bs) p" fix x let ?N = "λ t. Inum (x#bs) t" let ?q = "rlfm (simpfm p)" let ?U = "uset ?q" let ?Up = "alluopairs ?U" let ?g = "λ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)" let ?S = "map ?g ?Up" let ?SS = "map simp_num_pair ?S" let ?Y = "remdps ?SS" let ?f= "(λ (t,n). ?N t / real n)" let ?h = "λ ((t,n),(s,m)). (?N t/real n + ?N s/ real m) /2" let ?F = "λ p. ∃ a ∈ set (uset p). ∃ b ∈ set (uset p). ?I x (usubst p (?g(a,b)))" let ?ep = "evaldjf (simpfm o (usubst ?q)) ?Y" from rlfm_I[OF simpfm_qf[OF qf]] have lq: "isrlfm ?q" by blast from alluopairs_set1[where xs="?U"] have UpU: "set ?Up ≤ (set ?U × set ?U)" by simp from uset_l[OF lq] have U_l: "∀ (t,n) ∈ set ?U. numbound0 t ∧ n > 0" . 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 simp add: mult_pos_pos) have Y_l: "∀ (t,n) ∈ set ?Y. numbound0 t ∧ n > 0" proof- { 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) (rule_tac x="((aa,ba),(ab,bb))" in bexI, simp_all) then obtain t' n' where tn'S: "(t',n') ∈ set ?S" and tns: "simp_num_pair (t',n') = (t,n)" by blast from tn'S Snb have tnb: "numbound0 t'" and np: "n' > 0" by auto from simp_num_pair_l[OF tnb np tns] have "numbound0 t ∧ n > 0" . } thus ?thesis by blast qed have YU: "(?f ` set ?Y) = (?h ` (set ?U × set ?U))" proof- 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_compose) also have "… = (?f ` set ?S)" by (simp add: th) also have "… = ((?f o ?g) ` set ?Up)" by (simp only: set_map o_def image_compose[symmetric]) also have "… = (?h ` (set ?U × set ?U))" using uset_cong_aux[OF U_l, where x="x" and bs="bs", simplified set_map image_compose[symmetric]] by blast finally show ?thesis . qed have "∀ (t,n) ∈ set ?Y. bound0 (simpfm (usubst ?q (t,n)))" proof- { fix t n assume tnY: "(t,n) ∈ set ?Y" with Y_l have tnb: "numbound0 t" and np: "real n > 0" by auto from usubst_I[OF lq np tnb] have "bound0 (usubst ?q (t,n))" by simp hence "bound0 (simpfm (usubst ?q (t,n)))" using simpfm_bound0 by simp} thus ?thesis by blast qed hence ep_nb: "bound0 ?ep" using evaldjf_bound0[where xs="?Y" and f="simpfm o (usubst ?q)"] by auto let ?mp = "minusinf ?q" let ?pp = "plusinf ?q" let ?M = "?I x ?mp" let ?P = "?I x ?pp" let ?res = "disj ?mp (disj ?pp ?ep)" from rminusinf_bound0[OF lq] rplusinf_bound0[OF lq] ep_nb have nbth: "bound0 ?res" by auto from conjunct1[OF rlfm_I[OF simpfm_qf[OF qf]]] simpfm have th: "?lhs = (∃ x. ?I x ?q)" by auto from th fr_equsubst[OF lq, where bs="bs" and x="x"] have lhfr: "?lhs = (?M ∨ ?P ∨ ?F ?q)" by (simp only: split_def fst_conv snd_conv) also have "… = (?M ∨ ?P ∨ (∃ (t,n) ∈ set ?Y. ?I x (simpfm (usubst ?q (t,n)))))" using uset_cong[OF lq YU U_l Y_l] by (simp only: split_def fst_conv snd_conv simpfm) also have "… = (Ifm (x#bs) ?res)" using evaldjf_ex[where ps="?Y" and bs = "x#bs" and f="simpfm o (usubst ?q)",symmetric] by (simp add: split_def pair_collapse) finally have lheq: "?lhs = (Ifm bs (decr ?res))" using decr[OF nbth] by blast hence lr: "?lhs = ?rhs" apply (unfold ferrack_def Let_def) by (cases "?mp = T ∨ ?pp = T", auto) (simp add: disj_def)+ from decr_qf[OF nbth] have "qfree (ferrack p)" by (auto simp add: Let_def ferrack_def) with lr show ?thesis by blast qed constdefs linrqe:: "fm => fm" "linrqe ≡ (λ p. qelim (prep p) ferrack)" theorem linrqe: "(Ifm bs (linrqe p) = Ifm bs p) ∧ qfree (linrqe p)" using ferrack qelim_ci prep unfolding linrqe_def by auto definition ferrack_test :: "unit => fm" where "ferrack_test u = linrqe (A (A (Imp (Lt (Sub (Bound 1) (Bound 0))) (E (Eq (Sub (Add (Bound 0) (Bound 2)) (Bound 1)))))))" export_code linrqe ferrack_test in SML module_name Ferrack (*code_module Ferrack contains linrqe = linrqe test = ferrack_test*) ML {* Ferrack.ferrack_test () *} use "linreif.ML" oracle linr_oracle ("term") = ReflectedFerrack.linrqe_oracle use "linrtac.ML" setup LinrTac.setup end
lemma alluopairs_set1:
set (alluopairs xs) ⊆ {(x, y). x ∈ set xs ∧ y ∈ set xs}
lemma alluopairs_set:
[| x ∈ set xs; y ∈ set xs |]
==> (x, y) ∈ set (alluopairs xs) ∨ (y, x) ∈ set (alluopairs xs)
lemma alluopairs_ex:
∀x y. P x y = P y x
==> (∃x∈set xs. ∃y∈set xs. P x y) = (∃(x, y)∈set (alluopairs xs). P x y)
lemma nth_pos2:
0 < n ==> (x # xs) ! n = xs ! (n - 1)
lemma filter_length:
length (filter P xs) < Suc (length xs)
lemma remdps_set:
set (remdps xs) = set xs
lemma fmsize_pos:
0 < fmsize p
lemma IfmLeSub:
[| Inum bs s = s'; Inum bs t = t' |] ==> Ifm bs (Le (Sub s t)) = (s' ≤ t')
lemma IfmLtSub:
[| Inum bs s = s'; Inum bs t = t' |] ==> Ifm bs (Lt (Sub s t)) = (s' < t')
lemma IfmEqSub:
[| Inum bs s = s'; Inum bs t = t' |] ==> Ifm bs (Eq (Sub s t)) = (s' = t')
lemma IfmNOT:
Ifm bs p = P ==> Ifm bs (NOT p) = (¬ P)
lemma IfmAnd:
[| Ifm bs p = P; Ifm bs q = Q |] ==> Ifm bs (And p q) = (P ∧ Q)
lemma IfmOr:
[| Ifm bs p = P; Ifm bs q = Q |] ==> Ifm bs (Or p q) = (P ∨ Q)
lemma IfmImp:
[| Ifm bs p = P; Ifm bs q = Q |] ==> Ifm bs (Imp p q) = (P --> Q)
lemma IfmIff:
[| Ifm bs p = P; Ifm bs q = Q |] ==> Ifm bs (Iff p q) = (P = Q)
lemma IfmE:
(!!x. Ifm (x # bs) p = P x) ==> Ifm bs (E p) = (∃x. P x)
lemma IfmA:
(!!x. Ifm (x # bs) p = P x) ==> Ifm bs (A p) = (∀x. P x)
lemma not:
Ifm bs (not p) = Ifm bs (NOT p)
lemma conj:
Ifm bs (conj p q) = Ifm bs (And p q)
lemma disj:
Ifm bs (disj p q) = Ifm bs (Or p q)
lemma imp:
Ifm bs (imp p q) = Ifm bs (Imp p q)
lemma iff:
Ifm bs (ReflectedFerrack.iff p q) = Ifm bs (Iff p q)
lemma conj_simps:
conj F Q = F
conj P F = F
conj T Q = Q
conj P T = P
conj P P = P
[| P ≠ T; P ≠ F; Q ≠ T; Q ≠ F; P ≠ Q |] ==> conj P Q = And P Q
lemma disj_simps:
disj T Q = T
disj P T = T
disj F Q = Q
disj P F = P
disj P P = P
[| P ≠ T; P ≠ F; Q ≠ T; Q ≠ F; P ≠ Q |] ==> disj P Q = Or P Q
lemma imp_simps:
imp F Q = T
imp P T = T
imp T Q = Q
imp P F = not P
imp P P = T
[| P ≠ T; P ≠ F; P ≠ Q; Q ≠ T; Q ≠ F |] ==> imp P Q = Imp P Q
lemma trivNOT:
p ≠ NOT p
NOT p ≠ p
lemma iff_simps:
ReflectedFerrack.iff p p = T
ReflectedFerrack.iff p (NOT p) = F
ReflectedFerrack.iff (NOT p) p = F
ReflectedFerrack.iff p F = not p
ReflectedFerrack.iff F p = not p
p ≠ NOT T ==> ReflectedFerrack.iff T p = p
p ≠ NOT T ==> ReflectedFerrack.iff p T = p
[| p ≠ q; p ≠ NOT q; q ≠ NOT p; p ≠ F; q ≠ F; p ≠ T; q ≠ T |]
==> ReflectedFerrack.iff p q = Iff p q
lemma numbound0_I:
numbound0 a ==> Inum (b # bs) a = Inum (b' # bs) a
lemma bound0_I:
bound0 p ==> Ifm (b # bs) p = Ifm (b' # bs) p
lemma not_qf:
qfree p ==> qfree (not p)
lemma not_bn:
bound0 p ==> bound0 (not p)
lemma conj_qf:
[| qfree p; qfree q |] ==> qfree (conj p q)
lemma conj_nb:
[| bound0 p; bound0 q |] ==> bound0 (conj p q)
lemma disj_qf:
[| qfree p; qfree q |] ==> qfree (disj p q)
lemma disj_nb:
[| bound0 p; bound0 q |] ==> bound0 (disj p q)
lemma imp_qf:
[| qfree p; qfree q |] ==> qfree (imp p q)
lemma imp_nb:
[| bound0 p; bound0 q |] ==> bound0 (imp p q)
lemma iff_qf:
[| qfree p; qfree q |] ==> qfree (ReflectedFerrack.iff p q)
lemma iff_nb:
[| bound0 p; bound0 q |] ==> bound0 (ReflectedFerrack.iff p q)
lemma decrnum:
numbound0 t ==> Inum (x # bs) t = Inum bs (decrnum t)
lemma decr:
bound0 p ==> Ifm (x # bs) p = Ifm bs (decr p)
lemma decr_qf:
bound0 p ==> qfree (decr p)
lemma bound0_qf:
bound0 p ==> qfree p
lemma djf_Or:
Ifm bs (djf f p q) = Ifm bs (Or (f p) q)
lemma djf_simps:
djf f p T = T
djf f p F = f p
[| q ≠ T; q ≠ F |]
==> djf f p q =
Let (f p)
(fm_case T q (λnum. Or (f p) q) (λnum. Or (f p) q) (λnum. Or (f p) q)
(λnum. Or (f p) q) (λnum. Or (f p) q) (λnum. Or (f p) q)
(λfm. Or (f p) q) (λfm1 fm2. Or (f p) q) (λfm1 fm2. Or (f p) q)
(λfm1 fm2. Or (f p) q) (λfm1 fm2. Or (f p) q) (λfm. Or (f p) q)
(λfm. Or (f p) q))
lemma evaldjf_ex:
Ifm bs (evaldjf f ps) = (∃p∈set ps. Ifm bs (f p))
lemma evaldjf_bound0:
∀x∈set xs. bound0 (f x) ==> bound0 (evaldjf f xs)
lemma evaldjf_qf:
∀x∈set xs. qfree (f x) ==> qfree (evaldjf f xs)
lemma disjuncts:
(∃q∈set (disjuncts p). Ifm bs q) = Ifm bs p
lemma disjuncts_nb:
bound0 p ==> ∀q∈set (disjuncts p). bound0 q
lemma disjuncts_qf:
qfree p ==> ∀q∈set (disjuncts p). qfree q
lemma DJ:
[| ∀p q. Ifm bs (f (Or p q)) = Ifm bs (Or (f p) (f q)); f F = F |]
==> Ifm bs (DJ f p) = Ifm bs (f p)
lemma DJ_qf:
∀p. qfree p --> qfree (f p) ==> ∀p. qfree p --> qfree (DJ f p)
lemma DJ_qe:
∀bs p. qfree p --> qfree (qe p) ∧ Ifm bs (qe p) = Ifm bs (E p)
==> ∀bs p. qfree p --> qfree (DJ qe p) ∧ Ifm bs (DJ qe p) = Ifm bs (E p)
lemma maxcoeff_pos:
0 ≤ maxcoeff t
lemma dvdnumcoeff_trans:
[| g dvd g'; dvdnumcoeff t g' |] ==> dvdnumcoeff t g
lemma natabs0:
(nat ¦x¦ = 0) = (x = 0)
lemma numgcd0:
numgcd t = 0 ==> Inum bs t = 0
lemma numgcdh_pos:
0 ≤ g ==> 0 ≤ numgcdh t g
lemma numgcd_pos:
0 ≤ numgcd t
lemma reducecoeffh:
[| dvdnumcoeff t g; 0 < g |] ==> real g * Inum bs (reducecoeffh t g) = Inum bs t
lemma ismaxcoeff_mono:
[| ismaxcoeff t c; c ≤ c' |] ==> ismaxcoeff t c'
lemma maxcoeff_ismaxcoeff:
ismaxcoeff t (maxcoeff t)
lemma igcd_gt1:
1 < igcd i j ==> 1 < ¦i¦ ∧ 1 < ¦j¦ ∨ ¦i¦ = 0 ∧ 1 < ¦j¦ ∨ 1 < ¦i¦ ∧ ¦j¦ = 0
lemma numgcdh0:
numgcdh t m = 0 ==> m = 0
lemma dvdnumcoeff_aux:
[| ismaxcoeff t m; 0 ≤ m; 1 < numgcdh t m |] ==> dvdnumcoeff t (numgcdh t m)
lemma dvdnumcoeff_aux2:
1 < numgcd t ==> dvdnumcoeff t (numgcd t) ∧ 0 < numgcd t
lemma reducecoeff:
real (numgcd t) * Inum bs (reducecoeff t) = Inum bs t
lemma reducecoeffh_numbound0:
numbound0 t ==> numbound0 (reducecoeffh t g)
lemma reducecoeff_numbound0:
numbound0 t ==> numbound0 (reducecoeff t)
lemma numadd:
Inum bs (numadd (t, s)) = Inum bs (Add t s)
lemma numadd_nb:
[| numbound0 t; numbound0 s |] ==> numbound0 (numadd (t, s))
lemma nummul:
Inum bs (nummul t i) = Inum bs (Mul i t)
lemma nummul_nb:
numbound0 t ==> numbound0 (nummul t i)
lemma numneg:
Inum bs (numneg t) = Inum bs (Neg t)
lemma numneg_nb:
numbound0 t ==> numbound0 (numneg t)
lemma numsub:
Inum bs (numsub a b) = Inum bs (Sub a b)
lemma numsub_nb:
[| numbound0 t; numbound0 s |] ==> numbound0 (numsub t s)
lemma simpnum_ci:
Inum bs (simpnum t) = Inum bs t
lemma simpnum_numbound0:
numbound0 t ==> numbound0 (simpnum t)
lemma numadd_nz:
[| nozerocoeff a; nozerocoeff b |] ==> nozerocoeff (numadd (a, b))
lemma nummul_nz:
[| i ≠ 0; nozerocoeff a |] ==> nozerocoeff (nummul a i)
lemma numneg_nz:
nozerocoeff a ==> nozerocoeff (numneg a)
lemma numsub_nz:
[| nozerocoeff a; nozerocoeff b |] ==> nozerocoeff (numsub a b)
lemma simpnum_nz:
nozerocoeff (simpnum t)
lemma maxcoeff_nz:
[| nozerocoeff t; maxcoeff t = 0 |] ==> t = C 0
lemma numgcd_nz:
[| nozerocoeff t; numgcd t = 0 |] ==> t = C 0
lemma simp_num_pair_ci:
(λ(t, n). Inum bs t / real n) (simp_num_pair (t, n)) =
(λ(t, n). Inum bs t / real n) (t, n)
lemma simp_num_pair_l:
[| numbound0 t; 0 < n; simp_num_pair (t, n) = (t', n') |]
==> numbound0 t' ∧ 0 < n'
lemma simpfm:
Ifm bs (simpfm p) = Ifm bs p
lemma simpfm_bound0:
bound0 p ==> bound0 (simpfm p)
lemma simpfm_qf:
qfree p ==> qfree (simpfm p)
lemma prep:
Ifm bs (prep p) = Ifm bs p
lemma qelim_ci:
∀bs p. qfree p --> qfree (qe p) ∧ Ifm bs (qe p) = Ifm bs (E p)
==> qfree (qelim p qe) ∧ Ifm bs (qelim p qe) = Ifm bs p
lemma rsplit0:
Inum bs (split (CN 0) (rsplit0 t)) = Inum bs t ∧ numbound0 (snd (rsplit0 t))
lemma lt:
numnoabs t
==> Ifm bs (split lt (rsplit0 t)) = Ifm bs (Lt t) ∧
isrlfm (split lt (rsplit0 t))
lemma le:
numnoabs t
==> Ifm bs (split le (rsplit0 t)) = Ifm bs (Le t) ∧
isrlfm (split le (rsplit0 t))
lemma gt:
numnoabs t
==> Ifm bs (split gt (rsplit0 t)) = Ifm bs (Gt t) ∧
isrlfm (split gt (rsplit0 t))
lemma ge:
numnoabs t
==> Ifm bs (split ge (rsplit0 t)) = Ifm bs (Ge t) ∧
isrlfm (split ge (rsplit0 t))
lemma eq:
numnoabs t
==> Ifm bs (split eq (rsplit0 t)) = Ifm bs (Eq t) ∧
isrlfm (split eq (rsplit0 t))
lemma neq:
numnoabs t
==> Ifm bs (split neq (rsplit0 t)) = Ifm bs (NEq t) ∧
isrlfm (split neq (rsplit0 t))
lemma conj_lin:
[| isrlfm p; isrlfm q |] ==> isrlfm (conj p q)
lemma disj_lin:
[| isrlfm p; isrlfm q |] ==> isrlfm (disj p q)
lemma rlfm_I:
qfree p ==> Ifm bs (rlfm p) = Ifm bs p ∧ isrlfm (rlfm p)
lemma rminusinf_inf:
isrlfm p ==> ∃z. ∀x<z. Ifm (x # bs) (minusinf p) = Ifm (x # bs) p
lemma rplusinf_inf:
isrlfm p ==> ∃z. ∀x>z. Ifm (x # bs) (plusinf p) = Ifm (x # bs) p
lemma rminusinf_bound0:
isrlfm p ==> bound0 (minusinf p)
lemma rplusinf_bound0:
isrlfm p ==> bound0 (plusinf p)
lemma rminusinf_ex:
[| isrlfm p; Ifm (a # bs) (minusinf p) |] ==> ∃x. Ifm (x # bs) p
lemma rplusinf_ex:
[| isrlfm p; Ifm (a # bs) (plusinf p) |] ==> ∃x. Ifm (x # bs) p
lemma usubst_I:
[| isrlfm p; 0 < real n; numbound0 t |]
==> Ifm (x # bs) (usubst p (t, n)) = Ifm (Inum (x # bs) t / real n # bs) p ∧
bound0 (usubst p (t, n))
lemma uset_l:
isrlfm p ==> ∀(t, k)∈set (uset p). numbound0 t ∧ 0 < k
lemma rminusinf_uset:
[| isrlfm p; ¬ Ifm (a # bs) (minusinf p); Ifm (x # bs) p |]
==> ∃(s, m)∈set (uset p). Inum (a # bs) s / real m ≤ x
lemma rplusinf_uset:
[| isrlfm p; ¬ Ifm (a # bs) (plusinf p); Ifm (x # bs) p |]
==> ∃(s, m)∈set (uset p). x ≤ Inum (a # bs) s / real m
lemma lin_dense:
[| isrlfm p;
∀t. l < t ∧ t < u --> t ∉ (λ(t, n). Inum (x # bs) t / real n) ` set (uset p);
l < x; x < u; Ifm (x # bs) p; l < y; y < u |]
==> Ifm (y # bs) p
lemma finite_set_intervals:
[| P x; l ≤ x; x ≤ u; l ∈ S; u ∈ S; finite S; ∀x∈S. l ≤ x; ∀x∈S. x ≤ u |]
==> ∃a∈S. ∃b∈S. (∀y. a < y ∧ y < b --> y ∉ S) ∧ a ≤ x ∧ x ≤ b ∧ P x
lemma finite_set_intervals2:
[| P x; l ≤ x; x ≤ u; l ∈ S; u ∈ S; finite S; ∀x∈S. l ≤ x; ∀x∈S. x ≤ u |]
==> (∃s∈S. P s) ∨
(∃a∈S. ∃b∈S. (∀y. a < y ∧ y < b --> y ∉ S) ∧ a < x ∧ x < b ∧ P x)
lemma rinf_uset:
[| isrlfm p; ¬ Ifm (x # bs) (minusinf p); ¬ Ifm (x # bs) (plusinf p);
∃x. Ifm (x # bs) p |]
==> ∃(l, n)∈set (uset p).
∃(s, m)∈set (uset p).
Ifm ((Inum (x # bs) l / real n + Inum (x # bs) s / real m) / 2 # bs) p
theorem fr_eq:
isrlfm p
==> (∃x. Ifm (x # bs) p) =
(Ifm (x # bs) (minusinf p) ∨
Ifm (x # bs) (plusinf p) ∨
(∃(t, n)∈set (uset p).
∃(s, m)∈set (uset p).
Ifm ((Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2 # bs)
p))
lemma fr_equsubst:
isrlfm p
==> (∃x. Ifm (x # bs) p) =
(Ifm (x # bs) (minusinf p) ∨
Ifm (x # bs) (plusinf p) ∨
(∃(t, k)∈set (uset p).
∃(s, l)∈set (uset p).
Ifm (x # bs) (usubst p (Add (Mul l t) (Mul k s), 2 * k * l))))
lemma uset_cong_aux:
∀(t, n)∈set U. numbound0 t ∧ 0 < n
==> (λ(t, n). Inum (x # bs) t / real n) `
set (map (λ((t, n), s, m). (Add (Mul m t) (Mul n s), 2 * n * m))
(alluopairs U)) =
(λ((t, n), s, m).
(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) `
(set U × set U)
lemma uset_cong:
[| isrlfm p;
(λ(t, n). Inum (x # bs) t / real n) ` U' =
(λ((t, n), s, m).
(Inum (x # bs) t / real n + Inum (x # bs) s / real m) / 2) `
(U × U);
∀(t, n)∈U. numbound0 t ∧ 0 < n; ∀(t, n)∈U'. numbound0 t ∧ 0 < n |]
==> (∃(t, n)∈U.
∃(s, m)∈U.
Ifm (x # bs) (usubst p (Add (Mul m t) (Mul n s), 2 * n * m))) =
(∃(t, n)∈U'. Ifm (x # bs) (usubst p (t, n)))
lemma ferrack:
qfree p ==> qfree (ferrack p) ∧ Ifm bs (ferrack p) = (∃x. Ifm (x # bs) p)
theorem linrqe:
Ifm bs (linrqe p) = Ifm bs p ∧ qfree (linrqe p)