Up to index of Isabelle/HOL/ex
theory Reflected_Presburgertheory Reflected_Presburger imports GCD Efficient_Nat uses ("coopereif.ML") ("coopertac.ML") begin function iupt :: "int => int => int list" where "iupt i j = (if j < i then [] else i # iupt (i+1) j)" by pat_completeness auto termination by (relation "measure (λ (i, j). nat (j-i+1))") auto lemma iupt_set: "set (iupt i j) = {i..j}" by (induct rule: iupt.induct) (simp add: simp_from_to) (* Periodicity of dvd *) (*********************************************************************************) (**** 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 (CN n c a) = 4 + num_size a" "num_size (Mul c a) = 1 + num_size a" consts Inum :: "int list => num => int" primrec "Inum bs (C c) = c" "Inum bs (Bound n) = bs!n" "Inum bs (CN n c a) = 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) = c* Inum bs a" 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 | Closed nat | NClosed nat (* 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 (Dvd i t) = 2" "fmsize (NDvd i t) = 2" "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 ::"bool list => int list => fm => bool" primrec "Ifm bbs bs T = True" "Ifm bbs bs F = False" "Ifm bbs bs (Lt a) = (Inum bs a < 0)" "Ifm bbs bs (Gt a) = (Inum bs a > 0)" "Ifm bbs bs (Le a) = (Inum bs a ≤ 0)" "Ifm bbs bs (Ge a) = (Inum bs a ≥ 0)" "Ifm bbs bs (Eq a) = (Inum bs a = 0)" "Ifm bbs bs (NEq a) = (Inum bs a ≠ 0)" "Ifm bbs bs (Dvd i b) = (i dvd Inum bs b)" "Ifm bbs bs (NDvd i b) = (¬(i dvd Inum bs b))" "Ifm bbs bs (NOT p) = (¬ (Ifm bbs bs p))" "Ifm bbs bs (And p q) = (Ifm bbs bs p ∧ Ifm bbs bs q)" "Ifm bbs bs (Or p q) = (Ifm bbs bs p ∨ Ifm bbs bs q)" "Ifm bbs bs (Imp p q) = ((Ifm bbs bs p) --> (Ifm bbs bs q))" "Ifm bbs bs (Iff p q) = (Ifm bbs bs p = Ifm bbs bs q)" "Ifm bbs bs (E p) = (∃ x. Ifm bbs (x#bs) p)" "Ifm bbs bs (A p) = (∀ x. Ifm bbs (x#bs) p)" "Ifm bbs bs (Closed n) = bbs!n" "Ifm bbs bs (NClosed n) = (¬ bbs!n)" consts prep :: "fm => fm" recdef prep "measure fmsize" "prep (E T) = T" "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 (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)))" "prep (E p) = E (prep p)" "prep (A (And p q)) = And (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)) = Or (prep (NOT p)) (prep (NOT q))" "prep (NOT (A p)) = prep (E (NOT p))" "prep (NOT (Or p q)) = And (prep (NOT p)) (prep (NOT q))" "prep (NOT (Imp p q)) = And (prep p) (prep (NOT q))" "prep (NOT (Iff p q)) = Or (prep (And p (NOT q))) (prep (And (NOT p) q))" "prep (NOT p) = NOT (prep p)" "prep (Or p q) = Or (prep p) (prep q)" "prep (And p q) = And (prep p) (prep q)" "prep (Imp p q) = prep (Or (NOT p) q)" "prep (Iff p q) = Or (prep (And p q)) (prep (And (NOT p) (NOT q)))" "prep p = p" (hints simp add: fmsize_pos) lemma prep: "Ifm bbs bs (prep p) = Ifm bbs bs p" by (induct p arbitrary: bs rule: prep.induct, 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 *) subst0:: "num => fm => fm" (* substitue a num into a formula for Bound 0 *) primrec "numbound0 (C c) = True" "numbound0 (Bound n) = (n>0)" "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 (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: gr0_conv_Suc) 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 (Dvd i a) = numbound0 a" "bound0 (NDvd i 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" "bound0 (Closed P) = True" "bound0 (NClosed P) = True" lemma bound0_I: assumes bp: "bound0 p" shows "Ifm bbs (b#bs) p = Ifm bbs (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: gr0_conv_Suc) fun numsubst0:: "num => num => num" where "numsubst0 t (C c) = (C c)" | "numsubst0 t (Bound n) = (if n=0 then t else Bound n)" | "numsubst0 t (CN 0 i a) = Add (Mul i t) (numsubst0 t a)" | "numsubst0 t (CN n i a) = CN n i (numsubst0 t a)" | "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 (Mul i a) = Mul i (numsubst0 t a)" lemma numsubst0_I: "Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b#bs) a)#bs) t" by (induct t rule: numsubst0.induct,auto simp:nth_Cons') lemma numsubst0_I': "numbound0 a ==> Inum (b#bs) (numsubst0 a t) = Inum ((Inum (b'#bs) a)#bs) t" by (induct t rule: numsubst0.induct, auto simp: nth_Cons' numbound0_I[where b="b" and b'="b'"]) primrec "subst0 t T = T" "subst0 t F = F" "subst0 t (Lt a) = Lt (numsubst0 t a)" "subst0 t (Le a) = Le (numsubst0 t a)" "subst0 t (Gt a) = Gt (numsubst0 t a)" "subst0 t (Ge a) = Ge (numsubst0 t a)" "subst0 t (Eq a) = Eq (numsubst0 t a)" "subst0 t (NEq a) = NEq (numsubst0 t a)" "subst0 t (Dvd i a) = Dvd i (numsubst0 t a)" "subst0 t (NDvd i a) = NDvd i (numsubst0 t a)" "subst0 t (NOT p) = NOT (subst0 t p)" "subst0 t (And p q) = And (subst0 t p) (subst0 t q)" "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)" "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)" "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)" "subst0 t (Closed P) = (Closed P)" "subst0 t (NClosed P) = (NClosed P)" lemma subst0_I: assumes qfp: "qfree p" shows "Ifm bbs (b#bs) (subst0 a p) = Ifm bbs ((Inum (b#bs) a)#bs) p" using qfp numsubst0_I[where b="b" and bs="bs" and a="a"] by (induct p) (simp_all add: gr0_conv_Suc) 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 i a) = (CN (n - 1) i (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 (Dvd i a) = Dvd i (decrnum a)" "decr (NDvd i a) = NDvd i (decrnum a)" "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)" "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, auto simp add: gr0_conv_Suc) lemma decr: assumes nb: "bound0 p" shows "Ifm bbs (x#bs) p = Ifm bbs bs (decr p)" using nb by (induct p rule: decr.induct, simp_all add: gr0_conv_Suc 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 (Dvd i b) = True" "isatom (NDvd i b) = True" "isatom (Closed P) = True" "isatom (NClosed P) = True" "isatom p = False" lemma numsubst0_numbound0: assumes nb: "numbound0 t" shows "numbound0 (numsubst0 t a)" using nb apply (induct a rule: numbound0.induct) apply simp_all apply (case_tac n, simp_all) done lemma subst0_bound0: assumes qf: "qfree p" and nb: "numbound0 t" shows "bound0 (subst0 t p)" using qf numsubst0_numbound0[OF nb] by (induct p rule: subst0.induct, auto) 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 bbs bs (djf f p q) = Ifm bbs 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 evaldjf_ex: "Ifm bbs bs (evaldjf f ps) = (∃ p ∈ set ps. Ifm bbs 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 bbs bs q) = Ifm bbs 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. f (Or p q) = Or (f p) (f q)" and fF: "f F = F" shows "Ifm bbs bs (DJ f p) = Ifm bbs bs (f p)" proof- have "Ifm bbs bs (DJ f p) = (∃ q ∈ set (disjuncts p). Ifm bbs bs (f q))" by (simp add: DJ_def evaldjf_ex) also have "… = Ifm bbs 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 bbs bs (qe p) = Ifm bbs bs (E p))" shows "∀ bs p. qfree p --> qfree (DJ qe p) ∧ (Ifm bbs bs ((DJ qe p)) = Ifm bbs 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 bbs bs (DJ qe p) = (∃ q∈ set (disjuncts p). Ifm bbs bs (qe q))" by (simp add: DJ_def evaldjf_ex) also have "… = (∃ q ∈ set(disjuncts p). Ifm bbs bs (E q))" using qe disjuncts_qf[OF qf] by auto also have "… = Ifm bbs bs (E p)" by (induct p rule: disjuncts.induct, auto) finally show "qfree (DJ qe p) ∧ Ifm bbs bs (DJ qe p) = Ifm bbs bs (E p)" using qfth by blast qed (* Simplification *) (* Algebraic simplifications for nums *) consts bnds:: "num => nat list" lex_ns:: "nat list × nat list => bool" recdef bnds "measure size" "bnds (Bound n) = [n]" "bnds (CN n c a) = n#(bnds a)" "bnds (Neg a) = bnds a" "bnds (Add a b) = (bnds a)@(bnds b)" "bnds (Sub a b) = (bnds a)@(bnds b)" "bnds (Mul i a) = bnds a" "bnds a = []" recdef lex_ns "measure (λ (xs,ys). length xs + length ys)" "lex_ns ([], ms) = True" "lex_ns (ns, []) = False" "lex_ns (n#ns, m#ms) = (n<m ∨ ((n = m) ∧ lex_ns (ns,ms))) " constdefs lex_bnd :: "num => num => bool" "lex_bnd t s ≡ lex_ns (bnds t, bnds s)" consts numadd:: "num × num => num" recdef numadd "measure (λ (t,s). num_size t + num_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,Add (Mul c2 (Bound n2)) r2)) else CN n2 c2 (numadd (Add (Mul c1 (Bound n1)) 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" (*function (sequential) numadd :: "num => num => num" where "numadd (Add (Mul c1 (Bound n1)) r1) (Add (Mul c2 (Bound n2)) r2) = (if n1 = n2 then (let c = c1 + c2 in (if c = 0 then numadd r1 r2 else Add (Mul c (Bound n1)) (numadd r1 r2))) else if n1 ≤ n2 then Add (Mul c1 (Bound n1)) (numadd r1 (Add (Mul c2 (Bound n2)) r2)) else Add (Mul c2 (Bound n2)) (numadd (Add (Mul c1 (Bound n1)) r1) r2))" | "numadd (Add (Mul c1 (Bound n1)) r1) t = Add (Mul c1 (Bound n1)) (numadd r1 t)" | "numadd t (Add (Mul c2 (Bound n2)) r2) = Add (Mul c2 (Bound n2)) (numadd t r2)" | "numadd (C b1) (C b2) = C (b1 + b2)" | "numadd a b = Add a b" apply pat_completeness apply auto*) lemma numadd: "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") apply(simp_all add: ring_simps) apply(simp add: left_distrib[symmetric]) done lemma numadd_nb: "[| numbound0 t ; numbound0 s|] ==> numbound0 (numadd (t,s))" by (induct t s rule: numadd.induct, auto simp add: Let_def) fun nummul :: "int => num => num" where "nummul i (C j) = C (i * j)" | "nummul i (CN n c t) = CN n (c*i) (nummul i t)" | "nummul i t = Mul i t" lemma nummul: "!! i. Inum bs (nummul i t) = Inum bs (Mul i t)" by (induct t rule: nummul.induct, auto simp add: ring_simps numadd) lemma nummul_nb: "!! i. numbound0 t ==> numbound0 (nummul i t)" by (induct t rule: nummul.induct, auto simp add: numadd_nb) constdefs numneg :: "num => num" "numneg t ≡ nummul (- 1) t" constdefs numsub :: "num => num => num" "numsub s t ≡ (if s = t then C 0 else numadd (s, numneg t))" lemma numneg: "Inum bs (numneg t) = Inum bs (Neg t)" using numneg_def nummul by simp lemma numneg_nb: "numbound0 t ==> numbound0 (numneg t)" using numneg_def nummul_nb by simp lemma numsub: "Inum bs (numsub a b) = Inum bs (Sub a b)" using numneg numadd numsub_def by simp lemma numsub_nb: "[| numbound0 t ; numbound0 s|] ==> numbound0 (numsub t s)" using numsub_def numadd_nb numneg_nb by simp fun simpnum :: "num => num" where "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 i (simpnum t))" | "simpnum t = t" lemma simpnum_ci: "Inum bs (simpnum t) = Inum bs t" by (induct t rule: simpnum.induct, auto simp add: numneg numadd numsub nummul) lemma simpnum_numbound0: "numbound0 t ==> numbound0 (simpnum t)" by (induct t rule: simpnum.induct, auto simp add: numadd_nb numsub_nb nummul_nb numneg_nb) fun not :: "fm => fm" where "not (NOT p) = p" | "not T = F" | "not F = T" | "not p = NOT p" lemma not: "Ifm bbs bs (not p) = Ifm bbs bs (NOT p)" by (cases p) auto lemma not_qf: "qfree p ==> qfree (not p)" by (cases p, auto) lemma not_bn: "bound0 p ==> bound0 (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 And p q)" lemma conj: "Ifm bbs bs (conj p q) = Ifm bbs bs (And p q)" by (cases "p=F ∨ q=F",simp_all add: conj_def) (cases p,simp_all) lemma conj_qf: "[|qfree p ; qfree q|] ==> qfree (conj p q)" using conj_def by auto lemma conj_nb: "[|bound0 p ; bound0 q|] ==> bound0 (conj p q)" using conj_def by auto 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 Or p q)" lemma disj: "Ifm bbs bs (disj p q) = Ifm bbs bs (Or p q)" by (cases "p=T ∨ q=T",simp_all add: disj_def) (cases p,simp_all) lemma disj_qf: "[|qfree p ; qfree q|] ==> qfree (disj p q)" using disj_def by auto lemma disj_nb: "[|bound0 p ; bound0 q|] ==> bound0 (disj p q)" using disj_def by auto constdefs imp :: "fm => fm => fm" "imp p q ≡ (if (p = F ∨ q=T) then T else if p=T then q else if q=F then not p else Imp p q)" lemma imp: "Ifm bbs bs (imp p q) = Ifm bbs bs (Imp p q)" by (cases "p=F ∨ q=T",simp_all add: imp_def,cases p) (simp_all add: not) lemma imp_qf: "[|qfree p ; qfree q|] ==> qfree (imp p q)" using imp_def by (cases "p=F ∨ q=T",simp_all add: imp_def,cases p) (simp_all add: not_qf) lemma imp_nb: "[|bound0 p ; bound0 q|] ==> bound0 (imp p q)" using imp_def by (cases "p=F ∨ q=T",simp_all add: imp_def,cases p) simp_all 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: "Ifm bbs bs (iff p q) = Ifm bbs bs (Iff p q)" by (unfold iff_def,cases "p=q", simp,cases "p=not q", simp add:not) (cases "not p= q", auto simp add:not) lemma iff_qf: "[|qfree p ; qfree q|] ==> qfree (iff p q)" by (unfold iff_def,cases "p=q", auto simp add: not_qf) lemma iff_nb: "[|bound0 p ; bound0 q|] ==> bound0 (iff p q)" using iff_def by (unfold iff_def,cases "p=q", auto simp add: not_bn) function (sequential) simpfm :: "fm => fm" where "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 (Dvd i a) = (if i=0 then simpfm (Eq a) else if (abs i = 1) then T else let a' = simpnum a in case a' of C v => if (i dvd v) then T else F | _ => Dvd i a')" | "simpfm (NDvd i a) = (if i=0 then simpfm (NEq a) else if (abs i = 1) then F else let a' = simpnum a in case a' of C v => if (¬(i dvd v)) then T else F | _ => NDvd i a')" | "simpfm p = p" by pat_completeness auto termination by (relation "measure fmsize") auto lemma simpfm: "Ifm bbs bs (simpfm p) = Ifm bbs 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 next case (12 i a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp have "i=0 ∨ abs i = 1 ∨ (i≠0 ∧ (abs i ≠ 1))" by auto {assume "i=0" hence ?case using "12.hyps" by (simp add: dvd_def Let_def)} moreover {assume i1: "abs i = 1" from zdvd_1_left[where m = "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"] have ?case using i1 apply (cases "i=0", simp_all add: Let_def) by (cases "i > 0", simp_all)} moreover {assume inz: "i≠0" and cond: "abs i ≠ 1" {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond by (cases "abs i = 1", auto) } moreover {assume "¬ (∃ v. ?sa = C v)" hence "simpfm (Dvd i a) = Dvd i ?sa" using inz cond by (cases ?sa, auto simp add: Let_def) hence ?case using sa by simp} ultimately have ?case by blast} ultimately show ?case by blast next case (13 i a) let ?sa = "simpnum a" from simpnum_ci have sa: "Inum bs ?sa = Inum bs a" by simp have "i=0 ∨ abs i = 1 ∨ (i≠0 ∧ (abs i ≠ 1))" by auto {assume "i=0" hence ?case using "13.hyps" by (simp add: dvd_def Let_def)} moreover {assume i1: "abs i = 1" from zdvd_1_left[where m = "Inum bs a"] uminus_dvd_conv[where d="1" and t="Inum bs a"] have ?case using i1 apply (cases "i=0", simp_all add: Let_def) apply (cases "i > 0", simp_all) done} moreover {assume inz: "i≠0" and cond: "abs i ≠ 1" {fix v assume "?sa = C v" hence ?case using sa[symmetric] inz cond by (cases "abs i = 1", auto) } moreover {assume "¬ (∃ v. ?sa = C v)" hence "simpfm (NDvd i a) = NDvd i ?sa" using inz cond by (cases ?sa, auto simp add: Let_def) hence ?case using sa by simp} ultimately have ?case by blast} 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) next case (12 i 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 (13 i 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)+ (* 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)" (*function (sequential) qelim :: "(fm => fm) => fm => fm" where "qelim qe (E p) = DJ qe (qelim qe p)" | "qelim qe (A p) = not (qe ((qelim qe (NOT p))))" | "qelim qe (NOT p) = not (qelim qe p)" | "qelim qe (And p q) = conj (qelim qe p) (qelim qe q)" | "qelim qe (Or p q) = disj (qelim qe p) (qelim qe q)" | "qelim qe (Imp p q) = imp (qelim qe p) (qelim qe q)" | "qelim qe (Iff p q) = iff (qelim qe p) (qelim qe q)" | "qelim qe p = simpfm p" by pat_completeness auto termination by (relation "measure (fmsize o snd)") auto*) lemma qelim_ci: assumes qe_inv: "∀ bs p. qfree p --> qfree (qe p) ∧ (Ifm bbs bs (qe p) = Ifm bbs bs (E p))" shows "!! bs. qfree (qelim p qe) ∧ (Ifm bbs bs (qelim p qe) = Ifm bbs 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) (* Linearity for fm where Bound 0 ranges over \<int> *) fun zsplit0 :: "num => int × num" (* splits the bounded from the unbounded part*) where "zsplit0 (C c) = (0,C c)" | "zsplit0 (Bound n) = (if n=0 then (1, C 0) else (0,Bound n))" | "zsplit0 (CN n i a) = (let (i',a') = zsplit0 a in if n=0 then (i+i', a') else (i',CN n i a'))" | "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 in (ia+ib, Add a' 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'))" lemma zsplit0_I: shows "!! n a. zsplit0 t = (n,a) ==> (Inum ((x::int) #bs) (CN 0 n a) = Inum (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 next case (2 m n a) thus ?case by (cases "m=0") auto next case (3 m i a n a') let ?j = "fst (zsplit0 a)" let ?b = "snd (zsplit0 a)" have abj: "zsplit0 a = (?j,?b)" by simp {assume "m≠0" with prems(1)[OF abj] prems(2) have ?case by (auto simp add: Let_def split_def)} moreover {assume m0: "m =0" from abj have th: "a'=?b ∧ n=i+?j" using prems by (simp add: Let_def split_def) from abj prems have th2: "(?I x (CN 0 ?j ?b) = ?I x a) ∧ ?N ?b" by blast from th have "?I x (CN 0 n a') = ?I x (CN 0 (i+?j) ?b)" by simp also from th2 have "… = ?I x (CN 0 i (CN 0 ?j ?b))" by (simp add: left_distrib) finally have "?I x (CN 0 n a') = ?I x (CN 0 i a)" using th2 by simp with th2 th have ?case using m0 by blast} ultimately show ?case by blast next case (4 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 prems by (simp add: Let_def split_def) from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) ∧ ?N ?at" by blast from th2[simplified] th[simplified] show ?case by simp next case (5 s t n a) let ?ns = "fst (zsplit0 s)" 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 ultimately have th: "a=Add ?as ?at ∧ n=?ns + ?nt" using prems by (simp add: Let_def split_def) from abjs[symmetric] have bluddy: "∃ x y. (x,y) = zsplit0 s" by blast from prems have "(∃ x y. (x,y) = zsplit0 s) --> (∀xa xb. zsplit0 t = (xa, xb) --> Inum (x # bs) (CN 0 xa xb) = Inum (x # bs) t ∧ numbound0 xb)" by auto with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) ∧ ?N ?at" by blast from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) ∧ ?N ?as" by blast from th3[simplified] th2[simplified] th[simplified] show ?case by (simp add: left_distrib) next case (6 s t n a) let ?ns = "fst (zsplit0 s)" 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 ultimately have th: "a=Sub ?as ?at ∧ n=?ns - ?nt" using prems by (simp add: Let_def split_def) from abjs[symmetric] have bluddy: "∃ x y. (x,y) = zsplit0 s" by blast from prems have "(∃ x y. (x,y) = zsplit0 s) --> (∀xa xb. zsplit0 t = (xa, xb) --> Inum (x # bs) (CN 0 xa xb) = Inum (x # bs) t ∧ numbound0 xb)" by auto with bluddy abjt have th3: "(?I x (CN 0 ?nt ?at) = ?I x t) ∧ ?N ?at" by blast from abjs prems have th2: "(?I x (CN 0 ?ns ?as) = ?I x s) ∧ ?N ?as" by blast from th3[simplified] th2[simplified] th[simplified] show ?case by (simp add: left_diff_distrib) next case (7 i t n a) let ?nt = "fst (zsplit0 t)" let ?at = "snd (zsplit0 t)" have abj: "zsplit0 t = (?nt,?at)" by simp hence th: "a=Mul i ?at ∧ n=i*?nt" using prems by (simp add: Let_def split_def) from abj prems have th2: "(?I x (CN 0 ?nt ?at) = ?I x t) ∧ ?N ?at" by blast hence " ?I x (Mul i t) = i * ?I x (CN 0 ?nt ?at)" by simp also have "… = ?I x (CN 0 (i*?nt) (Mul i ?at))" by (simp add: right_distrib) finally show ?case using th th2 by simp qed consts iszlfm :: "fm => bool" (* Linearity test for fm *) recdef iszlfm "measure size" "iszlfm (And p q) = (iszlfm p ∧ iszlfm q)" "iszlfm (Or p q) = (iszlfm p ∧ iszlfm q)" "iszlfm (Eq (CN 0 c e)) = (c>0 ∧ numbound0 e)" "iszlfm (NEq (CN 0 c e)) = (c>0 ∧ numbound0 e)" "iszlfm (Lt (CN 0 c e)) = (c>0 ∧ numbound0 e)" "iszlfm (Le (CN 0 c e)) = (c>0 ∧ numbound0 e)" "iszlfm (Gt (CN 0 c e)) = (c>0 ∧ numbound0 e)" "iszlfm (Ge (CN 0 c e)) = ( c>0 ∧ numbound0 e)" "iszlfm (Dvd i (CN 0 c e)) = (c>0 ∧ i>0 ∧ numbound0 e)" "iszlfm (NDvd i (CN 0 c e))= (c>0 ∧ i>0 ∧ numbound0 e)" "iszlfm p = (isatom p ∧ (bound0 p))" lemma zlin_qfree: "iszlfm p ==> qfree p" by (induct p rule: iszlfm.induct) auto consts zlfm :: "fm => fm" (* Linearity transformation for fm *) recdef zlfm "measure fmsize" "zlfm (And p q) = And (zlfm p) (zlfm q)" "zlfm (Or p q) = Or (zlfm p) (zlfm q)" "zlfm (Imp p q) = Or (zlfm (NOT p)) (zlfm q)" "zlfm (Iff p q) = Or (And (zlfm p) (zlfm q)) (And (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 (Lt (CN 0 c r)) else (Gt (CN 0 (- c) (Neg r))))" "zlfm (Le a) = (let (c,r) = zsplit0 a in if c=0 then Le r else if c>0 then (Le (CN 0 c r)) else (Ge (CN 0 (- c) (Neg r))))" "zlfm (Gt a) = (let (c,r) = zsplit0 a in if c=0 then Gt r else if c>0 then (Gt (CN 0 c r)) else (Lt (CN 0 (- c) (Neg r))))" "zlfm (Ge a) = (let (c,r) = zsplit0 a in if c=0 then Ge r else if c>0 then (Ge (CN 0 c r)) else (Le (CN 0 (- c) (Neg r))))" "zlfm (Eq a) = (let (c,r) = zsplit0 a in if c=0 then Eq r else if c>0 then (Eq (CN 0 c r)) else (Eq (CN 0 (- c) (Neg r))))" "zlfm (NEq a) = (let (c,r) = zsplit0 a in if c=0 then NEq r else if c>0 then (NEq (CN 0 c r)) else (NEq (CN 0 (- c) (Neg 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 (Dvd (abs i) (CN 0 c r)) else (Dvd (abs i) (CN 0 (- c) (Neg 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 (NDvd (abs i) (CN 0 c r)) else (NDvd (abs i) (CN 0 (- c) (Neg r)))))" "zlfm (NOT (And p q)) = Or (zlfm (NOT p)) (zlfm (NOT q))" "zlfm (NOT (Or p q)) = And (zlfm (NOT p)) (zlfm (NOT q))" "zlfm (NOT (Imp p q)) = And (zlfm p) (zlfm (NOT q))" "zlfm (NOT (Iff p q)) = Or (And(zlfm p) (zlfm(NOT q))) (And (zlfm(NOT p)) (zlfm q))" "zlfm (NOT (NOT p)) = zlfm p" "zlfm (NOT T) = F" "zlfm (NOT F) = T" "zlfm (NOT (Lt a)) = zlfm (Ge a)" "zlfm (NOT (Le a)) = zlfm (Gt a)" "zlfm (NOT (Gt a)) = zlfm (Le a)" "zlfm (NOT (Ge a)) = zlfm (Lt a)" "zlfm (NOT (Eq a)) = zlfm (NEq a)" "zlfm (NOT (NEq a)) = zlfm (Eq a)" "zlfm (NOT (Dvd i a)) = zlfm (NDvd i a)" "zlfm (NOT (NDvd i a)) = zlfm (Dvd i a)" "zlfm (NOT (Closed P)) = NClosed P" "zlfm (NOT (NClosed P)) = Closed P" "zlfm p = p" (hints simp add: fmsize_pos) lemma zlfm_I: assumes qfp: "qfree p" shows "(Ifm bbs (i#bs) (zlfm p) = Ifm bbs (i# bs) p) ∧ iszlfm (zlfm p)" (is "(?I (?l p) = ?I p) ∧ ?L (?l p)") using qfp proof(induct p rule: zlfm.induct) 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 (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "λ t. Inum (i#bs) t" from prems Ia nb show ?case apply (auto simp add: Let_def split_def ring_simps) apply (cases "?r",auto) apply (case_tac nat, auto) done next case (6 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 (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "λ t. Inum (i#bs) t" from prems Ia nb show ?case apply (auto simp add: Let_def split_def ring_simps) apply (cases "?r",auto) apply (case_tac nat, auto) done next 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 (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "λ t. Inum (i#bs) t" from prems Ia nb show ?case apply (auto simp add: Let_def split_def ring_simps) apply (cases "?r",auto) apply (case_tac nat, auto) done next case (8 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 (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "λ t. Inum (i#bs) t" from prems Ia nb show ?case apply (auto simp add: Let_def split_def ring_simps) apply (cases "?r",auto) apply (case_tac nat, auto) done next case (9 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 (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "λ t. Inum (i#bs) t" from prems Ia nb show ?case apply (auto simp add: Let_def split_def ring_simps) apply (cases "?r",auto) apply (case_tac nat, auto) done next case (10 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 (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "λ t. Inum (i#bs) t" from prems Ia nb show ?case apply (auto simp add: Let_def split_def ring_simps) apply (cases "?r",auto) apply (case_tac nat, auto) done next case (11 j 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 (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "λ t. Inum (i#bs) t" have "j=0 ∨ (j≠0 ∧ ?c = 0) ∨ (j≠0 ∧ ?c >0) ∨ (j≠ 0 ∧ ?c<0)" by arith moreover {assume "j=0" hence z: "zlfm (Dvd j a) = (zlfm (Eq a))" by (simp add: Let_def) hence ?case using prems by (simp del: zlfm.simps add: zdvd_0_left)} moreover {assume "?c=0" and "j≠0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] zdvd_abs1[where i="j"] apply (auto simp add: Let_def split_def ring_simps) apply (cases "?r",auto) apply (case_tac nat, auto) done} moreover {assume cp: "?c > 0" and jnz: "j≠0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def) hence ?case using Ia cp jnz by (simp add: Let_def split_def zdvd_abs1[where i="j" and j="(?c*i) + ?N ?r", symmetric])} moreover {assume cn: "?c < 0" and jnz: "j≠0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def) hence ?case using Ia cn jnz zdvd_zminus_iff[where m="abs j" and n="?c*i + ?N ?r" ] by (simp add: Let_def split_def zdvd_abs1[where i="j" and j="(?c*i) + ?N ?r", symmetric])} ultimately show ?case by blast next case (12 j 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 (i # bs) a = Inum (i #bs) (CN 0 ?c ?r)" and nb: "numbound0 ?r" by auto let ?N = "λ t. Inum (i#bs) t" have "j=0 ∨ (j≠0 ∧ ?c = 0) ∨ (j≠0 ∧ ?c >0) ∨ (j≠ 0 ∧ ?c<0)" by arith moreover {assume "j=0" hence z: "zlfm (NDvd j a) = (zlfm (NEq a))" by (simp add: Let_def) hence ?case using prems by (simp del: zlfm.simps add: zdvd_0_left)} moreover {assume "?c=0" and "j≠0" hence ?case using zsplit0_I[OF spl, where x="i" and bs="bs"] zdvd_abs1[where i="j"] apply (auto simp add: Let_def split_def ring_simps) apply (cases "?r",auto) apply (case_tac nat, auto) done} moreover {assume cp: "?c > 0" and jnz: "j≠0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def) hence ?case using Ia cp jnz by (simp add: Let_def split_def zdvd_abs1[where i="j" and j="(?c*i) + ?N ?r", symmetric])} moreover {assume cn: "?c < 0" and jnz: "j≠0" hence l: "?L (?l (Dvd j a))" by (simp add: nb Let_def split_def) hence ?case using Ia cn jnz zdvd_zminus_iff[where m="abs j" and n="?c*i + ?N ?r" ] by (simp add: Let_def split_def zdvd_abs1[where i="j" and j="(?c*i) + ?N ?r", symmetric])} ultimately show ?case by blast qed auto consts plusinf:: "fm => fm" (* Virtual substitution of +∞*) minusinf:: "fm => fm" (* Virtual substitution of -∞*) δ :: "fm => int" (* Compute lcm {d| N? Dvd c*x+t ∈ p}*) dδ :: "fm => int => bool" (* checks if a given l divides all the ds above*) recdef minusinf "measure size" "minusinf (And p q) = And (minusinf p) (minusinf q)" "minusinf (Or p q) = Or (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" lemma minusinf_qfree: "qfree p ==> qfree (minusinf p)" by (induct p rule: minusinf.induct, auto) recdef plusinf "measure size" "plusinf (And p q) = And (plusinf p) (plusinf q)" "plusinf (Or p q) = Or (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" recdef δ "measure size" "δ (And p q) = ilcm (δ p) (δ q)" "δ (Or p q) = ilcm (δ p) (δ q)" "δ (Dvd i (CN 0 c e)) = i" "δ (NDvd i (CN 0 c e)) = i" "δ p = 1" recdef dδ "measure size" "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: assumes lin: "iszlfm p" and d: "d dvd d'" and ad: "dδ p d" shows "dδ p d'" using lin ad d proof(induct p rule: iszlfm.induct) case (9 i c e) thus ?case using d by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) next case (10 i c e) thus ?case using d by (simp add: zdvd_trans[where m="i" and n="d" and k="d'"]) qed simp_all lemma δ : assumes lin:"iszlfm p" shows "dδ p (δ p) ∧ δ p >0" using lin proof (induct p rule: iszlfm.induct) case (1 p q) let ?d = "δ (And p q)" from prems ilcm_pos have dp: "?d >0" by simp have d1: "δ p dvd δ (And p q)" using prems by simp hence th: "dδ p ?d" using delta_mono prems(3-4) by(simp del:dvd_ilcm_self1) have "δ q dvd δ (And p q)" using prems by simp hence th': "dδ q ?d" using delta_mono prems by(simp del:dvd_ilcm_self2) from th th' dp show ?case by simp next case (2 p q) let ?d = "δ (And p q)" from prems ilcm_pos have dp: "?d >0" by simp have "δ p dvd δ (And p q)" using prems by simp hence th: "dδ p ?d" using delta_mono prems by(simp del:dvd_ilcm_self1) have "δ q dvd δ (And p q)" using prems by simp hence th': "dδ q ?d" using delta_mono prems by(simp del:dvd_ilcm_self2) from th th' dp show ?case by simp qed simp_all 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*) β :: "fm => num list" α :: "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β (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)))" "aβ (Le (CN 0 c e)) = (λ k. Le (CN 0 1 (Mul (k div c) e)))" "aβ (Gt (CN 0 c e)) = (λ k. Gt (CN 0 1 (Mul (k div c) e)))" "aβ (Ge (CN 0 c e)) = (λ k. Ge (CN 0 1 (Mul (k div c) e)))" "aβ (Dvd i (CN 0 c e)) =(λ k. Dvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" "aβ (NDvd i (CN 0 c e))=(λ k. NDvd ((k div c)*i) (CN 0 1 (Mul (k div c) e)))" "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β (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)" "dβ (Le (CN 0 c e)) = (λ k. c dvd k)" "dβ (Gt (CN 0 c e)) = (λ k. c dvd k)" "dβ (Ge (CN 0 c e)) = (λ k. c dvd k)" "dβ (Dvd i (CN 0 c e)) =(λ k. c dvd k)" "dβ (NDvd i (CN 0 c e))=(λ k. c dvd k)" "dβ p = (λ k. True)" recdef ζ "measure size" "ζ (And p q) = ilcm (ζ p) (ζ q)" "ζ (Or p q) = ilcm (ζ p) (ζ q)" "ζ (Eq (CN 0 c e)) = c" "ζ (NEq (CN 0 c e)) = c" "ζ (Lt (CN 0 c e)) = c" "ζ (Le (CN 0 c e)) = c" "ζ (Gt (CN 0 c e)) = c" "ζ (Ge (CN 0 c e)) = c" "ζ (Dvd i (CN 0 c e)) = c" "ζ (NDvd i (CN 0 c e))= c" "ζ p = 1" recdef β "measure size" "β (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)) = []" "β (Le (CN 0 c e)) = []" "β (Gt (CN 0 c e)) = [Neg e]" "β (Ge (CN 0 c e)) = [Sub (C -1) e]" "β p = []" recdef α "measure size" "α (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]" "α (Le (CN 0 c e)) = [Add (C -1) e]" "α (Gt (CN 0 c e)) = []" "α (Ge (CN 0 c e)) = []" "α 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 (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))" "mirror (Le (CN 0 c e)) = Ge (CN 0 c (Neg e))" "mirror (Gt (CN 0 c e)) = Lt (CN 0 c (Neg e))" "mirror (Ge (CN 0 c e)) = Le (CN 0 c (Neg e))" "mirror (Dvd i (CN 0 c e)) = Dvd i (CN 0 c (Neg e))" "mirror (NDvd i (CN 0 c e)) = NDvd i (CN 0 c (Neg e))" "mirror p = p" (* Lemmas for the correctness of σρ *) lemma dvd1_eq1: "x >0 ==> (x::int) dvd 1 = (x = 1)" by auto lemma minusinf_inf: assumes linp: "iszlfm p" and u: "dβ p 1" shows "∃ (z::int). ∀ x < z. Ifm bbs (x#bs) (minusinf p) = Ifm bbs (x#bs) p" (is "?P p" is "∃ (z::int). ∀ x < z. ?I x (?M p) = ?I x p") using linp u proof (induct p rule: minusinf.induct) case (1 p q) thus ?case by (auto simp add: dvd1_eq1) (rule_tac x="min z za" in exI,simp) next case (2 p q) thus ?case by (auto simp add: dvd1_eq1) (rule_tac x="min z za" in exI,simp) next case (3 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ hence "∀ x<(- Inum (a#bs) e). c*x + Inum (x#bs) e ≠ 0" proof(clarsimp) fix x assume "x < (- Inum (a#bs) e)" and"x + Inum (x#bs) e = 0" with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] show "False" by simp qed thus ?case by auto next case (4 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ hence "∀ x<(- Inum (a#bs) e). c*x + Inum (x#bs) e ≠ 0" proof(clarsimp) fix x assume "x < (- Inum (a#bs) e)" and"x + Inum (x#bs) e = 0" with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] show "False" by simp qed thus ?case by auto next case (5 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ hence "∀ x<(- Inum (a#bs) e). c*x + Inum (x#bs) e < 0" proof(clarsimp) fix x assume "x < (- Inum (a#bs) e)" with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] show "x + Inum (x#bs) e < 0" by simp qed thus ?case by auto next case (6 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ hence "∀ x<(- Inum (a#bs) e). c*x + Inum (x#bs) e ≤ 0" proof(clarsimp) fix x assume "x < (- Inum (a#bs) e)" with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] show "x + Inum (x#bs) e ≤ 0" by simp qed thus ?case by auto next case (7 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ hence "∀ x<(- Inum (a#bs) e). ¬ (c*x + Inum (x#bs) e > 0)" proof(clarsimp) fix x assume "x < (- Inum (a#bs) e)" and"x + Inum (x#bs) e > 0" with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] show "False" by simp qed thus ?case by auto next case (8 c e) hence c1: "c=1" and nb: "numbound0 e" using dvd1_eq1 by simp+ hence "∀ x<(- Inum (a#bs) e). ¬ (c*x + Inum (x#bs) e ≥ 0)" proof(clarsimp) fix x assume "x < (- Inum (a#bs) e)" and"x + Inum (x#bs) e ≥ 0" with numbound0_I[OF nb, where bs="bs" and b="a" and b'="x"] show "False" by simp qed thus ?case by auto qed auto lemma minusinf_repeats: assumes d: "dδ p d" and linp: "iszlfm p" shows "Ifm bbs ((x - k*d)#bs) (minusinf p) = Ifm bbs (x #bs) (minusinf p)" using linp d 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 proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib, rule iffI) assume "i dvd c * x - c*(k*d) + Inum (x # bs) e" (is "?ri dvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri dvd ?rt") hence "∃ (l::int). ?rt = i * l" by (simp add: dvd_def) hence "∃ (l::int). c*x+ ?I x e = i*l+c*(k * i*di)" by (simp add: ring_simps di_def) hence "∃ (l::int). c*x+ ?I x e = i*(l + c*k*di)" by (simp add: ring_simps) hence "∃ (l::int). c*x+ ?I x e = i*l" by blast thus "i dvd c*x + Inum (x # bs) e" by (simp add: dvd_def) next assume "i dvd c*x + Inum (x # bs) e" (is "?ri dvd ?rc*?rx+?e") hence "∃ (l::int). c*x+?e = i*l" by (simp add: dvd_def) hence "∃ (l::int). c*x - c*(k*d) +?e = i*l - c*(k*d)" by simp hence "∃ (l::int). c*x - c*(k*d) +?e = i*l - c*(k*i*di)" by (simp add: di_def) hence "∃ (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: ring_simps) hence "∃ (l::int). c*x - c * (k*d) +?e = i*l" by blast thus "i dvd c*x - c*(k*d) + Inum (x # bs) e" by (simp add: dvd_def) qed next 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 proof(simp add: numbound0_I[OF nbe,where bs="bs" and b="x - k * d" and b'="x"] right_diff_distrib, rule iffI) assume "i dvd c * x - c*(k*d) + Inum (x # bs) e" (is "?ri dvd ?rc*?rx - ?rc*(?rk*?rd) + ?I x e" is "?ri dvd ?rt") hence "∃ (l::int). ?rt = i * l" by (simp add: dvd_def) hence "∃ (l::int). c*x+ ?I x e = i*l+c*(k * i*di)" by (simp add: ring_simps di_def) hence "∃ (l::int). c*x+ ?I x e = i*(l + c*k*di)" by (simp add: ring_simps) hence "∃ (l::int). c*x+ ?I x e = i*l" by blast thus "i dvd c*x + Inum (x # bs) e" by (simp add: dvd_def) next assume "i dvd c*x + Inum (x # bs) e" (is "?ri dvd ?rc*?rx+?e") hence "∃ (l::int). c*x+?e = i*l" by (simp add: dvd_def) hence "∃ (l::int). c*x - c*(k*d) +?e = i*l - c*(k*d)" by simp hence "∃ (l::int). c*x - c*(k*d) +?e = i*l - c*(k*i*di)" by (simp add: di_def) hence "∃ (l::int). c*x - c*(k*d) +?e = i*((l - c*k*di))" by (simp add: ring_simps) hence "∃ (l::int). c*x - c * (k*d) +?e = i*l" by blast thus "i dvd c*x - c*(k*d) + Inum (x # bs) e" by (simp add: dvd_def) qed qed (auto simp add: gr0_conv_Suc numbound0_I[where bs="bs" and b="x - k*d" and b'="x"]) (* Is'nt this beautiful?*) lemma minusinf_ex: assumes lin: "iszlfm p" and u: "dβ p 1" and exmi: "∃ (x::int). Ifm bbs (x#bs) (minusinf p)" (is "∃ x. ?P1 x") shows "∃ (x::int). Ifm bbs (x#bs) p" (is "∃ x. ?P x") proof- let ?d = "δ p" from δ [OF lin] have dpos: "?d >0" by simp from δ [OF lin] have alld: "dδ p ?d" by simp from minusinf_repeats[OF alld lin] have th1:"∀ x k. ?P1 x = ?P1 (x - (k * ?d))" by simp from minusinf_inf[OF lin u] have th2:"∃ z. ∀ x. x<z --> (?P x = ?P1 x)" by blast from minusinfinity [OF dpos th1 th2] exmi show ?thesis by blast qed (* And This ???*) lemma minusinf_bex: assumes lin: "iszlfm p" shows "(∃ (x::int). Ifm bbs (x#bs) (minusinf p)) = (∃ (x::int)∈ {1..δ p}. Ifm bbs (x#bs) (minusinf p))" (is "(∃ x. ?P x) = _") proof- let ?d = "δ p" from δ [OF lin] have dpos: "?d >0" by simp from δ [OF lin] have alld: "dδ p ?d" by simp from minusinf_repeats[OF alld lin] have th1:"∀ x k. ?P x = ?P (x - (k * ?d))" by simp from periodic_finite_ex[OF dpos th1] show ?thesis by blast qed lemma mirrorαβ: assumes lp: "iszlfm p" shows "(Inum (i#bs)) ` set (α p) = (Inum (i#bs)) ` set (β (mirror p))" using lp by (induct p rule: mirror.induct, auto) lemma mirror: assumes lp: "iszlfm p" shows "Ifm bbs (x#bs) (mirror p) = Ifm bbs ((- x)#bs) p" using lp proof(induct p rule: iszlfm.induct) case (9 j c e) hence nb: "numbound0 e" by simp have "Ifm bbs (x#bs) (mirror (Dvd j (CN 0 c e))) = (j dvd c*x - Inum (x#bs) e)" (is "_ = (j dvd c*x - ?e)") by simp also have "… = (j dvd (- (c*x - ?e)))" by (simp only: zdvd_zminus_iff) also have "… = (j dvd (c* (- x)) + ?e)" apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_def zadd_ac zminus_zadd_distrib) by (simp add: ring_simps) also have "… = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))" using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp finally show ?case . next case (10 j c e) hence nb: "numbound0 e" by simp have "Ifm bbs (x#bs) (mirror (Dvd j (CN 0 c e))) = (j dvd c*x - Inum (x#bs) e)" (is "_ = (j dvd c*x - ?e)") by simp also have "… = (j dvd (- (c*x - ?e)))" by (simp only: zdvd_zminus_iff) also have "… = (j dvd (c* (- x)) + ?e)" apply (simp only: minus_mult_right[symmetric] minus_mult_left[symmetric] diff_def zadd_ac zminus_zadd_distrib) by (simp add: ring_simps) also have "… = Ifm bbs ((- x)#bs) (Dvd j (CN 0 c e))" using numbound0_I[OF nb, where bs="bs" and b="x" and b'="- x"] by simp finally show ?case by simp qed (auto simp add: numbound0_I[where bs="bs" and b="x" and b'="- x"] gr0_conv_Suc) lemma mirror_l: "iszlfm p ∧ dβ p 1 ==> iszlfm (mirror p) ∧ dβ (mirror p) 1" by (induct p rule: mirror.induct, auto) lemma mirror_δ: "iszlfm p ==> δ (mirror p) = δ p" by (induct p rule: mirror.induct,auto) lemma β_numbound0: assumes lp: "iszlfm p" shows "∀ b∈ set (β p). numbound0 b" using lp by (induct p rule: β.induct,auto) lemma dβ_mono: assumes linp: "iszlfm p" and dr: "dβ p l" and d: "l dvd l'" shows "dβ p l'" using dr linp zdvd_trans[where n="l" and k="l'", simplified d] by (induct p rule: iszlfm.induct) simp_all lemma α_l: assumes lp: "iszlfm p" shows "∀ b∈ set (α p). numbound0 b" using lp by(induct p rule: α.induct, auto) lemma ζ: assumes linp: "iszlfm p" shows "ζ p > 0 ∧ dβ p (ζ p)" using linp proof(induct p rule: iszlfm.induct) case (1 p q) from prems have dl1: "ζ p dvd ilcm (ζ p) (ζ q)" by simp from prems have dl2: "ζ q dvd ilcm (ζ p) (ζ q)" by simp from prems dβ_mono[where p = "p" and l="ζ p" and l'="ilcm (ζ p) (ζ q)"] dβ_mono[where p = "q" and l="ζ q" and l'="ilcm (ζ p) (ζ q)"] dl1 dl2 show ?case by (auto simp add: ilcm_pos) next case (2 p q) from prems have dl1: "ζ p dvd ilcm (ζ p) (ζ q)" by simp from prems have dl2: "ζ q dvd ilcm (ζ p) (ζ q)" by simp from prems dβ_mono[where p = "p" and l="ζ p" and l'="ilcm (ζ p) (ζ q)"] dβ_mono[where p = "q" and l="ζ q" and l'="ilcm (ζ p) (ζ q)"] dl1 dl2 show ?case by (auto simp add: ilcm_pos) qed (auto simp add: ilcm_pos) lemma aβ: assumes linp: "iszlfm p" and d: "dβ p l" and lp: "l > 0" shows "iszlfm (aβ p l) ∧ dβ (aβ p l) 1 ∧ (Ifm bbs (l*x #bs) (aβ p l) = Ifm bbs (x#bs) p)" using linp d proof (induct p rule: iszlfm.induct) case (5 c e) hence cp: "c>0" and be: "numbound0 e" 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" by (simp add: zdiv_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(l*x + (l div c) * Inum (x # bs) e < 0) = ((c * (l div c)) * x + (l div c) * Inum (x # bs) e < 0)" by simp also have "… = ((l div c) * (c*x + Inum (x # bs) e) < (l div c) * 0)" by (simp add: ring_simps) also have "… = (c*x + Inum (x # bs) e < 0)" using mult_less_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] be by simp next case (6 c e) hence cp: "c>0" and be: "numbound0 e" 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" by (simp add: zdiv_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(l*x + (l div c) * Inum (x# bs) e ≤ 0) = ((c * (l div c)) * x + (l div c) * Inum (x # bs) e ≤ 0)" by simp also have "… = ((l div c) * (c * x + Inum (x # bs) e) ≤ ((l div c)) * 0)" by (simp add: ring_simps) also have "… = (c*x + Inum (x # bs) e ≤ 0)" using mult_le_0_iff [where a="(l div c)" and b="c*x + Inum (x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] be by simp next case (7 c e) hence cp: "c>0" and be: "numbound0 e" 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" by (simp add: zdiv_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(l*x + (l div c)* Inum (x # bs) e > 0) = ((c * (l div c)) * x + (l div c) * Inum (x # bs) e > 0)" by simp also have "… = ((l div c) * (c * x + Inum (x # bs) e) > ((l div c)) * 0)" by (simp add: ring_simps) also have "… = (c * x + Inum (x # bs) e > 0)" using zero_less_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be by simp next case (8 c e) hence cp: "c>0" and be: "numbound0 e" 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" by (simp add: zdiv_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(l*x + (l div c)* Inum (x # bs) e ≥ 0) = ((c*(l div c))*x + (l div c)* Inum (x # bs) e ≥ 0)" by simp also have "… = ((l div c)*(c*x + Inum (x # bs) e) ≥ ((l div c)) * 0)" by (simp add: ring_simps) also have "… = (c*x + Inum (x # bs) e ≥ 0)" using ldcp zero_le_mult_iff [where a="l div c" and b="c*x + Inum (x # bs) e"] by simp finally show ?case using be numbound0_I[OF be,where b="l*x" and b'="x" and bs="bs"] by simp next case (3 c e) hence cp: "c>0" and be: "numbound0 e" 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" by (simp add: zdiv_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(l * x + (l div c) * Inum (x # bs) e = 0) = ((c * (l div c)) * x + (l div c) * Inum (x # bs) e = 0)" by simp also have "… = ((l div c) * (c * x + Inum (x # bs) e) = ((l div c)) * 0)" by (simp add: ring_simps) also have "… = (c * x + Inum (x # bs) e = 0)" using mult_eq_0_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be by simp next case (4 c e) hence cp: "c>0" and be: "numbound0 e" 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" by (simp add: zdiv_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(l * x + (l div c) * Inum (x # bs) e ≠ 0) = ((c * (l div c)) * x + (l div c) * Inum (x # bs) e ≠ 0)" by simp also have "… = ((l div c) * (c * x + Inum (x # bs) e) ≠ ((l div c)) * 0)" by (simp add: ring_simps) also have "… = (c * x + Inum (x # bs) e ≠ 0)" using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e"] ldcp by simp finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be by simp next case (9 j c e) hence cp: "c>0" and be: "numbound0 e" 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" by (simp add: zdiv_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(∃ (k::int). l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) = (∃ (k::int). (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)" by simp also have "… = (∃ (k::int). (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c)*0)" by (simp add: ring_simps) also have "… = (∃ (k::int). c * x + Inum (x # bs) e - j * k = 0)" using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k"] ldcp by simp also have "… = (∃ (k::int). c * x + Inum (x # bs) e = j * k)" by simp finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be mult_strict_mono[OF ldcp jp ldcp ] by (simp add: dvd_def) next case (10 j c e) hence cp: "c>0" and be: "numbound0 e" 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" by (simp add: zdiv_self[OF cnz]) have "c * (l div c) = c* (l div c) + l mod c" using d' zdvd_iff_zmod_eq_0[where m="c" and n="l"] by simp hence cl:"c * (l div c) =l" using zmod_zdiv_equality[where a="l" and b="c", symmetric] by simp hence "(∃ (k::int). l * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k) = (∃ (k::int). (c * (l div c)) * x + (l div c) * Inum (x # bs) e = ((l div c) * j) * k)" by simp also have "… = (∃ (k::int). (l div c) * (c * x + Inum (x # bs) e - j * k) = (l div c)*0)" by (simp add: ring_simps) also have "… = (∃ (k::int). c * x + Inum (x # bs) e - j * k = 0)" using zero_le_mult_iff [where a="(l div c)" and b="c * x + Inum (x # bs) e - j * k"] ldcp by simp also have "… = (∃ (k::int). c * x + Inum (x # bs) e = j * k)" by simp finally show ?case using numbound0_I[OF be,where b="(l * x)" and b'="x" and bs="bs"] be mult_strict_mono[OF ldcp jp ldcp ] by (simp add: dvd_def) qed (auto simp add: gr0_conv_Suc numbound0_I[where bs="bs" and b="(l * x)" and b'="x"]) lemma aβ_ex: assumes linp: "iszlfm p" and d: "dβ p l" and lp: "l>0" shows "(∃ x. l dvd x ∧ Ifm bbs (x #bs) (aβ p l)) = (∃ (x::int). Ifm bbs (x#bs) p)" (is "(∃ x. l dvd x ∧ ?P x) = (∃ x. ?P' x)") proof- 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 . qed lemma β: assumes lp: "iszlfm p" and u: "dβ p 1" and d: "dδ p d" and dp: "d > 0" and nob: "¬(∃(j::int) ∈ {1 .. d}. ∃ b∈ (Inum (a#bs)) ` set(β p). x = b + j)" and p: "Ifm bbs (x#bs) p" (is "?P x") shows "?P (x - d)" using lp u d dp nob p proof(induct p rule: iszlfm.induct) case (5 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ with dp p c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] prems show ?case by simp next case (6 c e) hence c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ with dp p c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] prems show ?case by simp next case (7 c e) hence p: "Ifm bbs (x #bs) (Gt (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (x # bs) e" {assume "(x-d) +?e > 0" hence ?case using c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] by simp} moreover {assume H: "¬ (x-d) + ?e > 0" let ?v="Neg e" have vb: "?v ∈ set (β (Gt (CN 0 c e)))" by simp from prems(11)[simplified simp_thms Inum.simps β.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]] have nob: "¬ (∃ j∈ {1 ..d}. x = - ?e + j)" by auto from H p have "x + ?e > 0 ∧ x + ?e ≤ d" by (simp add: c1) hence "x + ?e ≥ 1 ∧ x + ?e ≤ d" by simp hence "∃ (j::int) ∈ {1 .. d}. j = x + ?e" by simp hence "∃ (j::int) ∈ {1 .. d}. x = (- ?e + j)" by (simp add: ring_simps) with nob have ?case by auto} ultimately show ?case by blast next case (8 c e) hence p: "Ifm bbs (x #bs) (Ge (CN 0 c e))" and c1: "c=1" and bn:"numbound0 e" using dvd1_eq1[where x="c"] by simp+ let ?e = "Inum (x # bs) e" {assume "(x-d) +?e ≥ 0" hence ?case using c1 numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] by simp} moreover {assume H: "¬ (x-d) + ?e ≥ 0" let ?v="Sub (C -1) e" have vb: "?v ∈ set (β (Ge (CN 0 c e)))" by simp from prems(11)[simplified simp_thms Inum.simps β.simps set.simps bex_simps numbound0_I[OF bn,where b="a" and b'="x" and bs="bs"]] have nob: "¬ (∃ j∈ {1 ..d}. x = - ?e - 1 + j)" by auto from H p have "x + ?e ≥ 0 ∧ x + ?e < d" by (simp add: c1) hence "x + ?e +1 ≥ 1 ∧ x + ?e + 1 ≤ d" by simp hence "∃ (j::int) ∈ {1 .. d}. j = x + ?e + 1" by simp hence "∃ (j::int) ∈ {1 .. d}. x= - ?e - 1 + j" by (simp add: ring_simps) with nob have ?case by simp } ultimately show ?case by blast next case (3 c e) hence p: "Ifm bbs (x #bs) (Eq (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 (x # bs) e" let ?v="(Sub (C -1) e)" have vb: "?v ∈ set (β (Eq (CN 0 c e)))" by simp from p have "x= - ?e" by (simp add: c1) with prems(11) show ?case using dp by simp (erule ballE[where x="1"], simp_all add:ring_simps numbound0_I[OF bn,where b="x"and b'="a"and bs="bs"]) next case (4 c e)hence p: "Ifm bbs (x #bs) (NEq (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 (x # bs) e" let ?v="Neg e" have vb: "?v ∈ set (β (NEq (CN 0 c e)))" by simp {assume "x - d + Inum (((x -d)) # bs) e ≠ 0" hence ?case by (simp add: c1)} moreover {assume H: "x - d + Inum (((x -d)) # bs) e = 0" hence "x = - Inum (((x -d)) # bs) e + d" by simp hence "x = - Inum (a # bs) e + d" by (simp add: numbound0_I[OF bn,where b="x - d"and b'="a"and bs="bs"]) with prems(11) have ?case using dp by simp} ultimately show ?case by blast next case (9 j c e) hence p: "Ifm bbs (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 (x # bs) e" from prems have id: "j dvd d" by simp from c1 have "?p x = (j dvd (x+ ?e))" by simp also have "… = (j dvd x - d + ?e)" using zdvd_period[OF id, where x="x" and c="-1" and t="?e"] by simp finally show ?case using numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] c1 p by simp next case (10 j c e) hence p: "Ifm bbs (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 (x # bs) e" from prems have id: "j dvd d" by simp from c1 have "?p x = (¬ j dvd (x+ ?e))" by simp also have "… = (¬ j dvd x - d + ?e)" using zdvd_period[OF id, where x="x" and c="-1" and t="?e"] by simp finally show ?case using numbound0_I[OF bn,where b="(x-d)" and b'="x" and bs="bs"] c1 p by simp qed (auto simp add: numbound0_I[where bs="bs" and b="(x - d)" and b'="x"] gr0_conv_Suc) lemma β': assumes lp: "iszlfm p" and u: "dβ p 1" and d: "dδ p d" and dp: "d > 0" shows "∀ x. ¬(∃(j::int) ∈ {1 .. d}. ∃ b∈ set(β p). Ifm bbs ((Inum (a#bs) b + j) #bs) p) --> Ifm bbs (x#bs) p --> Ifm bbs ((x - d)#bs) p" (is "∀ x. ?b --> ?P x --> ?P (x - d)") proof(clarify) fix x assume nb:"?b" and px: "?P x" hence nb2: "¬(∃(j::int) ∈ {1 .. d}. ∃ b∈ (Inum (a#bs)) ` set(β p). x = b + j)" by auto from β[OF lp u d dp nb2 px] show "?P (x -d )" . qed 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::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) prefer 2 apply(drule minusinfinity) apply assumption+ apply(fastsimp) apply clarsimp apply(subgoal_tac "!!k. 0<=k ==> !x. P x --> P (x - k*D)") apply(frule_tac x = x and z=z in decr_lemma) apply(subgoal_tac "P1(x - (¦x - z¦ + 1) * D)") prefer 2 apply(subgoal_tac "0 <= (¦x - z¦ + 1)") prefer 2 apply arith apply fastsimp apply(drule (1) periodic_finite_ex) apply blast apply(blast dest:decr_mult_lemma) done theorem cp_thm: assumes lp: "iszlfm p" and u: "dβ p 1" and d: "dδ p d" and dp: "d > 0" shows "(∃ (x::int). Ifm bbs (x #bs) p) = (∃ j∈ {1.. d}. Ifm bbs (j #bs) (minusinf p) ∨ (∃ b ∈ set (β p). Ifm bbs ((Inum (i#bs) b + j) #bs) p))" (is "(∃ (x::int). ?P (x)) = (∃ j∈ ?D. ?M j ∨ (∃ b∈ ?B. ?P (?I b + j)))") proof- from minusinf_inf[OF lp u] have th: "∃(z::int). ∀x<z. ?P (x) = ?M x" by blast let ?B' = "{?I b | b. b∈ ?B}" have BB': "(∃j∈?D. ∃b∈ ?B. ?P (?I b +j)) = (∃ j ∈ ?D. ∃ b ∈ ?B'. ?P (b + j))" by auto hence th2: "∀ x. ¬ (∃ j ∈ ?D. ∃ b ∈ ?B'. ?P ((b + j))) --> ?P (x) --> ?P ((x - d))" using β'[OF lp u d dp, where a="i" and bbs = "bbs"] by blast from minusinf_repeats[OF d lp] have th3: "∀ x k. ?M x = ?M (x-k*d)" by simp from cpmi_eq[OF dp th th2 th3] BB' show ?thesis by blast qed (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *) lemma mirror_ex: assumes lp: "iszlfm p" shows "(∃ x. Ifm bbs (x#bs) (mirror p)) = (∃ x. Ifm bbs (x#bs) p)" (is "(∃ x. ?I x ?mp) = (∃ x. ?I x p)") proof(auto) 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" using mirror[OF lp, where x="- x", symmetric] by auto thus "∃ x. ?I x ?mp" by blast qed lemma cp_thm': assumes lp: "iszlfm p" and up: "dβ p 1" and dd: "dδ p d" and dp: "d > 0" shows "(∃ x. Ifm bbs (x#bs) p) = ((∃ j∈ {1 .. d}. Ifm bbs (j#bs) (minusinf p)) ∨ (∃ j∈ {1.. d}. ∃ b∈ (Inum (i#bs)) ` set (β p). Ifm bbs ((b+j)#bs) p))" using cp_thm[OF lp up dd dp,where i="i"] by auto constdefs unit:: "fm => fm × num list × int" "unit p ≡ (let p' = zlfm p ; l = ζ p' ; q = And (Dvd l (CN 0 1 (C 0))) (aβ p' l); d = δ q; B = remdups (map simpnum (β q)) ; a = remdups (map simpnum (α q)) in if length B ≤ length a then (q,B,d) else (mirror q, a,d))" lemma unit: assumes qf: "qfree p" shows "!! q B d. unit p = (q,B,d) ==> ((∃ x. Ifm bbs (x#bs) p) = (∃ x. Ifm bbs (x#bs) q)) ∧ (Inum (i#bs)) ` set B = (Inum (i#bs)) ` set (β q) ∧ dβ q 1 ∧ dδ q d ∧ d >0 ∧ iszlfm q ∧ (∀ b∈ set B. numbound0 b)" proof- fix q B d assume qBd: "unit p = (q,B,d)" let ?thes = "((∃ x. Ifm bbs (x#bs) p) = (∃ x. Ifm bbs (x#bs) q)) ∧ Inum (i#bs) ` set B = Inum (i#bs) ` set (β q) ∧ dβ q 1 ∧ dδ q d ∧ 0 < d ∧ iszlfm q ∧ (∀ b∈ set B. numbound0 b)" let ?I = "λ x p. Ifm bbs (x#bs) p" let ?p' = "zlfm p" let ?l = "ζ ?p'" let ?q = "And (Dvd ?l (CN 0 1 (C 0))) (aβ ?p' ?l)" let ?d = "δ ?q" let ?B = "set (β ?q)" 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"]] have pp': "∀ i. ?I i ?p' = ?I i p" by auto from conjunct2[OF zlfm_I[OF qf, where bs="bs" and i="i"]] have lp': "iszlfm ?p'" . from lp' ζ[where p="?p'"] 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 from lp' lp aβ[OF lp' dl lp] have lq:"iszlfm ?q" and uq: "dβ ?q 1" by auto from δ[OF lq] have dp:"?d >0" and dd: "dδ ?q ?d" by blast+ let ?N = "λ t. Inum (i#bs) t" have "?N ` set ?B' = ((?N o simpnum) ` ?B)" by auto also have "… = ?N ` ?B" using simpnum_ci[where bs="i#bs"] by auto finally have BB': "?N ` set ?B' = ?N ` ?B" . have "?N ` set ?A' = ((?N o simpnum) ` ?A)" by auto also have "… = ?N ` ?A" using simpnum_ci[where bs="i#bs"] by auto finally have AA': "?N ` set ?A' = ?N ` ?A" . from β_numbound0[OF lq] have B_nb:"∀ b∈ set ?B'. numbound0 b" by (simp add: simpnum_numbound0) from α_l[OF lq] have A_nb: "∀ b∈ set ?A'. numbound0 b" by (simp add: simpnum_numbound0) {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)" and bn: "∀b∈ set B. numbound0 b" by simp+ with pq_ex dp uq dd lq q d have ?thes by simp} 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)" and bn: "∀b∈ set B. numbound0 b" by simp+ 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_l[where p="?q"] have lq': "iszlfm q" and uq: "dβ q 1" by auto from δ[OF lq'] mirror_δ[OF lq] q d have dq:"dδ q d " by auto from pqm_eq b bn uq lq' dp dq q dp d have ?thes by simp } ultimately show ?thes by blast qed (* Cooper's Algorithm *) constdefs cooper :: "fm => fm" "cooper p ≡ (let (q,B,d) = unit p; js = iupt 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 (λ (b,j). simpfm (subst0 (Add b (C j)) q)) [(b,j). b\<leftarrow>B,j\<leftarrow>js] in decr (disj md qd)))" lemma cooper: assumes qf: "qfree p" shows "((∃ x. Ifm bbs (x#bs) p) = (Ifm bbs bs (cooper p))) ∧ qfree (cooper p)" (is "(?lhs = ?rhs) ∧ _") proof- let ?I = "λ x p. Ifm bbs (x#bs) p" let ?q = "fst (unit p)" let ?B = "fst (snd(unit p))" let ?d = "snd (snd (unit p))" let ?js = "iupt 1 ?d" let ?mq = "minusinf ?q" let ?smq = "simpfm ?mq" let ?md = "evaldjf (λ j. simpfm (subst0 (C j) ?smq)) ?js" let ?N = "λ t. Inum (i#bs) t" let ?Bjs = "[(b,j). b\<leftarrow>?B,j\<leftarrow>?js]" let ?qd = "evaldjf (λ (b,j). simpfm (subst0 (Add b (C j)) ?q)) ?Bjs" 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" 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)" by (auto simp only: subst0_bound0[OF qfmq]) hence th: "∀ j∈ set ?js. bound0 (simpfm (subst0 (C j) ?smq))" by (auto simp add: simpfm_bound0) 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. bound0 (subst0 (Add b (C j)) ?q)" using subst0_bound0[OF qfq] by blast hence "∀ (b,j) ∈ set ?Bjs. bound0 (simpfm (subst0 (Add b (C j)) ?q))" using simpfm_bound0 by blast hence th': "∀ x ∈ set ?Bjs. bound0 ((λ (b,j). simpfm (subst0 (Add b (C j)) ?q)) x)" by auto from evaldjf_bound0 [OF th'] have qdb: "bound0 ?qd" by simp 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,where i="i"]] B have "?lhs = (∃ j∈ {1.. ?d}. ?I j ?mq ∨ (∃ b∈ ?N ` set ?B. Ifm bbs ((b+ j)#bs) ?q))" by auto also have "… = (∃ j∈ {1.. ?d}. ?I j ?mq ∨ (∃ b∈ set ?B. Ifm bbs ((?N b+ j)#bs) ?q))" by simp also have "… = ((∃ j∈ {1.. ?d}. ?I j ?mq ) ∨ (∃ j∈ {1.. ?d}. ∃ b∈ set ?B. Ifm bbs ((?N (Add b (C j)))#bs) ?q))" by (simp only: Inum.simps) blast also have "… = ((∃ j∈ {1.. ?d}. ?I j ?smq ) ∨ (∃ j∈ {1.. ?d}. ∃ b∈ set ?B. Ifm bbs ((?N (Add b (C j)))#bs) ?q))" by (simp add: simpfm) also have "… = ((∃ j∈ set ?js. (λ j. ?I i (simpfm (subst0 (C j) ?smq))) j) ∨ (∃ j∈ set ?js. ∃ b∈ set ?B. Ifm bbs ((?N (Add b (C j)))#bs) ?q))" by (simp only: simpfm subst0_I[OF qfmq] iupt_set) auto also have "… = (?I i (evaldjf (λ j. simpfm (subst0 (C j) ?smq)) ?js) ∨ (∃ j∈ set ?js. ∃ b∈ set ?B. ?I i (subst0 (Add b (C j)) ?q)))" by (simp only: evaldjf_ex subst0_I[OF qfq]) also have "…= (?I i ?md ∨ (∃ (b,j) ∈ set ?Bjs. (λ (b,j). ?I i (simpfm (subst0 (Add b (C j)) ?q))) (b,j)))" by (simp only: simpfm set_concat set_map concat_map_singleton UN_simps) blast also have "… = (?I i ?md ∨ (?I i (evaldjf (λ (b,j). simpfm (subst0 (Add b (C j)) ?q)) ?Bjs)))" by (simp only: evaldjf_ex[where bs="i#bs" and f="λ (b,j). simpfm (subst0 (Add b (C j)) ?q)" and ps="?Bjs"]) (auto simp add: split_def) finally have mdqd: "?lhs = (?I i ?md ∨ ?I i ?qd)" by simp also have "… = (?I i (disj ?md ?qd))" by (simp add: disj) also have "… = (Ifm bbs bs (decr (disj ?md ?qd)))" by (simp only: decr [OF mdqdb]) finally have mdqd2: "?lhs = (Ifm bbs bs (decr (disj ?md ?qd)))" . {assume mdT: "?md = T" hence cT:"cooper p = T" by (simp only: cooper_def unit_def split_def Let_def if_True) simp from mdT have lhs:"?lhs" using mdqd by simp 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) with mdqd2 decr_qf[OF mdqdb] have ?thesis by simp } ultimately show ?thesis by blast qed constdefs pa:: "fm => fm" "pa ≡ (λ p. qelim (prep p) cooper)" theorem mirqe: "(Ifm bbs bs (pa p) = Ifm bbs bs p) ∧ qfree (pa p)" using qelim_ci cooper prep by (auto simp add: pa_def) definition cooper_test :: "unit => fm" where "cooper_test u = pa (E (A (Imp (Ge (Sub (Bound 0) (Bound 1))) (E (E (Eq (Sub (Add (Mul 3 (Bound 1)) (Mul 5 (Bound 0))) (Bound 2))))))))" code_reserved SML oo export_code pa cooper_test in SML module_name GeneratedCooper (*export_code pa in SML module_name GeneratedCooper file "~~/src/HOL/Tools/Qelim/raw_generated_cooper.ML"*) ML {* GeneratedCooper.cooper_test () *} use "coopereif.ML" oracle linzqe_oracle ("term") = Coopereif.cooper_oracle use "coopertac.ML" setup "LinZTac.setup" (* Tests *) lemma "∃ (j::int). ∀ x≥j. (∃ a b. x = 3*a+5*b)" by cooper lemma "ALL (x::int) >=8. EX i j. 5*i + 3*j = x" by cooper theorem "(∀(y::int). 3 dvd y) ==> ∀(x::int). b < x --> a ≤ x" by cooper theorem "!! (y::int) (z::int) (n::int). 3 dvd z ==> 2 dvd (y::int) ==> (∃(x::int). 2*x = y) & (∃(k::int). 3*k = z)" by cooper theorem "!! (y::int) (z::int) n. Suc(n::nat) < 6 ==> 3 dvd z ==> 2 dvd (y::int) ==> (∃(x::int). 2*x = y) & (∃(k::int). 3*k = z)" by cooper theorem "∀(x::nat). ∃(y::nat). (0::nat) ≤ 5 --> y = 5 + x " by cooper lemma "ALL (x::int) >=8. EX i j. 5*i + 3*j = x" by cooper lemma "ALL (y::int) (z::int) (n::int). 3 dvd z --> 2 dvd (y::int) --> (EX (x::int). 2*x = y) & (EX (k::int). 3*k = z)" by cooper lemma "ALL(x::int) y. x < y --> 2 * x + 1 < 2 * y" by cooper lemma "ALL(x::int) y. 2 * x + 1 ~= 2 * y" by cooper lemma "EX(x::int) y. 0 < x & 0 <= y & 3 * x - 5 * y = 1" by cooper lemma "~ (EX(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)" by cooper lemma "ALL(x::int). (2 dvd x) --> (EX(y::int). x = 2*y)" by cooper lemma "ALL(x::int). (2 dvd x) --> (EX(y::int). x = 2*y)" by cooper lemma "ALL(x::int). (2 dvd x) = (EX(y::int). x = 2*y)" by cooper lemma "ALL(x::int). ((2 dvd x) = (ALL(y::int). x ~= 2*y + 1))" by cooper lemma "~ (ALL(x::int). ((2 dvd x) = (ALL(y::int). x ~= 2*y+1) | (EX(q::int) (u::int) i. 3*i + 2*q - u < 17) --> 0 < x | ((~ 3 dvd x) &(x + 8 = 0))))" by cooper lemma "~ (ALL(i::int). 4 <= i --> (EX x y. 0 <= x & 0 <= y & 3 * x + 5 * y = i))" by cooper lemma "EX j. ALL (x::int) >= j. EX i j. 5*i + 3*j = x" by cooper theorem "(∀(y::int). 3 dvd y) ==> ∀(x::int). b < x --> a ≤ x" by cooper theorem "!! (y::int) (z::int) (n::int). 3 dvd z ==> 2 dvd (y::int) ==> (∃(x::int). 2*x = y) & (∃(k::int). 3*k = z)" by cooper theorem "!! (y::int) (z::int) n. Suc(n::nat) < 6 ==> 3 dvd z ==> 2 dvd (y::int) ==> (∃(x::int). 2*x = y) & (∃(k::int). 3*k = z)" by cooper theorem "∀(x::nat). ∃(y::nat). (0::nat) ≤ 5 --> y = 5 + x " by cooper theorem "∀(x::nat). ∃(y::nat). y = 5 + x | x div 6 + 1= 2" by cooper theorem "∃(x::int). 0 < x" by cooper theorem "∀(x::int) y. x < y --> 2 * x + 1 < 2 * y" by cooper theorem "∀(x::int) y. 2 * x + 1 ≠ 2 * y" by cooper theorem "∃(x::int) y. 0 < x & 0 ≤ y & 3 * x - 5 * y = 1" by cooper theorem "~ (∃(x::int) (y::int) (z::int). 4*x + (-6::int)*y = 1)" by cooper theorem "~ (∃(x::int). False)" by cooper theorem "∀(x::int). (2 dvd x) --> (∃(y::int). x = 2*y)" by cooper theorem "∀(x::int). (2 dvd x) --> (∃(y::int). x = 2*y)" by cooper theorem "∀(x::int). (2 dvd x) = (∃(y::int). x = 2*y)" by cooper theorem "∀(x::int). ((2 dvd x) = (∀(y::int). x ≠ 2*y + 1))" by cooper theorem "~ (∀(x::int). ((2 dvd x) = (∀(y::int). x ≠ 2*y+1) | (∃(q::int) (u::int) i. 3*i + 2*q - u < 17) --> 0 < x | ((~ 3 dvd x) &(x + 8 = 0))))" by cooper theorem "~ (∀(i::int). 4 ≤ i --> (∃x y. 0 ≤ x & 0 ≤ y & 3 * x + 5 * y = i))" by cooper theorem "∀(i::int). 8 ≤ i --> (∃x y. 0 ≤ x & 0 ≤ y & 3 * x + 5 * y = i)" by cooper theorem "∃(j::int). ∀i. j ≤ i --> (∃x y. 0 ≤ x & 0 ≤ y & 3 * x + 5 * y = i)" by cooper theorem "~ (∀j (i::int). j ≤ i --> (∃x y. 0 ≤ x & 0 ≤ y & 3 * x + 5 * y = i))" by cooper theorem "(∃m::nat. n = 2 * m) --> (n + 1) div 2 = n div 2" by cooper end
lemma iupt_set:
set (iupt i j) = {i..j}
lemma fmsize_pos:
0 < fmsize p
lemma prep:
Ifm bbs bs (prep p) = Ifm bbs bs p
lemma numbound0_I:
numbound0 a ==> Inum (b # bs) a = Inum (b' # bs) a
lemma bound0_I:
bound0 p ==> Ifm bbs (b # bs) p = Ifm bbs (b' # bs) p
lemma numsubst0_I:
Inum (b # bs) (numsubst0 a t) = Inum (Inum (b # bs) a # bs) t
lemma numsubst0_I':
numbound0 a ==> Inum (b # bs) (numsubst0 a t) = Inum (Inum (b' # bs) a # bs) t
lemma subst0_I:
qfree p ==> Ifm bbs (b # bs) (subst0 a p) = Ifm bbs (Inum (b # bs) a # bs) p
lemma decrnum:
numbound0 t ==> Inum (x # bs) t = Inum bs (decrnum t)
lemma decr:
bound0 p ==> Ifm bbs (x # bs) p = Ifm bbs bs (decr p)
lemma decr_qf:
bound0 p ==> qfree (decr p)
lemma numsubst0_numbound0:
numbound0 t ==> numbound0 (numsubst0 t a)
lemma subst0_bound0:
[| qfree p; numbound0 t |] ==> bound0 (subst0 t p)
lemma bound0_qf:
bound0 p ==> qfree p
lemma djf_Or:
Ifm bbs bs (djf f p q) = Ifm bbs bs (Or (f p) q)
lemma evaldjf_ex:
Ifm bbs bs (evaldjf f ps) = (∃p∈set ps. Ifm bbs 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 bbs bs q) = Ifm bbs 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. f (Or p q) = Or (f p) (f q); f F = F |]
==> Ifm bbs bs (DJ f p) = Ifm bbs 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 bbs bs (qe p) = Ifm bbs bs (E p)
==> ∀bs p. qfree p --> qfree (DJ qe p) ∧ Ifm bbs bs (DJ qe p) = Ifm bbs bs (E p)
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 i t) = Inum bs (Mul i t)
lemma nummul_nb:
numbound0 t ==> numbound0 (nummul i t)
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 not:
Ifm bbs bs (not p) = Ifm bbs bs (NOT p)
lemma not_qf:
qfree p ==> qfree (not p)
lemma not_bn:
bound0 p ==> bound0 (not p)
lemma conj:
Ifm bbs bs (conj p q) = Ifm bbs bs (And p q)
lemma conj_qf:
[| qfree p; qfree q |] ==> qfree (conj p q)
lemma conj_nb:
[| bound0 p; bound0 q |] ==> bound0 (conj p q)
lemma disj:
Ifm bbs bs (disj p q) = Ifm bbs bs (Or 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:
Ifm bbs bs (imp p q) = Ifm bbs bs (Imp 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:
Ifm bbs bs (Reflected_Presburger.iff p q) = Ifm bbs bs (Iff p q)
lemma iff_qf:
[| qfree p; qfree q |] ==> qfree (Reflected_Presburger.iff p q)
lemma iff_nb:
[| bound0 p; bound0 q |] ==> bound0 (Reflected_Presburger.iff p q)
lemma simpfm:
Ifm bbs bs (simpfm p) = Ifm bbs bs p
lemma simpfm_bound0:
bound0 p ==> bound0 (simpfm p)
lemma simpfm_qf:
qfree p ==> qfree (simpfm p)
lemma qelim_ci:
∀bs p. qfree p --> qfree (qe p) ∧ Ifm bbs bs (qe p) = Ifm bbs bs (E p)
==> qfree (qelim p qe) ∧ Ifm bbs bs (qelim p qe) = Ifm bbs bs p
lemma zsplit0_I:
zsplit0 t = (n, a) ==> Inum (x # bs) (CN 0 n a) = Inum (x # bs) t ∧ numbound0 a
lemma zlin_qfree:
iszlfm p ==> qfree p
lemma zlfm_I:
qfree p ==> Ifm bbs (i # bs) (zlfm p) = Ifm bbs (i # bs) p ∧ iszlfm (zlfm p)
lemma minusinf_qfree:
qfree p ==> qfree (minusinf p)
lemma delta_mono:
[| iszlfm p; d dvd d'; dδ p d |] ==> dδ p d'
lemma δ:
iszlfm p ==> dδ p (δ p) ∧ 0 < δ p
lemma dvd1_eq1:
0 < x ==> (x dvd 1) = (x = 1)
lemma minusinf_inf:
[| iszlfm p; dβ p 1 |]
==> ∃z. ∀x<z. Ifm bbs (x # bs) (minusinf p) = Ifm bbs (x # bs) p
lemma minusinf_repeats:
[| dδ p d; iszlfm p |]
==> Ifm bbs ((x - k * d) # bs) (minusinf p) = Ifm bbs (x # bs) (minusinf p)
lemma minusinf_ex:
[| iszlfm p; dβ p 1; ∃x. Ifm bbs (x # bs) (minusinf p) |]
==> ∃x. Ifm bbs (x # bs) p
lemma minusinf_bex:
iszlfm p
==> (∃x. Ifm bbs (x # bs) (minusinf p)) =
(∃x∈{1..δ p}. Ifm bbs (x # bs) (minusinf p))
lemma mirrorαβ:
iszlfm p ==> Inum (i # bs) ` set (α p) = Inum (i # bs) ` set (β (mirror p))
lemma mirror:
iszlfm p ==> Ifm bbs (x # bs) (mirror p) = Ifm bbs (- x # bs) p
lemma mirror_l:
iszlfm p ∧ dβ p 1 ==> iszlfm (mirror p) ∧ dβ (mirror p) 1
lemma mirror_δ:
iszlfm p ==> δ (mirror p) = δ p
lemma β_numbound0:
iszlfm p ==> ∀b∈set (β p). numbound0 b
lemma dβ_mono:
[| iszlfm p; dβ p l; l dvd l' |] ==> dβ p l'
lemma α_l:
iszlfm p ==> ∀b∈set (α p). numbound0 b
lemma ζ:
iszlfm p ==> 0 < ζ p ∧ dβ p (ζ p)
lemma aβ:
[| iszlfm p; dβ p l; 0 < l |]
==> iszlfm (aβ p l) ∧
dβ (aβ p l) 1 ∧ Ifm bbs (l * x # bs) (aβ p l) = Ifm bbs (x # bs) p
lemma aβ_ex:
[| iszlfm p; dβ p l; 0 < l |]
==> (∃x. l dvd x ∧ Ifm bbs (x # bs) (aβ p l)) = (∃x. Ifm bbs (x # bs) p)
lemma β:
[| iszlfm p; dβ p 1; dδ p d; 0 < d;
¬ (∃j∈{1..d}. ∃b∈Inum (a # bs) ` set (β p). x = b + j); Ifm bbs (x # bs) p |]
==> Ifm bbs ((x - d) # bs) p
lemma β':
[| iszlfm p; dβ p 1; dδ p d; 0 < d |]
==> ∀x. ¬ (∃j∈{1..d}. ∃b∈set (β p). Ifm bbs ((Inum (a # bs) b + j) # bs) p) -->
Ifm bbs (x # bs) p --> Ifm bbs ((x - d) # bs) p
lemma cpmi_eq:
[| 0 < D; ∃z. ∀x<z. P x = P1.0 x;
∀x. ¬ (∃j∈{1..D}. ∃b∈B. P (b + j)) --> P x --> P (x - D);
∀x k. P1.0 x = P1.0 (x - k * D) |]
==> (∃x. P x) = ((∃j∈{1..D}. P1.0 j) ∨ (∃j∈{1..D}. ∃b∈B. P (b + j)))
theorem cp_thm:
[| iszlfm p; dβ p 1; dδ p d; 0 < d |]
==> (∃x. Ifm bbs (x # bs) p) =
(∃j∈{1..d}.
Ifm bbs (j # bs) (minusinf p) ∨
(∃b∈set (β p). Ifm bbs ((Inum (i # bs) b + j) # bs) p))
lemma mirror_ex:
iszlfm p ==> (∃x. Ifm bbs (x # bs) (mirror p)) = (∃x. Ifm bbs (x # bs) p)
lemma cp_thm':
[| iszlfm p; dβ p 1; dδ p d; 0 < d |]
==> (∃x. Ifm bbs (x # bs) p) =
((∃j∈{1..d}. Ifm bbs (j # bs) (minusinf p)) ∨
(∃j∈{1..d}. ∃b∈Inum (i # bs) ` set (β p). Ifm bbs ((b + j) # bs) p))
lemma unit:
[| qfree p; Reflected_Presburger.unit p = (q, B, d) |]
==> (∃x. Ifm bbs (x # bs) p) = (∃x. Ifm bbs (x # bs) q) ∧
Inum (i # bs) ` set B = Inum (i # bs) ` set (β q) ∧
dβ q 1 ∧ dδ q d ∧ 0 < d ∧ iszlfm q ∧ (∀b∈set B. numbound0 b)
lemma cooper:
qfree p ==> (∃x. Ifm bbs (x # bs) p) = Ifm bbs bs (cooper p) ∧ qfree (cooper p)
theorem mirqe:
Ifm bbs bs (pa p) = Ifm bbs bs p ∧ qfree (pa p)
lemma
∃j. ∀x≥j. ∃a b. x = 3 * a + 5 * b [!]
lemma
∀x≥8. ∃i j. 5 * i + 3 * j = x [!]
theorem
∀y. 3 dvd y ==> ∀x>b. a ≤ x [!]
theorem
[| 3 dvd z; 2 dvd y |] ==> (∃x. 2 * x = y) ∧ (∃k. 3 * k = z) [!]
theorem
[| Suc n < 6; 3 dvd z; 2 dvd y |] ==> (∃x. 2 * x = y) ∧ (∃k. 3 * k = z) [!]
theorem
∀x. ∃y. 0 ≤ 5 --> y = 5 + x
lemma
∀x≥8. ∃i j. 5 * i + 3 * j = x [!]
lemma
∀y z n. 3 dvd z --> 2 dvd y --> (∃x. 2 * x = y) ∧ (∃k. 3 * k = z) [!]
lemma
∀x y. x < y --> 2 * x + 1 < 2 * y
lemma
∀x y. 2 * x + 1 ≠ 2 * y [!]
lemma
∃x y. 0 < x ∧ 0 ≤ y ∧ 3 * x - 5 * y = 1 [!]
lemma
¬ (∃x y z. 4 * x + -6 * y = 1) [!]
lemma
∀x. 2 dvd x --> (∃y. x = 2 * y) [!]
lemma
∀x. 2 dvd x --> (∃y. x = 2 * y) [!]
lemma
∀x. (2 dvd x) = (∃y. x = 2 * y) [!]
lemma
∀x. (2 dvd x) = (∀y. x ≠ 2 * y + 1) [!]
lemma
¬ (∀x. (2 dvd x) = (∀y. x ≠ 2 * y + 1) ∨ (∃q u i. 3 * i + 2 * q - u < 17) -->
0 < x ∨ ¬ 3 dvd x ∧ x + 8 = 0)
[!]
lemma
¬ (∀i≥4. ∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i) [!]
lemma
∃j. ∀x≥j. ∃i j. 5 * i + 3 * j = x [!]
theorem
∀y. 3 dvd y ==> ∀x>b. a ≤ x [!]
theorem
[| 3 dvd z; 2 dvd y |] ==> (∃x. 2 * x = y) ∧ (∃k. 3 * k = z) [!]
theorem
[| Suc n < 6; 3 dvd z; 2 dvd y |] ==> (∃x. 2 * x = y) ∧ (∃k. 3 * k = z) [!]
theorem
∀x. ∃y. 0 ≤ 5 --> y = 5 + x
theorem
∀x. ∃y. y = 5 + x ∨ x div 6 + 1 = 2 [!]
theorem
∃x. 0 < x [!]
theorem
∀x y. x < y --> 2 * x + 1 < 2 * y
theorem
∀x y. 2 * x + 1 ≠ 2 * y [!]
theorem
∃x y. 0 < x ∧ 0 ≤ y ∧ 3 * x - 5 * y = 1 [!]
theorem
¬ (∃x y z. 4 * x + -6 * y = 1) [!]
theorem
¬ (∃x. False)
theorem
∀x. 2 dvd x --> (∃y. x = 2 * y) [!]
theorem
∀x. 2 dvd x --> (∃y. x = 2 * y) [!]
theorem
∀x. (2 dvd x) = (∃y. x = 2 * y) [!]
theorem
∀x. (2 dvd x) = (∀y. x ≠ 2 * y + 1) [!]
theorem
¬ (∀x. (2 dvd x) = (∀y. x ≠ 2 * y + 1) ∨ (∃q u i. 3 * i + 2 * q - u < 17) -->
0 < x ∨ ¬ 3 dvd x ∧ x + 8 = 0)
[!]
theorem
¬ (∀i≥4. ∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i) [!]
theorem
∀i≥8. ∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i [!]
theorem
∃j. ∀i≥j. ∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i [!]
theorem
¬ (∀j i. j ≤ i --> (∃x y. 0 ≤ x ∧ 0 ≤ y ∧ 3 * x + 5 * y = i)) [!]
theorem
(∃m. n = 2 * m) --> (n + 1) div 2 = n div 2 [!]