# HG changeset patch # User chaieb # Date 1256457455 -3600 # Node ID 241cfaed158f52a50a0f31b2f6fd86935bb7a658 # Parent fe29679cabc23a34bebb7e8ca6a25928d7720260 Add a quantifier elimination for parametric linear arithmetic over ordered fields (parameters are multivariate polynomials) diff -r fe29679cabc2 -r 241cfaed158f src/HOL/Decision_Procs/Decision_Procs.thy --- a/src/HOL/Decision_Procs/Decision_Procs.thy Fri Oct 23 10:11:56 2009 +0200 +++ b/src/HOL/Decision_Procs/Decision_Procs.thy Sun Oct 25 08:57:35 2009 +0100 @@ -1,7 +1,7 @@ header {* Various decision procedures. typically involving reflection *} theory Decision_Procs -imports Cooper Ferrack MIR Approximation Dense_Linear_Order "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex" +imports Cooper Ferrack MIR Approximation Dense_Linear_Order "ex/Approximation_Ex" "ex/Dense_Linear_Order_Ex" Parametric_Ferrante_Rackoff begin end \ No newline at end of file diff -r fe29679cabc2 -r 241cfaed158f src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Sun Oct 25 08:57:35 2009 +0100 @@ -0,0 +1,3227 @@ +(* Title: HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy + Author: Amine Chaieb +*) + +header{* A formalization of Ferrante and Rackoff's procedure with polynomial parameters, see Paper in CALCULEMUS 2008 *} + +theory Parametric_Ferrante_Rackoff +imports Reflected_Multivariate_Polynomial + "~~/src/HOL/Decision_Procs/Dense_Linear_Order" +begin + + +subsection {* Terms *} + +datatype tm = CP poly | Bound nat | Add tm tm | Mul poly tm + | Neg tm | Sub tm tm | CNP nat poly tm + (* A size for poly to make inductive proofs simpler*) + +consts tmsize :: "tm \ nat" +primrec + "tmsize (CP c) = polysize c" + "tmsize (Bound n) = 1" + "tmsize (Neg a) = 1 + tmsize a" + "tmsize (Add a b) = 1 + tmsize a + tmsize b" + "tmsize (Sub a b) = 3 + tmsize a + tmsize b" + "tmsize (Mul c a) = 1 + polysize c + tmsize a" + "tmsize (CNP n c a) = 3 + polysize c + tmsize a " + + (* Semantics of terms tm *) +consts Itm :: "'a::{ring_char_0,division_by_zero,field} list \ 'a list \ tm \ 'a" +primrec + "Itm vs bs (CP c) = (Ipoly vs c)" + "Itm vs bs (Bound n) = bs!n" + "Itm vs bs (Neg a) = -(Itm vs bs a)" + "Itm vs bs (Add a b) = Itm vs bs a + Itm vs bs b" + "Itm vs bs (Sub a b) = Itm vs bs a - Itm vs bs b" + "Itm vs bs (Mul c a) = (Ipoly vs c) * Itm vs bs a" + "Itm vs bs (CNP n c t) = (Ipoly vs c)*(bs!n) + Itm vs bs t" + + +fun allpolys:: "(poly \ bool) \ tm \ bool" where + "allpolys P (CP c) = P c" +| "allpolys P (CNP n c p) = (P c \ allpolys P p)" +| "allpolys P (Mul c p) = (P c \ allpolys P p)" +| "allpolys P (Neg p) = allpolys P p" +| "allpolys P (Add p q) = (allpolys P p \ allpolys P q)" +| "allpolys P (Sub p q) = (allpolys P p \ allpolys P q)" +| "allpolys P p = True" + +consts + tmboundslt:: "nat \ tm \ bool" + tmbound0:: "tm \ bool" (* a tm is INDEPENDENT of Bound 0 *) + tmbound:: "nat \ tm \ bool" (* a tm is INDEPENDENT of Bound n *) + incrtm0:: "tm \ tm" + incrtm:: "nat \ tm \ tm" + decrtm0:: "tm \ tm" + decrtm:: "nat \ tm \ tm" +primrec + "tmboundslt n (CP c) = True" + "tmboundslt n (Bound m) = (m < n)" + "tmboundslt n (CNP m c a) = (m < n \ tmboundslt n a)" + "tmboundslt n (Neg a) = tmboundslt n a" + "tmboundslt n (Add a b) = (tmboundslt n a \ tmboundslt n b)" + "tmboundslt n (Sub a b) = (tmboundslt n a \ tmboundslt n b)" + "tmboundslt n (Mul i a) = tmboundslt n a" +primrec + "tmbound0 (CP c) = True" + "tmbound0 (Bound n) = (n>0)" + "tmbound0 (CNP n c a) = (n\0 \ tmbound0 a)" + "tmbound0 (Neg a) = tmbound0 a" + "tmbound0 (Add a b) = (tmbound0 a \ tmbound0 b)" + "tmbound0 (Sub a b) = (tmbound0 a \ tmbound0 b)" + "tmbound0 (Mul i a) = tmbound0 a" +lemma tmbound0_I: + assumes nb: "tmbound0 a" + shows "Itm vs (b#bs) a = Itm vs (b'#bs) a" +using nb +by (induct a rule: tmbound0.induct,auto simp add: nth_pos2) + +primrec + "tmbound n (CP c) = True" + "tmbound n (Bound m) = (n \ m)" + "tmbound n (CNP m c a) = (n\m \ tmbound n a)" + "tmbound n (Neg a) = tmbound n a" + "tmbound n (Add a b) = (tmbound n a \ tmbound n b)" + "tmbound n (Sub a b) = (tmbound n a \ tmbound n b)" + "tmbound n (Mul i a) = tmbound n a" +lemma tmbound0_tmbound_iff: "tmbound 0 t = tmbound0 t" by (induct t, auto) + +lemma tmbound_I: + assumes bnd: "tmboundslt (length bs) t" and nb: "tmbound n t" and le: "n \ length bs" + shows "Itm vs (bs[n:=x]) t = Itm vs bs t" + using nb le bnd + by (induct t rule: tmbound.induct , auto) + +recdef decrtm0 "measure size" + "decrtm0 (Bound n) = Bound (n - 1)" + "decrtm0 (Neg a) = Neg (decrtm0 a)" + "decrtm0 (Add a b) = Add (decrtm0 a) (decrtm0 b)" + "decrtm0 (Sub a b) = Sub (decrtm0 a) (decrtm0 b)" + "decrtm0 (Mul c a) = Mul c (decrtm0 a)" + "decrtm0 (CNP n c a) = CNP (n - 1) c (decrtm0 a)" + "decrtm0 a = a" +recdef incrtm0 "measure size" + "incrtm0 (Bound n) = Bound (n + 1)" + "incrtm0 (Neg a) = Neg (incrtm0 a)" + "incrtm0 (Add a b) = Add (incrtm0 a) (incrtm0 b)" + "incrtm0 (Sub a b) = Sub (incrtm0 a) (incrtm0 b)" + "incrtm0 (Mul c a) = Mul c (incrtm0 a)" + "incrtm0 (CNP n c a) = CNP (n + 1) c (incrtm0 a)" + "incrtm0 a = a" +lemma decrtm0: assumes nb: "tmbound0 t" + shows "Itm vs (x#bs) t = Itm vs bs (decrtm0 t)" + using nb by (induct t rule: decrtm0.induct, simp_all add: nth_pos2) +lemma incrtm0: "Itm vs (x#bs) (incrtm0 t) = Itm vs bs t" + by (induct t rule: decrtm0.induct, simp_all add: nth_pos2) + +primrec + "decrtm m (CP c) = (CP c)" + "decrtm m (Bound n) = (if n < m then Bound n else Bound (n - 1))" + "decrtm m (Neg a) = Neg (decrtm m a)" + "decrtm m (Add a b) = Add (decrtm m a) (decrtm m b)" + "decrtm m (Sub a b) = Sub (decrtm m a) (decrtm m b)" + "decrtm m (Mul c a) = Mul c (decrtm m a)" + "decrtm m (CNP n c a) = (if n < m then CNP n c (decrtm m a) else CNP (n - 1) c (decrtm m a))" + +consts removen:: "nat \ 'a list \ 'a list" +primrec + "removen n [] = []" + "removen n (x#xs) = (if n=0 then xs else (x#(removen (n - 1) xs)))" + +lemma removen_same: "n \ length xs \ removen n xs = xs" + by (induct xs arbitrary: n, auto) + +lemma nth_length_exceeds: "n \ length xs \ xs!n = []!(n - length xs)" + by (induct xs arbitrary: n, auto) + +lemma removen_length: "length (removen n xs) = (if n \ length xs then length xs else length xs - 1)" + by (induct xs arbitrary: n, auto) +lemma removen_nth: "(removen n xs)!m = (if n \ length xs then xs!m + else if m < n then xs!m else if m \ length xs then xs!(Suc m) else []!(m - (length xs - 1)))" +proof(induct xs arbitrary: n m) + case Nil thus ?case by simp +next + case (Cons x xs n m) + {assume nxs: "n \ length (x#xs)" hence ?case using removen_same[OF nxs] by simp} + moreover + {assume nxs: "\ (n \ length (x#xs))" + {assume mln: "m < n" hence ?case using prems by (cases m, auto)} + moreover + {assume mln: "\ (m < n)" + + {assume mxs: "m \ length (x#xs)" hence ?case using prems by (cases m, auto)} + moreover + {assume mxs: "\ (m \ length (x#xs))" + have th: "length (removen n (x#xs)) = length xs" + using removen_length[where n="n" and xs="x#xs"] nxs by simp + with mxs have mxs':"m \ length (removen n (x#xs))" by auto + hence "(removen n (x#xs))!m = [] ! (m - length xs)" + using th nth_length_exceeds[OF mxs'] by auto + hence th: "(removen n (x#xs))!m = [] ! (m - (length (x#xs) - 1))" + by auto + hence ?case using nxs mln mxs by auto } + ultimately have ?case by blast + } + ultimately have ?case by blast + + } ultimately show ?case by blast +qed + +lemma decrtm: assumes bnd: "tmboundslt (length bs) t" and nb: "tmbound m t" + and nle: "m \ length bs" + shows "Itm vs (removen m bs) (decrtm m t) = Itm vs bs t" + using bnd nb nle + by (induct t rule: decrtm.induct, auto simp add: removen_nth) + +consts tmsubst0:: "tm \ tm \ tm" +primrec + "tmsubst0 t (CP c) = CP c" + "tmsubst0 t (Bound n) = (if n=0 then t else Bound n)" + "tmsubst0 t (CNP n c a) = (if n=0 then Add (Mul c t) (tmsubst0 t a) else CNP n c (tmsubst0 t a))" + "tmsubst0 t (Neg a) = Neg (tmsubst0 t a)" + "tmsubst0 t (Add a b) = Add (tmsubst0 t a) (tmsubst0 t b)" + "tmsubst0 t (Sub a b) = Sub (tmsubst0 t a) (tmsubst0 t b)" + "tmsubst0 t (Mul i a) = Mul i (tmsubst0 t a)" +lemma tmsubst0: + shows "Itm vs (x#bs) (tmsubst0 t a) = Itm vs ((Itm vs (x#bs) t)#bs) a" +by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2) + +lemma tmsubst0_nb: "tmbound0 t \ tmbound0 (tmsubst0 t a)" +by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2) + +consts tmsubst:: "nat \ tm \ tm \ tm" + +primrec + "tmsubst n t (CP c) = CP c" + "tmsubst n t (Bound m) = (if n=m then t else Bound m)" + "tmsubst n t (CNP m c a) = (if n=m then Add (Mul c t) (tmsubst n t a) + else CNP m c (tmsubst n t a))" + "tmsubst n t (Neg a) = Neg (tmsubst n t a)" + "tmsubst n t (Add a b) = Add (tmsubst n t a) (tmsubst n t b)" + "tmsubst n t (Sub a b) = Sub (tmsubst n t a) (tmsubst n t b)" + "tmsubst n t (Mul i a) = Mul i (tmsubst n t a)" + +lemma tmsubst: assumes nb: "tmboundslt (length bs) a" and nlt: "n \ length bs" + shows "Itm vs bs (tmsubst n t a) = Itm vs (bs[n:= Itm vs bs t]) a" +using nb nlt +by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2) + +lemma tmsubst_nb0: assumes tnb: "tmbound0 t" +shows "tmbound0 (tmsubst 0 t a)" +using tnb +by (induct a rule: tmsubst.induct, auto) + +lemma tmsubst_nb: assumes tnb: "tmbound m t" +shows "tmbound m (tmsubst m t a)" +using tnb +by (induct a rule: tmsubst.induct, auto) +lemma incrtm0_tmbound: "tmbound n t \ tmbound (Suc n) (incrtm0 t)" + by (induct t, auto) + (* Simplification *) + +consts + simptm:: "tm \ tm" + tmadd:: "tm \ tm \ tm" + tmmul:: "tm \ poly \ tm" +recdef tmadd "measure (\ (t,s). size t + size s)" + "tmadd (CNP n1 c1 r1,CNP n2 c2 r2) = + (if n1=n2 then + (let c = c1 +\<^sub>p c2 + in if c = 0\<^sub>p then tmadd(r1,r2) else CNP n1 c (tmadd (r1,r2))) + else if n1 \ n2 then (CNP n1 c1 (tmadd (r1,CNP n2 c2 r2))) + else (CNP n2 c2 (tmadd (CNP n1 c1 r1,r2))))" + "tmadd (CNP n1 c1 r1,t) = CNP n1 c1 (tmadd (r1, t))" + "tmadd (t,CNP n2 c2 r2) = CNP n2 c2 (tmadd (t,r2))" + "tmadd (CP b1, CP b2) = CP (b1 +\<^sub>p b2)" + "tmadd (a,b) = Add a b" + +lemma tmadd[simp]: "Itm vs bs (tmadd (t,s)) = Itm vs bs (Add t s)" +apply (induct t s rule: tmadd.induct, simp_all add: Let_def) +apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p",case_tac "n1 \ n2", simp_all) +apply (case_tac "n1 = n2", simp_all add: ring_simps) +apply (simp only: right_distrib[symmetric]) +by (auto simp del: polyadd simp add: polyadd[symmetric]) + +lemma tmadd_nb0[simp]: "\ tmbound0 t ; tmbound0 s\ \ tmbound0 (tmadd (t,s))" +by (induct t s rule: tmadd.induct, auto simp add: Let_def) + +lemma tmadd_nb[simp]: "\ tmbound n t ; tmbound n s\ \ tmbound n (tmadd (t,s))" +by (induct t s rule: tmadd.induct, auto simp add: Let_def) +lemma tmadd_blt[simp]: "\tmboundslt n t ; tmboundslt n s\ \ tmboundslt n (tmadd (t,s))" +by (induct t s rule: tmadd.induct, auto simp add: Let_def) + +lemma tmadd_allpolys_npoly[simp]: "allpolys isnpoly t \ allpolys isnpoly s \ allpolys isnpoly (tmadd(t,s))" by (induct t s rule: tmadd.induct, simp_all add: Let_def polyadd_norm) + +recdef tmmul "measure size" + "tmmul (CP j) = (\ i. CP (i *\<^sub>p j))" + "tmmul (CNP n c a) = (\ i. CNP n (i *\<^sub>p c) (tmmul a i))" + "tmmul t = (\ i. Mul i t)" + +lemma tmmul[simp]: "Itm vs bs (tmmul t i) = Itm vs bs (Mul i t)" +by (induct t arbitrary: i rule: tmmul.induct, simp_all add: ring_simps) + +lemma tmmul_nb0[simp]: "tmbound0 t \ tmbound0 (tmmul t i)" +by (induct t arbitrary: i rule: tmmul.induct, auto ) + +lemma tmmul_nb[simp]: "tmbound n t \ tmbound n (tmmul t i)" +by (induct t arbitrary: n rule: tmmul.induct, auto ) +lemma tmmul_blt[simp]: "tmboundslt n t \ tmboundslt n (tmmul t i)" +by (induct t arbitrary: i rule: tmmul.induct, auto simp add: Let_def) + +lemma tmmul_allpolys_npoly[simp]: + assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" + shows "allpolys isnpoly t \ isnpoly c \ allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm) + +constdefs tmneg :: "tm \ tm" + "tmneg t \ tmmul t (C (- 1,1))" + +constdefs tmsub :: "tm \ tm \ tm" + "tmsub s t \ (if s = t then CP 0\<^sub>p else tmadd (s,tmneg t))" + +lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)" +using tmneg_def[of t] +apply simp +apply (subst number_of_Min) +apply (simp only: of_int_minus) +apply simp +done + +lemma tmneg_nb0[simp]: "tmbound0 t \ tmbound0 (tmneg t)" +using tmneg_def by simp + +lemma tmneg_nb[simp]: "tmbound n t \ tmbound n (tmneg t)" +using tmneg_def by simp +lemma tmneg_blt[simp]: "tmboundslt n t \ tmboundslt n (tmneg t)" +using tmneg_def by simp +lemma [simp]: "isnpoly (C (-1,1))" unfolding isnpoly_def by simp +lemma tmneg_allpolys_npoly[simp]: + assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" + shows "allpolys isnpoly t \ allpolys isnpoly (tmneg t)" + unfolding tmneg_def by auto + +lemma tmsub[simp]: "Itm vs bs (tmsub a b) = Itm vs bs (Sub a b)" +using tmsub_def by simp + +lemma tmsub_nb0[simp]: "\ tmbound0 t ; tmbound0 s\ \ tmbound0 (tmsub t s)" +using tmsub_def by simp +lemma tmsub_nb[simp]: "\ tmbound n t ; tmbound n s\ \ tmbound n (tmsub t s)" +using tmsub_def by simp +lemma tmsub_blt[simp]: "\tmboundslt n t ; tmboundslt n s\ \ tmboundslt n (tmsub t s )" +using tmsub_def by simp +lemma tmsub_allpolys_npoly[simp]: + assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" + shows "allpolys isnpoly t \ allpolys isnpoly s \ allpolys isnpoly (tmsub t s)" + unfolding tmsub_def by (simp add: isnpoly_def) + +recdef simptm "measure size" + "simptm (CP j) = CP (polynate j)" + "simptm (Bound n) = CNP n 1\<^sub>p (CP 0\<^sub>p)" + "simptm (Neg t) = tmneg (simptm t)" + "simptm (Add t s) = tmadd (simptm t,simptm s)" + "simptm (Sub t s) = tmsub (simptm t) (simptm s)" + "simptm (Mul i t) = (let i' = polynate i in if i' = 0\<^sub>p then CP 0\<^sub>p else tmmul (simptm t) i')" + "simptm (CNP n c t) = (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p ), simptm t))" + +lemma polynate_stupid: + assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" + shows "polynate t = 0\<^sub>p \ Ipoly bs t = (0::'a::{ring_char_0,division_by_zero, field})" +apply (subst polynate[symmetric]) +apply simp +done + +lemma simptm_ci[simp]: "Itm vs bs (simptm t) = Itm vs bs t" +by (induct t rule: simptm.induct, auto simp add: tmneg tmadd tmsub tmmul Let_def polynate_stupid) + +lemma simptm_tmbound0[simp]: + "tmbound0 t \ tmbound0 (simptm t)" +by (induct t rule: simptm.induct, auto simp add: Let_def) + +lemma simptm_nb[simp]: "tmbound n t \ tmbound n (simptm t)" +by (induct t rule: simptm.induct, auto simp add: Let_def) +lemma simptm_nlt[simp]: "tmboundslt n t \ tmboundslt n (simptm t)" +by (induct t rule: simptm.induct, auto simp add: Let_def) + +lemma [simp]: "isnpoly 0\<^sub>p" and [simp]: "isnpoly (C(1,1))" + by (simp_all add: isnpoly_def) +lemma simptm_allpolys_npoly[simp]: + assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" + shows "allpolys isnpoly (simptm p)" + by (induct p rule: simptm.induct, auto simp add: Let_def) + +consts split0 :: "tm \ (poly \ tm)" +recdef split0 "measure tmsize" + "split0 (Bound 0) = (1\<^sub>p, CP 0\<^sub>p)" + "split0 (CNP 0 c t) = (let (c',t') = split0 t in (c +\<^sub>p c',t'))" + "split0 (Neg t) = (let (c,t') = split0 t in (~\<^sub>p c,Neg t'))" + "split0 (CNP n c t) = (let (c',t') = split0 t in (c',CNP n c t'))" + "split0 (Add s t) = (let (c1,s') = split0 s ; (c2,t') = split0 t in (c1 +\<^sub>p c2, Add s' t'))" + "split0 (Sub s t) = (let (c1,s') = split0 s ; (c2,t') = split0 t in (c1 -\<^sub>p c2, Sub s' t'))" + "split0 (Mul c t) = (let (c',t') = split0 t in (c *\<^sub>p c', Mul c t'))" + "split0 t = (0\<^sub>p, t)" + +lemma split0_stupid[simp]: "\x y. (x,y) = split0 p" + apply (rule exI[where x="fst (split0 p)"]) + apply (rule exI[where x="snd (split0 p)"]) + by simp + +lemma split0: + "tmbound 0 (snd (split0 t)) \ (Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t)" + apply (induct t rule: split0.induct) + apply simp + apply (simp add: Let_def split_def ring_simps) + apply (simp add: Let_def split_def ring_simps) + apply (simp add: Let_def split_def ring_simps) + apply (simp add: Let_def split_def ring_simps) + apply (simp add: Let_def split_def ring_simps) + apply (simp add: Let_def split_def mult_assoc right_distrib[symmetric]) + apply (simp add: Let_def split_def ring_simps) + apply (simp add: Let_def split_def ring_simps) + done + +lemma split0_ci: "split0 t = (c',t') \ Itm vs bs t = Itm vs bs (CNP 0 c' t')" +proof- + fix c' t' + assume "split0 t = (c', t')" hence "c' = fst (split0 t)" and "t' = snd (split0 t)" by auto + with split0[where t="t" and bs="bs"] show "Itm vs bs t = Itm vs bs (CNP 0 c' t')" by simp +qed + +lemma split0_nb0: + assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" + shows "split0 t = (c',t') \ tmbound 0 t'" +proof- + fix c' t' + assume "split0 t = (c', t')" hence "c' = fst (split0 t)" and "t' = snd (split0 t)" by auto + with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'" by simp +qed + +lemma split0_nb0'[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" + shows "tmbound0 (snd (split0 t))" + using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"] by (simp add: tmbound0_tmbound_iff) + + +lemma split0_nb: assumes nb:"tmbound n t" shows "tmbound n (snd (split0 t))" + using nb by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid) + +lemma split0_blt: assumes nb:"tmboundslt n t" shows "tmboundslt n (snd (split0 t))" + using nb by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid) + +lemma tmbound_split0: "tmbound 0 t \ Ipoly vs (fst(split0 t)) = 0" + by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid) + +lemma tmboundslt_split0: "tmboundslt n t \ Ipoly vs (fst(split0 t)) = 0 \ n > 0" +by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid) + +lemma tmboundslt0_split0: "tmboundslt 0 t \ Ipoly vs (fst(split0 t)) = 0" + by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid) + +lemma allpolys_split0: "allpolys isnpoly p \ allpolys isnpoly (snd (split0 p))" +by (induct p rule: split0.induct, auto simp add: isnpoly_def Let_def split_def split0_stupid) + +lemma isnpoly_fst_split0: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" + shows + "allpolys isnpoly p \ isnpoly (fst (split0 p))" + by (induct p rule: split0.induct, + auto simp add: polyadd_norm polysub_norm polyneg_norm polymul_norm + Let_def split_def split0_stupid) + +subsection{* Formulae *} + +datatype fm = T| F| Le tm | Lt tm | Eq tm | NEq tm| + 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 ::"'a::{division_by_zero,ordered_field} list \ 'a list \ fm \ bool" +primrec + "Ifm vs bs T = True" + "Ifm vs bs F = False" + "Ifm vs bs (Lt a) = (Itm vs bs a < 0)" + "Ifm vs bs (Le a) = (Itm vs bs a \ 0)" + "Ifm vs bs (Eq a) = (Itm vs bs a = 0)" + "Ifm vs bs (NEq a) = (Itm vs bs a \ 0)" + "Ifm vs bs (NOT p) = (\ (Ifm vs bs p))" + "Ifm vs bs (And p q) = (Ifm vs bs p \ Ifm vs bs q)" + "Ifm vs bs (Or p q) = (Ifm vs bs p \ Ifm vs bs q)" + "Ifm vs bs (Imp p q) = ((Ifm vs bs p) \ (Ifm vs bs q))" + "Ifm vs bs (Iff p q) = (Ifm vs bs p = Ifm vs bs q)" + "Ifm vs bs (E p) = (\ x. Ifm vs (x#bs) p)" + "Ifm vs bs (A p) = (\ x. Ifm vs (x#bs) p)" + +consts not:: "fm \ fm" +recdef not "measure size" + "not (NOT (NOT p)) = not p" + "not (NOT p) = p" + "not T = F" + "not F = T" + "not (Lt t) = Le (tmneg t)" + "not (Le t) = Lt (tmneg t)" + "not (Eq t) = NEq t" + "not (NEq t) = Eq t" + "not p = NOT p" +lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)" +by (induct p rule: not.induct) 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 vs bs (conj p q) = Ifm vs 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 vs bs (disj p q) = Ifm vs 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 vs bs (imp p q) = Ifm vs 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 vs bs (iff p q) = Ifm vs bs (Iff p q)" + by (unfold iff_def,cases "p=q", simp,cases "p=NOT q", simp) (cases "NOT p= q", 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 boundslt :: "nat \ fm \ bool" +primrec + "boundslt n T = True" + "boundslt n F = True" + "boundslt n (Lt t) = (tmboundslt n t)" + "boundslt n (Le t) = (tmboundslt n t)" + "boundslt n (Eq t) = (tmboundslt n t)" + "boundslt n (NEq t) = (tmboundslt n t)" + "boundslt n (NOT p) = boundslt n p" + "boundslt n (And p q) = (boundslt n p \ boundslt n q)" + "boundslt n (Or p q) = (boundslt n p \ boundslt n q)" + "boundslt n (Imp p q) = ((boundslt n p) \ (boundslt n q))" + "boundslt n (Iff p q) = (boundslt n p \ boundslt n q)" + "boundslt n (E p) = boundslt (Suc n) p" + "boundslt n (A p) = boundslt (Suc n) p" + +consts + bound0:: "fm \ bool" (* A Formula is independent of Bound 0 *) + bound:: "nat \ fm \ bool" (* A Formula is independent of Bound n *) + decr0 :: "fm \ fm" + decr :: "nat \ fm \ fm" +recdef bound0 "measure size" + "bound0 T = True" + "bound0 F = True" + "bound0 (Lt a) = tmbound0 a" + "bound0 (Le a) = tmbound0 a" + "bound0 (Eq a) = tmbound0 a" + "bound0 (NEq a) = tmbound0 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 p = False" +lemma bound0_I: + assumes bp: "bound0 p" + shows "Ifm vs (b#bs) p = Ifm vs (b'#bs) p" +using bp tmbound0_I[where b="b" and bs="bs" and b'="b'"] +by (induct p rule: bound0.induct,auto simp add: nth_pos2) + +primrec + "bound m T = True" + "bound m F = True" + "bound m (Lt t) = tmbound m t" + "bound m (Le t) = tmbound m t" + "bound m (Eq t) = tmbound m t" + "bound m (NEq t) = tmbound m t" + "bound m (NOT p) = bound m p" + "bound m (And p q) = (bound m p \ bound m q)" + "bound m (Or p q) = (bound m p \ bound m q)" + "bound m (Imp p q) = ((bound m p) \ (bound m q))" + "bound m (Iff p q) = (bound m p \ bound m q)" + "bound m (E p) = bound (Suc m) p" + "bound m (A p) = bound (Suc m) p" + +lemma bound_I: + assumes bnd: "boundslt (length bs) p" and nb: "bound n p" and le: "n \ length bs" + shows "Ifm vs (bs[n:=x]) p = Ifm vs bs p" + using bnd nb le tmbound_I[where bs=bs and vs = vs] +proof(induct p arbitrary: bs n rule: bound.induct) + case (E p bs n) + {fix y + from prems have bnd: "boundslt (length (y#bs)) p" + and nb: "bound (Suc n) p" and le: "Suc n \ length (y#bs)" by simp+ + from E.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" . } + thus ?case by simp +next + case (A p bs n) {fix y + from prems have bnd: "boundslt (length (y#bs)) p" + and nb: "bound (Suc n) p" and le: "Suc n \ length (y#bs)" by simp+ + from A.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" . } + thus ?case by simp +qed auto + +recdef decr0 "measure size" + "decr0 (Lt a) = Lt (decrtm0 a)" + "decr0 (Le a) = Le (decrtm0 a)" + "decr0 (Eq a) = Eq (decrtm0 a)" + "decr0 (NEq a) = NEq (decrtm0 a)" + "decr0 (NOT p) = NOT (decr0 p)" + "decr0 (And p q) = conj (decr0 p) (decr0 q)" + "decr0 (Or p q) = disj (decr0 p) (decr0 q)" + "decr0 (Imp p q) = imp (decr0 p) (decr0 q)" + "decr0 (Iff p q) = iff (decr0 p) (decr0 q)" + "decr0 p = p" + +lemma decr0: assumes nb: "bound0 p" + shows "Ifm vs (x#bs) p = Ifm vs bs (decr0 p)" + using nb + by (induct p rule: decr0.induct, simp_all add: decrtm0) + +primrec + "decr m T = T" + "decr m F = F" + "decr m (Lt t) = (Lt (decrtm m t))" + "decr m (Le t) = (Le (decrtm m t))" + "decr m (Eq t) = (Eq (decrtm m t))" + "decr m (NEq t) = (NEq (decrtm m t))" + "decr m (NOT p) = NOT (decr m p)" + "decr m (And p q) = conj (decr m p) (decr m q)" + "decr m (Or p q) = disj (decr m p) (decr m q)" + "decr m (Imp p q) = imp (decr m p) (decr m q)" + "decr m (Iff p q) = iff (decr m p) (decr m q)" + "decr m (E p) = E (decr (Suc m) p)" + "decr m (A p) = A (decr (Suc m) p)" + +lemma decr: assumes bnd: "boundslt (length bs) p" and nb: "bound m p" + and nle: "m < length bs" + shows "Ifm vs (removen m bs) (decr m p) = Ifm vs bs p" + using bnd nb nle +proof(induct p arbitrary: bs m rule: decr.induct) + case (E p bs m) + {fix x + from prems have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p" + and nle: "Suc m < length (x#bs)" by auto + from prems(4)[OF bnd nb nle] have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p". + } thus ?case by auto +next + case (A p bs m) + {fix x + from prems have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p" + and nle: "Suc m < length (x#bs)" by auto + from prems(4)[OF bnd nb nle] have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p". + } thus ?case by auto +qed (auto simp add: decrtm removen_nth) + +consts + subst0:: "tm \ fm \ fm" + +primrec + "subst0 t T = T" + "subst0 t F = F" + "subst0 t (Lt a) = Lt (tmsubst0 t a)" + "subst0 t (Le a) = Le (tmsubst0 t a)" + "subst0 t (Eq a) = Eq (tmsubst0 t a)" + "subst0 t (NEq a) = NEq (tmsubst0 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 (E p) = E p" + "subst0 t (A p) = A p" + +lemma subst0: assumes qf: "qfree p" + shows "Ifm vs (x#bs) (subst0 t p) = Ifm vs ((Itm vs (x#bs) t)#bs) p" +using qf tmsubst0[where x="x" and bs="bs" and t="t"] +by (induct p rule: subst0.induct, auto) + +lemma subst0_nb: + assumes bp: "tmbound0 t" and qf: "qfree p" + shows "bound0 (subst0 t p)" +using qf tmsubst0_nb[OF bp] bp +by (induct p rule: subst0.induct, auto) + +consts subst:: "nat \ tm \ fm \ fm" +primrec + "subst n t T = T" + "subst n t F = F" + "subst n t (Lt a) = Lt (tmsubst n t a)" + "subst n t (Le a) = Le (tmsubst n t a)" + "subst n t (Eq a) = Eq (tmsubst n t a)" + "subst n t (NEq a) = NEq (tmsubst n t a)" + "subst n t (NOT p) = NOT (subst n t p)" + "subst n t (And p q) = And (subst n t p) (subst n t q)" + "subst n t (Or p q) = Or (subst n t p) (subst n t q)" + "subst n t (Imp p q) = Imp (subst n t p) (subst n t q)" + "subst n t (Iff p q) = Iff (subst n t p) (subst n t q)" + "subst n t (E p) = E (subst (Suc n) (incrtm0 t) p)" + "subst n t (A p) = A (subst (Suc n) (incrtm0 t) p)" + +lemma subst: assumes nb: "boundslt (length bs) p" and nlm: "n \ length bs" + shows "Ifm vs bs (subst n t p) = Ifm vs (bs[n:= Itm vs bs t]) p" + using nb nlm +proof (induct p arbitrary: bs n t rule: subst0.induct) + case (E p bs n) + {fix x + from prems have bn: "boundslt (length (x#bs)) p" by simp + from prems have nlm: "Suc n \ length (x#bs)" by simp + from prems(3)[OF bn nlm] have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p" by simp + hence "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p" + by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) } +thus ?case by simp +next + case (A p bs n) + {fix x + from prems have bn: "boundslt (length (x#bs)) p" by simp + from prems have nlm: "Suc n \ length (x#bs)" by simp + from prems(3)[OF bn nlm] have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p" by simp + hence "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p" + by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) } +thus ?case by simp +qed(auto simp add: tmsubst) + +lemma subst_nb: assumes tnb: "tmbound m t" +shows "bound m (subst m t p)" +using tnb tmsubst_nb incrtm0_tmbound +by (induct p arbitrary: m t rule: subst.induct, auto) + +lemma not_qf[simp]: "qfree p \ qfree (not p)" +by (induct p rule: not.induct, auto) +lemma not_bn0[simp]: "bound0 p \ bound0 (not p)" +by (induct p rule: not.induct, auto) +lemma not_nb[simp]: "bound n p \ bound n (not p)" +by (induct p rule: not.induct, auto) +lemma not_blt[simp]: "boundslt n p \ boundslt n (not p)" + by (induct p rule: not.induct, auto) + +lemma conj_qf[simp]: "\qfree p ; qfree q\ \ qfree (conj p q)" +using conj_def by auto +lemma conj_nb0[simp]: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" +using conj_def by auto +lemma conj_nb[simp]: "\bound n p ; bound n q\ \ bound n (conj p q)" +using conj_def by auto +lemma conj_blt[simp]: "boundslt n p \ boundslt n q \ boundslt n (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_nb0[simp]: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" +using disj_def by auto +lemma disj_nb[simp]: "\bound n p ; bound n q\ \ bound n (disj p q)" +using disj_def by auto +lemma disj_blt[simp]: "boundslt n p \ boundslt n q \ boundslt n (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_nb0[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 imp_nb[simp]: "\bound n p ; bound n q\ \ bound n (imp p q)" +using imp_def by (cases "p=F \ q=T \ p=q",simp_all add: imp_def) +lemma imp_blt[simp]: "boundslt n p \ boundslt n q \ boundslt n (imp p q)" +using imp_def by auto + +lemma iff_qf[simp]: "\qfree p ; qfree q\ \ qfree (iff p q)" + by (unfold iff_def,cases "p=q", auto) +lemma iff_nb0[simp]: "\bound0 p ; bound0 q\ \ bound0 (iff p q)" +using iff_def by (unfold iff_def,cases "p=q", auto) +lemma iff_nb[simp]: "\bound n p ; bound n q\ \ bound n (iff p q)" +using iff_def by (unfold iff_def,cases "p=q", auto) +lemma iff_blt[simp]: "boundslt n p \ boundslt n q \ boundslt n (iff p q)" +using iff_def by auto +lemma decr0_qf: "bound0 p \ qfree (decr0 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 (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 vs bs (djf f p q) = Ifm vs 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 vs bs (evaldjf f ps) = (\ p \ set ps. Ifm vs 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 vs bs q) = Ifm vs 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 vs bs (f (Or p q)) = Ifm vs bs (Or (f p) (f q))" + and fF: "f F = F" + shows "Ifm vs bs (DJ f p) = Ifm vs bs (f p)" +proof- + have "Ifm vs bs (DJ f p) = (\ q \ set (disjuncts p). Ifm vs bs (f q))" + by (simp add: DJ_def evaldjf_ex) + also have "\ = Ifm vs 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 vs bs (qe p) = Ifm vs bs (E p))" + shows "\ bs p. qfree p \ qfree (DJ qe p) \ (Ifm vs bs ((DJ qe p)) = Ifm vs 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 vs bs (DJ qe p) = (\ q\ set (disjuncts p). Ifm vs bs (qe q))" + by (simp add: DJ_def evaldjf_ex) + also have "\ = (\ q \ set(disjuncts p). Ifm vs bs (E q))" using qe disjuncts_qf[OF qf] by auto + also have "\ = Ifm vs bs (E p)" by (induct p rule: disjuncts.induct, auto) + finally show "qfree (DJ qe p) \ Ifm vs bs (DJ qe p) = Ifm vs bs (E p)" using qfth by blast +qed + +consts conjuncts :: "fm \ fm list" + +recdef conjuncts "measure size" + "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)" + "conjuncts T = []" + "conjuncts p = [p]" + +constdefs list_conj :: "fm list \ fm" + "list_conj ps \ foldr conj ps T" + +constdefs CJNB:: "(fm \ fm) \ fm \ fm" + "CJNB f p \ (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs + in conj (decr0 (list_conj yes)) (f (list_conj no)))" + +lemma conjuncts_qf: "qfree p \ \ q\ set (conjuncts p). qfree q" +proof- + assume qf: "qfree p" + hence "list_all qfree (conjuncts p)" + by (induct p rule: conjuncts.induct, auto) + thus ?thesis by (simp only: list_all_iff) +qed + +lemma conjuncts: "(\ q\ set (conjuncts p). Ifm vs bs q) = Ifm vs bs p" +by(induct p rule: conjuncts.induct, auto) + +lemma conjuncts_nb: "bound0 p \ \ q\ set (conjuncts p). bound0 q" +proof- + assume nb: "bound0 p" + hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto) + thus ?thesis by (simp only: list_all_iff) +qed + +fun islin :: "fm \ bool" where + "islin (And p q) = (islin p \ islin q \ p \ T \ p \ F \ q \ T \ q \ F)" +| "islin (Or p q) = (islin p \ islin q \ p \ T \ p \ F \ q \ T \ q \ F)" +| "islin (Eq (CNP 0 c s)) = (isnpoly c \ c \ 0\<^sub>p \ tmbound0 s \ allpolys isnpoly s)" +| "islin (NEq (CNP 0 c s)) = (isnpoly c \ c \ 0\<^sub>p \ tmbound0 s \ allpolys isnpoly s)" +| "islin (Lt (CNP 0 c s)) = (isnpoly c \ c \ 0\<^sub>p \ tmbound0 s \ allpolys isnpoly s)" +| "islin (Le (CNP 0 c s)) = (isnpoly c \ c \ 0\<^sub>p \ tmbound0 s \ allpolys isnpoly s)" +| "islin (NOT p) = False" +| "islin (Imp p q) = False" +| "islin (Iff p q) = False" +| "islin p = bound0 p" + +lemma islin_stupid: assumes nb: "tmbound0 p" + shows "islin (Lt p)" and "islin (Le p)" and "islin (Eq p)" and "islin (NEq p)" + using nb by (cases p, auto, case_tac nat, auto)+ + +definition "lt p = (case p of CP (C c) \ if 0>\<^sub>N c then T else F| _ \ Lt p)" +definition "le p = (case p of CP (C c) \ if 0\\<^sub>N c then T else F | _ \ Le p)" +definition "eq p = (case p of CP (C c) \ if c = 0\<^sub>N then T else F | _ \ Eq p)" +definition "neq p = not (eq p)" + +lemma lt: "allpolys isnpoly p \ Ifm vs bs (lt p) = Ifm vs bs (Lt p)" + apply(simp add: lt_def) + apply(cases p, simp_all) + apply (case_tac poly, simp_all add: isnpoly_def) + done + +lemma le: "allpolys isnpoly p \ Ifm vs bs (le p) = Ifm vs bs (Le p)" + apply(simp add: le_def) + apply(cases p, simp_all) + apply (case_tac poly, simp_all add: isnpoly_def) + done + +lemma eq: "allpolys isnpoly p \ Ifm vs bs (eq p) = Ifm vs bs (Eq p)" + apply(simp add: eq_def) + apply(cases p, simp_all) + apply (case_tac poly, simp_all add: isnpoly_def) + done + +lemma neq: "allpolys isnpoly p \ Ifm vs bs (neq p) = Ifm vs bs (NEq p)" + by(simp add: neq_def eq) + +lemma lt_lin: "tmbound0 p \ islin (lt p)" + apply (simp add: lt_def) + apply (cases p, simp_all) + apply (case_tac poly, simp_all) + apply (case_tac nat, simp_all) + done + +lemma le_lin: "tmbound0 p \ islin (le p)" + apply (simp add: le_def) + apply (cases p, simp_all) + apply (case_tac poly, simp_all) + apply (case_tac nat, simp_all) + done + +lemma eq_lin: "tmbound0 p \ islin (eq p)" + apply (simp add: eq_def) + apply (cases p, simp_all) + apply (case_tac poly, simp_all) + apply (case_tac nat, simp_all) + done + +lemma neq_lin: "tmbound0 p \ islin (neq p)" + apply (simp add: neq_def eq_def) + apply (cases p, simp_all) + apply (case_tac poly, simp_all) + apply (case_tac nat, simp_all) + done + +definition "simplt t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then lt s else Lt (CNP 0 c s))" +definition "simple t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then le s else Le (CNP 0 c s))" +definition "simpeq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then eq s else Eq (CNP 0 c s))" +definition "simpneq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then neq s else NEq (CNP 0 c s))" + +lemma simplt_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "islin (simplt t)" + unfolding simplt_def + using split0_nb0' +by (auto simp add: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly]) + +lemma simple_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "islin (simple t)" + unfolding simple_def + using split0_nb0' +by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] le_lin) +lemma simpeq_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "islin (simpeq t)" + unfolding simpeq_def + using split0_nb0' +by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] eq_lin) + +lemma simpneq_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "islin (simpneq t)" + unfolding simpneq_def + using split0_nb0' +by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] neq_lin) + +lemma really_stupid: "\ (\c1 s'. (c1, s') \ split0 s)" + by (cases "split0 s", auto) +lemma split0_npoly: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + and n: "allpolys isnpoly t" + shows "isnpoly (fst (split0 t))" and "allpolys isnpoly (snd (split0 t))" + using n + by (induct t rule: split0.induct, auto simp add: Let_def split_def polyadd_norm polymul_norm polyneg_norm polysub_norm really_stupid) +lemma simplt[simp]: + shows "Ifm vs bs (simplt t) = Ifm vs bs (Lt t)" +proof- + have n: "allpolys isnpoly (simptm t)" by simp + let ?t = "simptm t" + {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis + using split0[of "simptm t" vs bs] lt[OF split0_npoly(2)[OF n], of vs bs] + by (simp add: simplt_def Let_def split_def lt)} + moreover + {assume "fst (split0 ?t) \ 0\<^sub>p" + hence ?thesis using split0[of "simptm t" vs bs] by (simp add: simplt_def Let_def split_def) + } + ultimately show ?thesis by blast +qed + +lemma simple[simp]: + shows "Ifm vs bs (simple t) = Ifm vs bs (Le t)" +proof- + have n: "allpolys isnpoly (simptm t)" by simp + let ?t = "simptm t" + {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis + using split0[of "simptm t" vs bs] le[OF split0_npoly(2)[OF n], of vs bs] + by (simp add: simple_def Let_def split_def le)} + moreover + {assume "fst (split0 ?t) \ 0\<^sub>p" + hence ?thesis using split0[of "simptm t" vs bs] by (simp add: simple_def Let_def split_def) + } + ultimately show ?thesis by blast +qed + +lemma simpeq[simp]: + shows "Ifm vs bs (simpeq t) = Ifm vs bs (Eq t)" +proof- + have n: "allpolys isnpoly (simptm t)" by simp + let ?t = "simptm t" + {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis + using split0[of "simptm t" vs bs] eq[OF split0_npoly(2)[OF n], of vs bs] + by (simp add: simpeq_def Let_def split_def)} + moreover + {assume "fst (split0 ?t) \ 0\<^sub>p" + hence ?thesis using split0[of "simptm t" vs bs] by (simp add: simpeq_def Let_def split_def) + } + ultimately show ?thesis by blast +qed + +lemma simpneq[simp]: + shows "Ifm vs bs (simpneq t) = Ifm vs bs (NEq t)" +proof- + have n: "allpolys isnpoly (simptm t)" by simp + let ?t = "simptm t" + {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis + using split0[of "simptm t" vs bs] neq[OF split0_npoly(2)[OF n], of vs bs] + by (simp add: simpneq_def Let_def split_def )} + moreover + {assume "fst (split0 ?t) \ 0\<^sub>p" + hence ?thesis using split0[of "simptm t" vs bs] by (simp add: simpneq_def Let_def split_def) + } + ultimately show ?thesis by blast +qed + +lemma lt_nb: "tmbound0 t \ bound0 (lt t)" + apply (simp add: lt_def) + apply (cases t, auto) + apply (case_tac poly, auto) + done + +lemma le_nb: "tmbound0 t \ bound0 (le t)" + apply (simp add: le_def) + apply (cases t, auto) + apply (case_tac poly, auto) + done + +lemma eq_nb: "tmbound0 t \ bound0 (eq t)" + apply (simp add: eq_def) + apply (cases t, auto) + apply (case_tac poly, auto) + done + +lemma neq_nb: "tmbound0 t \ bound0 (neq t)" + apply (simp add: neq_def eq_def) + apply (cases t, auto) + apply (case_tac poly, auto) + done + +lemma simplt_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "tmbound0 t \ bound0 (simplt t)" + using split0 [of "simptm t" vs bs] +proof(simp add: simplt_def Let_def split_def) + assume nb: "tmbound0 t" + hence nb': "tmbound0 (simptm t)" by simp + let ?c = "fst (split0 (simptm t))" + from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]] + have th: "\bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto + from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]] + have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def) + from iffD1[OF isnpolyh_unique[OF ths] th] + have "fst (split0 (simptm t)) = 0\<^sub>p" . + thus "(fst (split0 (simptm t)) = 0\<^sub>p \ bound0 (lt (snd (split0 (simptm t))))) \ + fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def lt_nb) +qed + +lemma simple_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "tmbound0 t \ bound0 (simple t)" + using split0 [of "simptm t" vs bs] +proof(simp add: simple_def Let_def split_def) + assume nb: "tmbound0 t" + hence nb': "tmbound0 (simptm t)" by simp + let ?c = "fst (split0 (simptm t))" + from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]] + have th: "\bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto + from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]] + have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def) + from iffD1[OF isnpolyh_unique[OF ths] th] + have "fst (split0 (simptm t)) = 0\<^sub>p" . + thus "(fst (split0 (simptm t)) = 0\<^sub>p \ bound0 (le (snd (split0 (simptm t))))) \ + fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def le_nb) +qed + +lemma simpeq_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "tmbound0 t \ bound0 (simpeq t)" + using split0 [of "simptm t" vs bs] +proof(simp add: simpeq_def Let_def split_def) + assume nb: "tmbound0 t" + hence nb': "tmbound0 (simptm t)" by simp + let ?c = "fst (split0 (simptm t))" + from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]] + have th: "\bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto + from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]] + have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def) + from iffD1[OF isnpolyh_unique[OF ths] th] + have "fst (split0 (simptm t)) = 0\<^sub>p" . + thus "(fst (split0 (simptm t)) = 0\<^sub>p \ bound0 (eq (snd (split0 (simptm t))))) \ + fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpeq_def Let_def split_def eq_nb) +qed + +lemma simpneq_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "tmbound0 t \ bound0 (simpneq t)" + using split0 [of "simptm t" vs bs] +proof(simp add: simpneq_def Let_def split_def) + assume nb: "tmbound0 t" + hence nb': "tmbound0 (simptm t)" by simp + let ?c = "fst (split0 (simptm t))" + from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]] + have th: "\bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto + from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]] + have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def) + from iffD1[OF isnpolyh_unique[OF ths] th] + have "fst (split0 (simptm t)) = 0\<^sub>p" . + thus "(fst (split0 (simptm t)) = 0\<^sub>p \ bound0 (neq (snd (split0 (simptm t))))) \ + fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpneq_def Let_def split_def neq_nb) +qed + +consts conjs :: "fm \ fm list" +recdef conjs "measure size" + "conjs (And p q) = (conjs p)@(conjs q)" + "conjs T = []" + "conjs p = [p]" +lemma conjs_ci: "(\ q \ set (conjs p). Ifm vs bs q) = Ifm vs bs p" +by (induct p rule: conjs.induct, auto) +constdefs list_disj :: "fm list \ fm" + "list_disj ps \ foldr disj ps F" + +lemma list_conj: "Ifm vs bs (list_conj ps) = (\p\ set ps. Ifm vs bs p)" + by (induct ps, auto simp add: list_conj_def) +lemma list_conj_qf: " \p\ set ps. qfree p \ qfree (list_conj ps)" + by (induct ps, auto simp add: list_conj_def conj_qf) +lemma list_disj: "Ifm vs bs (list_disj ps) = (\p\ set ps. Ifm vs bs p)" + by (induct ps, auto simp add: list_disj_def) + +lemma conj_boundslt: "boundslt n p \ boundslt n q \ boundslt n (conj p q)" + unfolding conj_def by auto + +lemma conjs_nb: "bound n p \ \q\ set (conjs p). bound n q" + apply (induct p rule: conjs.induct) + apply (unfold conjs.simps) + apply (unfold set_append) + apply (unfold ball_Un) + apply (unfold bound.simps) + apply auto + done + +lemma conjs_boundslt: "boundslt n p \ \q\ set (conjs p). boundslt n q" + apply (induct p rule: conjs.induct) + apply (unfold conjs.simps) + apply (unfold set_append) + apply (unfold ball_Un) + apply (unfold boundslt.simps) + apply blast +by simp_all + +lemma list_conj_boundslt: " \p\ set ps. boundslt n p \ boundslt n (list_conj ps)" + unfolding list_conj_def + by (induct ps, auto simp add: conj_boundslt) + +lemma list_conj_nb: assumes bnd: "\p\ set ps. bound n p" + shows "bound n (list_conj ps)" + using bnd + unfolding list_conj_def + by (induct ps, auto simp add: conj_nb) + +lemma list_conj_nb': "\p\set ps. bound0 p \ bound0 (list_conj ps)" +unfolding list_conj_def by (induct ps , auto) + +lemma CJNB_qe: + assumes qe: "\ bs p. qfree p \ qfree (qe p) \ (Ifm vs bs (qe p) = Ifm vs bs (E p))" + shows "\ bs p. qfree p \ qfree (CJNB qe p) \ (Ifm vs bs ((CJNB qe p)) = Ifm vs bs (E p))" +proof(clarify) + fix bs p + assume qfp: "qfree p" + let ?cjs = "conjuncts p" + let ?yes = "fst (partition bound0 ?cjs)" + let ?no = "snd (partition bound0 ?cjs)" + let ?cno = "list_conj ?no" + let ?cyes = "list_conj ?yes" + have part: "partition bound0 ?cjs = (?yes,?no)" by simp + from partition_P[OF part] have "\ q\ set ?yes. bound0 q" by blast + hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb') + hence yes_qf: "qfree (decr0 ?cyes )" by (simp add: decr0_qf) + from conjuncts_qf[OF qfp] partition_set[OF part] + have " \q\ set ?no. qfree q" by auto + hence no_qf: "qfree ?cno"by (simp add: list_conj_qf) + with qe have cno_qf:"qfree (qe ?cno )" + and noE: "Ifm vs bs (qe ?cno) = Ifm vs bs (E ?cno)" by blast+ + from cno_qf yes_qf have qf: "qfree (CJNB qe p)" + by (simp add: CJNB_def Let_def conj_qf split_def) + {fix bs + from conjuncts have "Ifm vs bs p = (\q\ set ?cjs. Ifm vs bs q)" by blast + also have "\ = ((\q\ set ?yes. Ifm vs bs q) \ (\q\ set ?no. Ifm vs bs q))" + using partition_set[OF part] by auto + finally have "Ifm vs bs p = ((Ifm vs bs ?cyes) \ (Ifm vs bs ?cno))" using list_conj[of vs bs] by simp} + hence "Ifm vs bs (E p) = (\x. (Ifm vs (x#bs) ?cyes) \ (Ifm vs (x#bs) ?cno))" by simp + also have "\ = (\x. (Ifm vs (y#bs) ?cyes) \ (Ifm vs (x#bs) ?cno))" + using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast + also have "\ = (Ifm vs bs (decr0 ?cyes) \ Ifm vs bs (E ?cno))" + by (auto simp add: decr0[OF yes_nb]) + also have "\ = (Ifm vs bs (conj (decr0 ?cyes) (qe ?cno)))" + using qe[rule_format, OF no_qf] by auto + finally have "Ifm vs bs (E p) = Ifm vs bs (CJNB qe p)" + by (simp add: Let_def CJNB_def split_def) + with qf show "qfree (CJNB qe p) \ Ifm vs bs (CJNB qe p) = Ifm vs bs (E p)" by blast +qed + +consts simpfm :: "fm \ fm" +recdef simpfm "measure fmsize" + "simpfm (Lt t) = simplt (simptm t)" + "simpfm (Le t) = simple (simptm t)" + "simpfm (Eq t) = simpeq(simptm t)" + "simpfm (NEq t) = simpneq(simptm t)" + "simpfm (And p q) = conj (simpfm p) (simpfm q)" + "simpfm (Or p q) = disj (simpfm p) (simpfm q)" + "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)" + "simpfm (Iff p q) = disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))" + "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))" + "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))" + "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))" + "simpfm (NOT (Iff p q)) = disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))" + "simpfm (NOT (Eq t)) = simpneq t" + "simpfm (NOT (NEq t)) = simpeq t" + "simpfm (NOT (Le t)) = simplt (Neg t)" + "simpfm (NOT (Lt t)) = simple (Neg t)" + "simpfm (NOT (NOT p)) = simpfm p" + "simpfm (NOT T) = F" + "simpfm (NOT F) = T" + "simpfm p = p" + +lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p" +by(induct p arbitrary: bs rule: simpfm.induct, auto) + +lemma simpfm_bound0: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "bound0 p \ bound0 (simpfm p)" +by (induct p rule: simpfm.induct, auto) + +lemma lt_qf[simp]: "qfree (lt t)" + apply (cases t, auto simp add: lt_def) + by (case_tac poly, auto) + +lemma le_qf[simp]: "qfree (le t)" + apply (cases t, auto simp add: le_def) + by (case_tac poly, auto) + +lemma eq_qf[simp]: "qfree (eq t)" + apply (cases t, auto simp add: eq_def) + by (case_tac poly, auto) + +lemma neq_qf[simp]: "qfree (neq t)" by (simp add: neq_def) + +lemma simplt_qf[simp]: "qfree (simplt t)" by (simp add: simplt_def Let_def split_def) +lemma simple_qf[simp]: "qfree (simple t)" by (simp add: simple_def Let_def split_def) +lemma simpeq_qf[simp]: "qfree (simpeq t)" by (simp add: simpeq_def Let_def split_def) +lemma simpneq_qf[simp]: "qfree (simpneq t)" by (simp add: simpneq_def Let_def split_def) + +lemma simpfm_qf[simp]: "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) + +lemma disj_lin: "islin p \ islin q \ islin (disj p q)" by (simp add: disj_def) +lemma conj_lin: "islin p \ islin q \ islin (conj p q)" by (simp add: conj_def) + +lemma assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "qfree p \ islin (simpfm p)" + apply (induct p rule: simpfm.induct) + apply (simp_all add: conj_lin disj_lin) + done + +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: "Ifm vs bs (prep p) = Ifm vs bs p" +by (induct p arbitrary: bs rule: prep.induct, auto) + + + + (* Generic quantifier elimination *) +consts qelim :: "fm \ (fm \ fm) \ fm" +recdef qelim "measure fmsize" + "qelim (E p) = (\ qe. DJ (CJNB qe) (qelim p qe))" + "qelim (A p) = (\ qe. not (qe ((qelim (NOT p) qe))))" + "qelim (NOT p) = (\ qe. not (qelim p qe))" + "qelim (And p q) = (\ qe. conj (qelim p qe) (qelim q qe))" + "qelim (Or p q) = (\ qe. disj (qelim p qe) (qelim q qe))" + "qelim (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: + assumes qe_inv: "\ bs p. qfree p \ qfree (qe p) \ (Ifm vs bs (qe p) = Ifm vs bs (E p))" + shows "\ bs. qfree (qelim p qe) \ (Ifm vs bs (qelim p qe) = Ifm vs bs p)" +using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]] +by (induct p rule: qelim.induct) auto + +subsection{* Core Procedure *} + +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 (CNP 0 c e)) = conj (eq (CP c)) (eq e)" + "minusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))" + "minusinf (Lt (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP (~\<^sub>p c)))" + "minusinf (Le (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP (~\<^sub>p c)))" + "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 (CNP 0 c e)) = conj (eq (CP c)) (eq e)" + "plusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))" + "plusinf (Lt (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP c))" + "plusinf (Le (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP c))" + "plusinf p = p" + +lemma minusinf_inf: assumes lp:"islin p" + shows "\z. \x < z. Ifm vs (x#bs) (minusinf p) \ Ifm vs (x#bs) p" + using lp +proof (induct p rule: minusinf.induct) + case 1 thus ?case by (auto,rule_tac x="min z za" in exI, auto) +next + case 2 thus ?case by (auto,rule_tac x="min z za" in exI, auto) +next + case (3 c e) hence nbe: "tmbound0 e" by simp + from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all + note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a] + let ?c = "Ipoly vs c" + let ?e = "Itm vs (y#bs) e" + have "?c=0 \ ?c > 0 \ ?c < 0" by arith + moreover {assume "?c = 0" hence ?case + using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto} + moreover {assume cp: "?c > 0" + {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e" + using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e < 0" by simp + hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))" + using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto} + moreover {assume cp: "?c < 0" + {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e" + using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e > 0" by simp + hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))" + using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto} + ultimately show ?case by blast +next + case (4 c e) hence nbe: "tmbound0 e" by simp + from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all + note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a] + let ?c = "Ipoly vs c" + let ?e = "Itm vs (y#bs) e" + have "?c=0 \ ?c > 0 \ ?c < 0" by arith + moreover {assume "?c = 0" hence ?case using eqs by auto} + moreover {assume cp: "?c > 0" + {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e" + using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e < 0" by simp + hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))" + using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto} + moreover {assume cp: "?c < 0" + {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e" + using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e > 0" by simp + hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))" + using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto} + ultimately show ?case by blast +next + case (5 c e) hence nbe: "tmbound0 e" by simp + from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all + hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm) + note eqs = lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] lt[OF nc(2), where ?'a = 'a] + let ?c = "Ipoly vs c" + let ?e = "Itm vs (y#bs) e" + have "?c=0 \ ?c > 0 \ ?c < 0" by arith + moreover {assume "?c = 0" hence ?case using eqs by auto} + moreover {assume cp: "?c > 0" + {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e" + using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e < 0" by simp + hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))" + using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto} + moreover {assume cp: "?c < 0" + {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e" + using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e > 0" by simp + hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))" + using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto} + ultimately show ?case by blast +next + case (6 c e) hence nbe: "tmbound0 e" by simp + from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all + hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm) + note eqs = lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] le[OF nc(2), where ?'a = 'a] + let ?c = "Ipoly vs c" + let ?e = "Itm vs (y#bs) e" + have "?c=0 \ ?c > 0 \ ?c < 0" by arith + moreover {assume "?c = 0" hence ?case using eqs by auto} + moreover {assume cp: "?c > 0" + {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e" + using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e < 0" by simp + hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))" + using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto} + moreover {assume cp: "?c < 0" + {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e" + using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e > 0" by simp + hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))" + using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto} + ultimately show ?case by blast +qed (auto) + +lemma plusinf_inf: assumes lp:"islin p" + shows "\z. \x > z. Ifm vs (x#bs) (plusinf p) \ Ifm vs (x#bs) p" + using lp +proof (induct p rule: plusinf.induct) + case 1 thus ?case by (auto,rule_tac x="max z za" in exI, auto) +next + case 2 thus ?case by (auto,rule_tac x="max z za" in exI, auto) +next + case (3 c e) hence nbe: "tmbound0 e" by simp + from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all + note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a] + let ?c = "Ipoly vs c" + let ?e = "Itm vs (y#bs) e" + have "?c=0 \ ?c > 0 \ ?c < 0" by arith + moreover {assume "?c = 0" hence ?case + using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto} + moreover {assume cp: "?c > 0" + {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e" + using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e > 0" by simp + hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))" + using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto} + moreover {assume cp: "?c < 0" + {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e" + using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e < 0" by simp + hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))" + using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto} + ultimately show ?case by blast +next + case (4 c e) hence nbe: "tmbound0 e" by simp + from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all + note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a] + let ?c = "Ipoly vs c" + let ?e = "Itm vs (y#bs) e" + have "?c=0 \ ?c > 0 \ ?c < 0" by arith + moreover {assume "?c = 0" hence ?case using eqs by auto} + moreover {assume cp: "?c > 0" + {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e" + using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e > 0" by simp + hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))" + using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto} + moreover {assume cp: "?c < 0" + {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e" + using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e < 0" by simp + hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))" + using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto} + ultimately show ?case by blast +next + case (5 c e) hence nbe: "tmbound0 e" by simp + from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all + hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm) + note eqs = lt[OF nc(1), where ?'a = 'a] lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] lt[OF nc(2), where ?'a = 'a] + let ?c = "Ipoly vs c" + let ?e = "Itm vs (y#bs) e" + have "?c=0 \ ?c > 0 \ ?c < 0" by arith + moreover {assume "?c = 0" hence ?case using eqs by auto} + moreover {assume cp: "?c > 0" + {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e" + using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e > 0" by simp + hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))" + using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto} + moreover {assume cp: "?c < 0" + {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e" + using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e < 0" by simp + hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))" + using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto} + ultimately show ?case by blast +next + case (6 c e) hence nbe: "tmbound0 e" by simp + from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all + hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm) + note eqs = lt[OF nc(1), where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] le[OF nc(2), where ?'a = 'a] + let ?c = "Ipoly vs c" + let ?e = "Itm vs (y#bs) e" + have "?c=0 \ ?c > 0 \ ?c < 0" by arith + moreover {assume "?c = 0" hence ?case using eqs by auto} + moreover {assume cp: "?c > 0" + {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e" + using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e > 0" by simp + hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))" + using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto} + moreover {assume cp: "?c < 0" + {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e" + using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute) + hence "?c * x + ?e < 0" by simp + hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))" + using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto} + ultimately show ?case by blast +qed (auto) + +lemma minusinf_nb: "islin p \ bound0 (minusinf p)" + by (induct p rule: minusinf.induct, auto simp add: eq_nb lt_nb le_nb) +lemma plusinf_nb: "islin p \ bound0 (plusinf p)" + by (induct p rule: minusinf.induct, auto simp add: eq_nb lt_nb le_nb) + +lemma minusinf_ex: assumes lp: "islin p" and ex: "Ifm vs (x#bs) (minusinf p)" + shows "\x. Ifm vs (x#bs) p" +proof- + from bound0_I [OF minusinf_nb[OF lp], where b="a" and bs ="bs"] ex + have th: "\ x. Ifm vs (x#bs) (minusinf p)" by auto + from minusinf_inf[OF lp, where bs="bs"] + obtain z where z_def: "\xx. Ifm vs (x#bs) p" +proof- + from bound0_I [OF plusinf_nb[OF lp], where b="a" and bs ="bs"] ex + have th: "\ x. Ifm vs (x#bs) (plusinf p)" by auto + from plusinf_inf[OF lp, where bs="bs"] + obtain z where z_def: "\x>z. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p" by blast + from th have "Ifm vs ((z + 1)#bs) (plusinf p)" by simp + moreover have "z + 1 > z" by simp + ultimately show ?thesis using z_def by auto +qed + +fun uset :: "fm \ (poly \ tm) list" where + "uset (And p q) = uset p @ uset q" +| "uset (Or p q) = uset p @ uset q" +| "uset (Eq (CNP 0 a e)) = [(a,e)]" +| "uset (Le (CNP 0 a e)) = [(a,e)]" +| "uset (Lt (CNP 0 a e)) = [(a,e)]" +| "uset (NEq (CNP 0 a e)) = [(a,e)]" +| "uset p = []" + +lemma uset_l: + assumes lp: "islin p" + shows "\ (c,s) \ set (uset p). isnpoly c \ c \ 0\<^sub>p \ tmbound0 s \ allpolys isnpoly s" +using lp by(induct p rule: uset.induct,auto) + +lemma minusinf_uset0: + assumes lp: "islin p" + and nmi: "\ (Ifm vs (x#bs) (minusinf p))" + and ex: "Ifm vs (x#bs) p" (is "?I x p") + shows "\ (c,s) \ set (uset p). x \ - Itm vs (x#bs) s / Ipoly vs c" +proof- + have "\ (c,s) \ set (uset p). (Ipoly vs c < 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s) \ (Ipoly vs c > 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s)" + using lp nmi ex + apply (induct p rule: minusinf.induct, auto simp add: eq le lt nth_pos2 polyneg_norm) + apply (auto simp add: linorder_not_less order_le_less) + done + then obtain c s where csU: "(c,s) \ set (uset p)" and x: "(Ipoly vs c < 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s) \ (Ipoly vs c > 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s)" by blast + hence "x \ (- Itm vs (x#bs) s) / Ipoly vs c" + using divide_le_eq[of "- Itm vs (x#bs) s" "Ipoly vs c" x] + by (auto simp add: mult_commute del: divide_minus_left) + thus ?thesis using csU by auto +qed + +lemma minusinf_uset: + assumes lp: "islin p" + and nmi: "\ (Ifm vs (a#bs) (minusinf p))" + and ex: "Ifm vs (x#bs) p" (is "?I x p") + shows "\ (c,s) \ set (uset p). x \ - Itm vs (a#bs) s / Ipoly vs c" +proof- + from nmi have nmi': "\ (Ifm vs (x#bs) (minusinf p))" + by (simp add: bound0_I[OF minusinf_nb[OF lp], where b=x and b'=a]) + from minusinf_uset0[OF lp nmi' ex] + obtain c s where csU: "(c,s) \ set (uset p)" and th: "x \ - Itm vs (x#bs) s / Ipoly vs c" by blast + from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s" by simp + from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis by auto +qed + + +lemma plusinf_uset0: + assumes lp: "islin p" + and nmi: "\ (Ifm vs (x#bs) (plusinf p))" + and ex: "Ifm vs (x#bs) p" (is "?I x p") + shows "\ (c,s) \ set (uset p). x \ - Itm vs (x#bs) s / Ipoly vs c" +proof- + have "\ (c,s) \ set (uset p). (Ipoly vs c < 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s) \ (Ipoly vs c > 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s)" + using lp nmi ex + apply (induct p rule: minusinf.induct, auto simp add: eq le lt nth_pos2 polyneg_norm) + apply (auto simp add: linorder_not_less order_le_less) + done + then obtain c s where csU: "(c,s) \ set (uset p)" and x: "(Ipoly vs c < 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s) \ (Ipoly vs c > 0 \ Ipoly vs c * x \ - Itm vs (x#bs) s)" by blast + hence "x \ (- Itm vs (x#bs) s) / Ipoly vs c" + using le_divide_eq[of x "- Itm vs (x#bs) s" "Ipoly vs c"] + by (auto simp add: mult_commute del: divide_minus_left) + thus ?thesis using csU by auto +qed + +lemma plusinf_uset: + assumes lp: "islin p" + and nmi: "\ (Ifm vs (a#bs) (plusinf p))" + and ex: "Ifm vs (x#bs) p" (is "?I x p") + shows "\ (c,s) \ set (uset p). x \ - Itm vs (a#bs) s / Ipoly vs c" +proof- + from nmi have nmi': "\ (Ifm vs (x#bs) (plusinf p))" + by (simp add: bound0_I[OF plusinf_nb[OF lp], where b=x and b'=a]) + from plusinf_uset0[OF lp nmi' ex] + obtain c s where csU: "(c,s) \ set (uset p)" and th: "x \ - Itm vs (x#bs) s / Ipoly vs c" by blast + from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s" by simp + from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis by auto +qed + +lemma lin_dense: + assumes lp: "islin p" + and noS: "\ t. l < t \ t< u \ t \ (\ (c,t). - Itm vs (x#bs) t / Ipoly vs c) ` set (uset p)" + (is "\ t. _ \ _ \ t \ (\ (c,t). - ?Nt x t / ?N c) ` ?U p") + and lx: "l < x" and xu:"x < u" and px:" Ifm vs (x#bs) p" + and ly: "l < y" and yu: "y < u" + shows "Ifm vs (y#bs) p" +using lp px noS +proof (induct p rule: islin.induct) + case (5 c s) + from "5.prems" + have lin: "isnpoly c" "c \ 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s" + and px: "Ifm vs (x # bs) (Lt (CNP 0 c s))" + and noS: "\t. l < t \ t < u \ t \ - Itm vs (x # bs) s / \c\\<^sub>p\<^bsup>vs\<^esup>" by simp_all + from ly yu noS have yne: "y \ - ?Nt x s / \c\\<^sub>p\<^bsup>vs\<^esup>" by simp + hence ycs: "y < - ?Nt x s / ?N c \ y > -?Nt x s / ?N c" by auto + have ccs: "?N c = 0 \ ?N c < 0 \ ?N c > 0" by dlo + moreover + {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])} + moreover + {assume c: "?N c > 0" + from px pos_less_divide_eq[OF c, where a="x" and b="-?Nt x s"] + have px': "x < - ?Nt x s / ?N c" + by (auto simp add: not_less ring_simps) + {assume y: "y < - ?Nt x s / ?N c" + hence "y * ?N c < - ?Nt x s" + by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric]) + hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps) + hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp} + moreover + {assume y: "y > -?Nt x s / ?N c" + with yu have eu: "u > - ?Nt x s / ?N c" by auto + with noS ly yu have th: "- ?Nt x s / ?N c \ l" by (cases "- ?Nt x s / ?N c > l", auto) + with lx px' have "False" by simp hence ?case by simp } + ultimately have ?case using ycs by blast + } + moreover + {assume c: "?N c < 0" + from px neg_divide_less_eq[OF c, where a="x" and b="-?Nt x s"] + have px': "x > - ?Nt x s / ?N c" + by (auto simp add: not_less ring_simps) + {assume y: "y > - ?Nt x s / ?N c" + hence "y * ?N c < - ?Nt x s" + by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric]) + hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps) + hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp} + moreover + {assume y: "y < -?Nt x s / ?N c" + with ly have eu: "l < - ?Nt x s / ?N c" by auto + with noS ly yu have th: "- ?Nt x s / ?N c \ u" by (cases "- ?Nt x s / ?N c < u", auto) + with xu px' have "False" by simp hence ?case by simp } + ultimately have ?case using ycs by blast + } + ultimately show ?case by blast +next + case (6 c s) + from "6.prems" + have lin: "isnpoly c" "c \ 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s" + and px: "Ifm vs (x # bs) (Le (CNP 0 c s))" + and noS: "\t. l < t \ t < u \ t \ - Itm vs (x # bs) s / \c\\<^sub>p\<^bsup>vs\<^esup>" by simp_all + from ly yu noS have yne: "y \ - ?Nt x s / \c\\<^sub>p\<^bsup>vs\<^esup>" by simp + hence ycs: "y < - ?Nt x s / ?N c \ y > -?Nt x s / ?N c" by auto + have ccs: "?N c = 0 \ ?N c < 0 \ ?N c > 0" by dlo + moreover + {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])} + moreover + {assume c: "?N c > 0" + from px pos_le_divide_eq[OF c, where a="x" and b="-?Nt x s"] + have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less ring_simps) + {assume y: "y < - ?Nt x s / ?N c" + hence "y * ?N c < - ?Nt x s" + by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric]) + hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps) + hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp} + moreover + {assume y: "y > -?Nt x s / ?N c" + with yu have eu: "u > - ?Nt x s / ?N c" by auto + with noS ly yu have th: "- ?Nt x s / ?N c \ l" by (cases "- ?Nt x s / ?N c > l", auto) + with lx px' have "False" by simp hence ?case by simp } + ultimately have ?case using ycs by blast + } + moreover + {assume c: "?N c < 0" + from px neg_divide_le_eq[OF c, where a="x" and b="-?Nt x s"] + have px': "x >= - ?Nt x s / ?N c" by (simp add: ring_simps) + {assume y: "y > - ?Nt x s / ?N c" + hence "y * ?N c < - ?Nt x s" + by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric]) + hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps) + hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp} + moreover + {assume y: "y < -?Nt x s / ?N c" + with ly have eu: "l < - ?Nt x s / ?N c" by auto + with noS ly yu have th: "- ?Nt x s / ?N c \ u" by (cases "- ?Nt x s / ?N c < u", auto) + with xu px' have "False" by simp hence ?case by simp } + ultimately have ?case using ycs by blast + } + ultimately show ?case by blast +next + case (3 c s) + from "3.prems" + have lin: "isnpoly c" "c \ 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s" + and px: "Ifm vs (x # bs) (Eq (CNP 0 c s))" + and noS: "\t. l < t \ t < u \ t \ - Itm vs (x # bs) s / \c\\<^sub>p\<^bsup>vs\<^esup>" by simp_all + from ly yu noS have yne: "y \ - ?Nt x s / \c\\<^sub>p\<^bsup>vs\<^esup>" by simp + hence ycs: "y < - ?Nt x s / ?N c \ y > -?Nt x s / ?N c" by auto + have ccs: "?N c = 0 \ ?N c < 0 \ ?N c > 0" by dlo + moreover + {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])} + moreover + {assume c: "?N c > 0" hence cnz: "?N c \ 0" by simp + from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz + have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps) + {assume y: "y < -?Nt x s / ?N c" + with ly have eu: "l < - ?Nt x s / ?N c" by auto + with noS ly yu have th: "- ?Nt x s / ?N c \ u" by (cases "- ?Nt x s / ?N c < u", auto) + with xu px' have "False" by simp hence ?case by simp } + moreover + {assume y: "y > -?Nt x s / ?N c" + with yu have eu: "u > - ?Nt x s / ?N c" by auto + with noS ly yu have th: "- ?Nt x s / ?N c \ l" by (cases "- ?Nt x s / ?N c > l", auto) + with lx px' have "False" by simp hence ?case by simp } + ultimately have ?case using ycs by blast + } + moreover + {assume c: "?N c < 0" hence cnz: "?N c \ 0" by simp + from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz + have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps) + {assume y: "y < -?Nt x s / ?N c" + with ly have eu: "l < - ?Nt x s / ?N c" by auto + with noS ly yu have th: "- ?Nt x s / ?N c \ u" by (cases "- ?Nt x s / ?N c < u", auto) + with xu px' have "False" by simp hence ?case by simp } + moreover + {assume y: "y > -?Nt x s / ?N c" + with yu have eu: "u > - ?Nt x s / ?N c" by auto + with noS ly yu have th: "- ?Nt x s / ?N c \ l" by (cases "- ?Nt x s / ?N c > l", auto) + with lx px' have "False" by simp hence ?case by simp } + ultimately have ?case using ycs by blast + } + ultimately show ?case by blast +next + case (4 c s) + from "4.prems" + have lin: "isnpoly c" "c \ 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s" + and px: "Ifm vs (x # bs) (NEq (CNP 0 c s))" + and noS: "\t. l < t \ t < u \ t \ - Itm vs (x # bs) s / \c\\<^sub>p\<^bsup>vs\<^esup>" by simp_all + from ly yu noS have yne: "y \ - ?Nt x s / \c\\<^sub>p\<^bsup>vs\<^esup>" by simp + hence ycs: "y < - ?Nt x s / ?N c \ y > -?Nt x s / ?N c" by auto + have ccs: "?N c = 0 \ ?N c \ 0" by dlo + moreover + {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])} + moreover + {assume c: "?N c \ 0" + from yne c eq_divide_eq[of "y" "- ?Nt x s" "?N c"] have ?case + by (simp add: ring_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric]) } + ultimately show ?case by blast +qed (auto simp add: nth_pos2 tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"] bound0_I[where vs=vs and bs="bs" and b="y" and b'="x"]) + +lemma one_plus_one_pos[simp]: "(1::'a::{ordered_field}) + 1 > 0" +proof- + have op: "(1::'a) > 0" by simp + from add_pos_pos[OF op op] show ?thesis . +qed + +lemma one_plus_one_nonzero[simp]: "(1::'a::{ordered_field}) + 1 \ 0" + using one_plus_one_pos[where ?'a = 'a] by (simp add: less_le) + +lemma half_sum_eq: "(u + u) / (1+1) = (u::'a::{ordered_field})" +proof- + have "(u + u) = (1 + 1) * u" by (simp add: ring_simps) + hence "(u + u) / (1+1) = (1 + 1)*u / (1 + 1)" by simp + with nonzero_mult_divide_cancel_left[OF one_plus_one_nonzero, of u] show ?thesis by simp +qed + +lemma inf_uset: + assumes lp: "islin p" + and nmi: "\ (Ifm vs (x#bs) (minusinf p))" (is "\ (Ifm vs (x#bs) (?M p))") + and npi: "\ (Ifm vs (x#bs) (plusinf p))" (is "\ (Ifm vs (x#bs) (?P p))") + and ex: "\ x. Ifm vs (x#bs) p" (is "\ x. ?I x p") + shows "\ (c,t) \ set (uset p). \ (d,s) \ set (uset p). ?I ((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) / (1 + 1)) p" +proof- + let ?Nt = "\ x t. Itm vs (x#bs) t" + let ?N = "Ipoly vs" + let ?U = "set (uset p)" + from ex obtain a where pa: "?I a p" by blast + from bound0_I[OF minusinf_nb[OF lp], where bs="bs" and b="x" and b'="a"] nmi + have nmi': "\ (?I a (?M p))" by simp + from bound0_I[OF plusinf_nb[OF lp], where bs="bs" and b="x" and b'="a"] npi + have npi': "\ (?I a (?P p))" by simp + have "\ (c,t) \ set (uset p). \ (d,s) \ set (uset p). ?I ((- ?Nt a t/?N c + - ?Nt a s /?N d) / (1 + 1)) p" + proof- + let ?M = "(\ (c,t). - ?Nt a t / ?N c) ` ?U" + have fM: "finite ?M" by auto + from minusinf_uset[OF lp nmi pa] plusinf_uset[OF lp npi pa] + have "\ (c,t) \ set (uset p). \ (d,s) \ set (uset p). a \ - ?Nt x t / ?N c \ a \ - ?Nt x s / ?N d" by blast + then obtain "c" "t" "d" "s" where + ctU: "(c,t) \ ?U" and dsU: "(d,s) \ ?U" + and xs1: "a \ - ?Nt x s / ?N d" and tx1: "a \ - ?Nt x t / ?N c" by blast + from uset_l[OF lp] ctU dsU tmbound0_I[where bs="bs" and b="x" and b'="a"] xs1 tx1 + have xs: "a \ - ?Nt a s / ?N d" and tx: "a \ - ?Nt a t / ?N c" by auto + from ctU 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 ctM: "- ?Nt a t / ?N c \ ?M" using ctU by auto + have dsM: "- ?Nt a s / ?N d \ ?M" using dsU 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 \ - ?Nt a t / ?N c" using ctM Mne by simp hence lx: "?l \ a" using tx by simp + have "- ?Nt a s / ?N d \ ?u" using dsM 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 "\ (nu,tu) \ ?U. u = - ?Nt a tu / ?N nu" by auto + then obtain "tu" "nu" where tuU: "(nu,tu) \ ?U" and tuu:"u= - ?Nt a tu / ?N nu" by blast + from half_sum_eq[of u] pu tuu + have "?I (((- ?Nt a tu / ?N nu) + (- ?Nt a tu / ?N nu)) / (1 + 1)) 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 "\ (t1n,t1u) \ ?U. t1 = - ?Nt a t1u / ?N t1n" by auto + then obtain "t1u" "t1n" where t1uU: "(t1n,t1u) \ ?U" and t1u: "t1 = - ?Nt a t1u / ?N t1n" by blast + from t2M have "\ (t2n,t2u) \ ?U. t2 = - ?Nt a t2u / ?N t2n" by auto + then obtain "t2u" "t2n" where t2uU: "(t2n,t2u) \ ?U" and t2u: "t2 = - ?Nt a t2u / ?N t2n" by blast + from t1x xt2 have t1t2: "t1 < t2" by simp + let ?u = "(t1 + t2) / (1 + 1)" + 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: "(n,l) \ ?U" and smU:"(m,s) \ ?U" + and pu: "?I ((- ?Nt a l / ?N n + - ?Nt a s / ?N m) / (1 + 1)) p" by blast + from lnU smU uset_l[OF lp] have nbl: "tmbound0 l" and nbs: "tmbound0 s" by auto + from tmbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"] + tmbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu + have "?I ((- ?Nt x l / ?N n + - ?Nt x s / ?N m) / (1 + 1)) p" by simp + with lnU smU + show ?thesis by auto +qed + + (* The Ferrante - Rackoff Theorem *) + +theorem fr_eq: + assumes lp: "islin p" + shows "(\ x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \ (Ifm vs (x#bs) (plusinf p)) \ (\ (n,t) \ set (uset p). \ (m,s) \ set (uset p). Ifm vs (((- Itm vs (x#bs) t / Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /(1 + 1))#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 inf_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 minusinf_ex[OF lp m] have "?E" .} + moreover {assume p: "?P" from plusinf_ex[OF lp p] have "?E" . } + moreover {assume f:"?F" hence "?E" by blast} + ultimately show "?E" by blast +qed + +section{* First implementation : Naive by encoding all case splits locally *} +definition "msubsteq c t d s a r = + evaldjf (split conj) + [(let cd = c *\<^sub>p d in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (conj (Eq (CP c)) (NEq (CP d)) , Eq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (NEq (CP c)) (Eq (CP d)) , Eq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (Eq (CP c)) (Eq (CP d)) , Eq r)]" + +lemma msubsteq_nb: assumes lp: "islin (Eq (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s" + shows "bound0 (msubsteq c t d s a r)" +proof- + have th: "\x\ set [(let cd = c *\<^sub>p d in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (conj (Eq (CP c)) (NEq (CP d)) , Eq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (NEq (CP c)) (Eq (CP d)) , Eq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (Eq (CP c)) (Eq (CP d)) , Eq r)]. bound0 (split conj x)" + using lp by (simp add: Let_def t s ) + from evaldjf_bound0[OF th] show ?thesis by (simp add: msubsteq_def) +qed + +lemma msubsteq: assumes lp: "islin (Eq (CNP 0 a r))" + shows "Ifm vs (x#bs) (msubsteq c t d s a r) = Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Eq (CNP 0 a r))" (is "?lhs = ?rhs") +proof- + let ?Nt = "\(x::'a) t. Itm vs (x#bs) t" + let ?N = "\p. Ipoly vs p" + let ?c = "?N c" + let ?d = "?N d" + let ?t = "?Nt x t" + let ?s = "?Nt x s" + let ?a = "?N a" + let ?r = "?Nt x r" + from lp have lin:"isnpoly a" "a \ 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all + note r= tmbound0_I[OF lin(3), of vs _ bs x] + have cd_cs: "?c * ?d \ 0 \ (?c = 0 \ ?d = 0) \ (?c = 0 \ ?d \ 0) \ (?c \ 0 \ ?d = 0)" by auto + moreover + {assume c: "?c = 0" and d: "?d=0" + hence ?thesis by (simp add: r[of 0] msubsteq_def Let_def evaldjf_ex)} + moreover + {assume c: "?c = 0" and d: "?d\0" + from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?s / ((1 + 1)*?d)" by simp + have "?rhs = Ifm vs (-?s / ((1 + 1)*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (-?s / ((1 + 1)*?d)) + ?r = 0" by (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \d\\<^sub>p\<^bsup>vs\<^esup>))"]) + also have "\ \ (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) = 0" + using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp + also have "\ \ (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r= 0" + by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib) + + also have "\ \ - (?a * ?s) + (1 + 1)*?d*?r = 0" using d by simp + finally have ?thesis using c d + apply (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \d\\<^sub>p\<^bsup>vs\<^esup>))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + apply simp + done} + moreover + {assume c: "?c \ 0" and d: "?d=0" + from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?t / ((1 + 1)*?c)" by simp + have "?rhs = Ifm vs (-?t / ((1 + 1)*?c) # bs) (Eq (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (-?t / ((1 + 1)*?c)) + ?r = 0" by (simp add: r[of "- (?t/ ((1 + 1)* ?c))"]) + also have "\ \ (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) = 0" + using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp + also have "\ \ (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r= 0" + by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib) + also have "\ \ - (?a * ?t) + (1 + 1)*?c*?r = 0" using c by simp + finally have ?thesis using c d + apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + apply simp + done } + moreover + {assume c: "?c \ 0" and d: "?d\0" hence dc: "?c * ?d *(1 + 1) \ 0" by simp + from add_frac_eq[OF c d, of "- ?t" "- ?s"] + have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" + by (simp add: ring_simps) + have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r = 0" + by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"]) + also have "\ \ ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) =0 " + using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r =0" + using nonzero_mult_divide_cancel_left[OF dc] c d + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using c d + apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex ring_simps) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + apply (simp add: ring_simps) + done } + ultimately show ?thesis by blast +qed + + +definition "msubstneq c t d s a r = + evaldjf (split conj) + [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (conj (Eq (CP c)) (NEq (CP d)) , NEq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (NEq (CP c)) (Eq (CP d)) , NEq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (Eq (CP c)) (Eq (CP d)) , NEq r)]" + +lemma msubstneq_nb: assumes lp: "islin (NEq (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s" + shows "bound0 (msubstneq c t d s a r)" +proof- + have th: "\x\ set [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (conj (Eq (CP c)) (NEq (CP d)) , NEq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (NEq (CP c)) (Eq (CP d)) , NEq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (Eq (CP c)) (Eq (CP d)) , NEq r)]. bound0 (split conj x)" + using lp by (simp add: Let_def t s ) + from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstneq_def) +qed + +lemma msubstneq: assumes lp: "islin (Eq (CNP 0 a r))" + shows "Ifm vs (x#bs) (msubstneq c t d s a r) = Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (NEq (CNP 0 a r))" (is "?lhs = ?rhs") +proof- + let ?Nt = "\(x::'a) t. Itm vs (x#bs) t" + let ?N = "\p. Ipoly vs p" + let ?c = "?N c" + let ?d = "?N d" + let ?t = "?Nt x t" + let ?s = "?Nt x s" + let ?a = "?N a" + let ?r = "?Nt x r" + from lp have lin:"isnpoly a" "a \ 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all + note r= tmbound0_I[OF lin(3), of vs _ bs x] + have cd_cs: "?c * ?d \ 0 \ (?c = 0 \ ?d = 0) \ (?c = 0 \ ?d \ 0) \ (?c \ 0 \ ?d = 0)" by auto + moreover + {assume c: "?c = 0" and d: "?d=0" + hence ?thesis by (simp add: r[of 0] msubstneq_def Let_def evaldjf_ex)} + moreover + {assume c: "?c = 0" and d: "?d\0" + from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?s / ((1 + 1)*?d)" by simp + have "?rhs = Ifm vs (-?s / ((1 + 1)*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (-?s / ((1 + 1)*?d)) + ?r \ 0" by (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \d\\<^sub>p\<^bsup>vs\<^esup>))"]) + also have "\ \ (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) \ 0" + using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp + also have "\ \ (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r\ 0" + by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib) + + also have "\ \ - (?a * ?s) + (1 + 1)*?d*?r \ 0" using d by simp + finally have ?thesis using c d + apply (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \d\\<^sub>p\<^bsup>vs\<^esup>))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + apply simp + done} + moreover + {assume c: "?c \ 0" and d: "?d=0" + from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?t / ((1 + 1)*?c)" by simp + have "?rhs = Ifm vs (-?t / ((1 + 1)*?c) # bs) (NEq (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (-?t / ((1 + 1)*?c)) + ?r \ 0" by (simp add: r[of "- (?t/ ((1 + 1)* ?c))"]) + also have "\ \ (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) \ 0" + using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp + also have "\ \ (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r \ 0" + by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib) + also have "\ \ - (?a * ?t) + (1 + 1)*?c*?r \ 0" using c by simp + finally have ?thesis using c d + apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + apply simp + done } + moreover + {assume c: "?c \ 0" and d: "?d\0" hence dc: "?c * ?d *(1 + 1) \ 0" by simp + from add_frac_eq[OF c d, of "- ?t" "- ?s"] + have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" + by (simp add: ring_simps) + have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r \ 0" + by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"]) + also have "\ \ ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) \ 0 " + using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r \ 0" + using nonzero_mult_divide_cancel_left[OF dc] c d + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using c d + apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex ring_simps) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + apply (simp add: ring_simps) + done } + ultimately show ?thesis by blast +qed + +definition "msubstlt c t d s a r = + evaldjf (split conj) + [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (let cd = c *\<^sub>p d in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Lt (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (lt (CP c)) (Eq (CP d)) , Lt (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Lt (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (lt (CP d)) (Eq (CP c)) , Lt (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (Eq (CP c)) (Eq (CP d)) , Lt r)]" + +lemma msubstlt_nb: assumes lp: "islin (Lt (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s" + shows "bound0 (msubstlt c t d s a r)" +proof- + have th: "\x\ set [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (let cd = c *\<^sub>p d in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Lt (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (lt (CP c)) (Eq (CP d)) , Lt (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Lt (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (lt (CP d)) (Eq (CP c)) , Lt (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (Eq (CP c)) (Eq (CP d)) , Lt r)]. bound0 (split conj x)" + using lp by (simp add: Let_def t s lt_nb ) + from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstlt_def) +qed + + +lemma msubstlt: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Lt (CNP 0 a r))" + shows "Ifm vs (x#bs) (msubstlt c t d s a r) \ + Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Lt (CNP 0 a r))" (is "?lhs = ?rhs") +proof- + let ?Nt = "\x t. Itm vs (x#bs) t" + let ?N = "\p. Ipoly vs p" + let ?c = "?N c" + let ?d = "?N d" + let ?t = "?Nt x t" + let ?s = "?Nt x s" + let ?a = "?N a" + let ?r = "?Nt x r" + from lp have lin:"isnpoly a" "a \ 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all + note r= tmbound0_I[OF lin(3), of vs _ bs x] + have cd_cs: "?c * ?d < 0 \ ?c * ?d > 0 \ (?c = 0 \ ?d = 0) \ (?c = 0 \ ?d < 0) \ (?c = 0 \ ?d > 0) \ (?c < 0 \ ?d = 0) \ (?c > 0 \ ?d = 0)" by auto + moreover + {assume c: "?c=0" and d: "?d=0" + hence ?thesis using nc nd by (simp add: polyneg_norm lt r[of 0] msubstlt_def Let_def evaldjf_ex)} + moreover + {assume dc: "?c*?d > 0" + from mult_pos_pos[OF one_plus_one_pos dc] have dc': "(1 + 1)*?c *?d > 0" by simp + hence c:"?c \ 0" and d: "?d\ 0" by auto + from dc' have dc'': "\ (1 + 1)*?c *?d < 0" by simp + from add_frac_eq[OF c d, of "- ?t" "- ?s"] + have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" + by (simp add: ring_simps) + have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0" + by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"]) + also have "\ \ ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) < 0" + + using dc' dc'' mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r < 0" + using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using dc c d nc nd dc' + apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + by (simp add: ring_simps order_less_not_sym[OF dc])} + moreover + {assume dc: "?c*?d < 0" + + from dc one_plus_one_pos[where ?'a='a] have dc': "(1 + 1)*?c *?d < 0" + apply (simp add: mult_less_0_iff field_simps) + apply (rule add_neg_neg) + apply (simp_all add: mult_less_0_iff) + done + hence c:"?c \ 0" and d: "?d\ 0" by auto + from add_frac_eq[OF c d, of "- ?t" "- ?s"] + have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" + by (simp add: ring_simps) + have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0" + by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"]) + + also have "\ \ ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) > 0" + + using dc' order_less_not_sym[OF dc'] mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp + also have "\ \ ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r < 0" + using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using dc c d nc nd + apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + by (simp add: ring_simps order_less_not_sym[OF dc]) } + moreover + {assume c: "?c > 0" and d: "?d=0" + from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff) + from c have c': "(1 + 1)*?c \ 0" by simp + from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps) + have "?rhs \ Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"]) + also have "\ \ (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) < 0" + using c mult_less_cancel_left_disj[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp + also have "\ \ - ?a*?t+ (1 + 1)*?c *?r < 0" + using nonzero_mult_divide_cancel_left[OF c'] c + by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib) + finally have ?thesis using c d nc nd + apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + using c order_less_not_sym[OF c] less_imp_neq[OF c] + by (simp add: ring_simps ) } + moreover + {assume c: "?c < 0" and d: "?d=0" hence c': "(1 + 1)*?c \ 0" by simp + from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff) + from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps) + have "?rhs \ Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"]) + also have "\ \ (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) > 0" + using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_less_cancel_left_disj[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp + also have "\ \ ?a*?t - (1 + 1)*?c *?r < 0" + using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using c d nc nd + apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + using c order_less_not_sym[OF c] less_imp_neq[OF c] + by (simp add: ring_simps ) } + moreover + moreover + {assume c: "?c = 0" and d: "?d>0" + from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff) + from d have d': "(1 + 1)*?d \ 0" by simp + from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps) + have "?rhs \ Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"]) + also have "\ \ (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) < 0" + using d mult_less_cancel_left_disj[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp + also have "\ \ - ?a*?s+ (1 + 1)*?d *?r < 0" + using nonzero_mult_divide_cancel_left[OF d'] d + by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib) + finally have ?thesis using c d nc nd + apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + using d order_less_not_sym[OF d] less_imp_neq[OF d] + by (simp add: ring_simps ) } + moreover + {assume c: "?c = 0" and d: "?d<0" hence d': "(1 + 1)*?d \ 0" by simp + from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff) + from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps) + have "?rhs \ Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"]) + also have "\ \ (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) > 0" + using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_less_cancel_left_disj[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp + also have "\ \ ?a*?s - (1 + 1)*?d *?r < 0" + using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using c d nc nd + apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + using d order_less_not_sym[OF d] less_imp_neq[OF d] + by (simp add: ring_simps ) } +ultimately show ?thesis by blast +qed + +definition "msubstle c t d s a r = + evaldjf (split conj) + [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (let cd = c *\<^sub>p d in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Le (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (lt (CP c)) (Eq (CP d)) , Le (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Le (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (lt (CP d)) (Eq (CP c)) , Le (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (Eq (CP c)) (Eq (CP d)) , Le r)]" + +lemma msubstle_nb: assumes lp: "islin (Le (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s" + shows "bound0 (msubstle c t d s a r)" +proof- + have th: "\x\ set [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (let cd = c *\<^sub>p d in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))), + (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Le (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (lt (CP c)) (Eq (CP d)) , Le (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))), + (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Le (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (lt (CP d)) (Eq (CP c)) , Le (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))), + (conj (Eq (CP c)) (Eq (CP d)) , Le r)]. bound0 (split conj x)" + using lp by (simp add: Let_def t s lt_nb ) + from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstle_def) +qed + +lemma msubstle: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Le (CNP 0 a r))" + shows "Ifm vs (x#bs) (msubstle c t d s a r) \ + Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Le (CNP 0 a r))" (is "?lhs = ?rhs") +proof- + let ?Nt = "\x t. Itm vs (x#bs) t" + let ?N = "\p. Ipoly vs p" + let ?c = "?N c" + let ?d = "?N d" + let ?t = "?Nt x t" + let ?s = "?Nt x s" + let ?a = "?N a" + let ?r = "?Nt x r" + from lp have lin:"isnpoly a" "a \ 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all + note r= tmbound0_I[OF lin(3), of vs _ bs x] + have cd_cs: "?c * ?d < 0 \ ?c * ?d > 0 \ (?c = 0 \ ?d = 0) \ (?c = 0 \ ?d < 0) \ (?c = 0 \ ?d > 0) \ (?c < 0 \ ?d = 0) \ (?c > 0 \ ?d = 0)" by auto + moreover + {assume c: "?c=0" and d: "?d=0" + hence ?thesis using nc nd by (simp add: polyneg_norm polymul_norm lt r[of 0] msubstle_def Let_def evaldjf_ex)} + moreover + {assume dc: "?c*?d > 0" + from mult_pos_pos[OF one_plus_one_pos dc] have dc': "(1 + 1)*?c *?d > 0" by simp + hence c:"?c \ 0" and d: "?d\ 0" by auto + from dc' have dc'': "\ (1 + 1)*?c *?d < 0" by simp + from add_frac_eq[OF c d, of "- ?t" "- ?s"] + have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" + by (simp add: ring_simps) + have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0" + by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"]) + also have "\ \ ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) <= 0" + + using dc' dc'' mult_le_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r <= 0" + using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using dc c d nc nd dc' + apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + by (simp add: ring_simps order_less_not_sym[OF dc])} + moreover + {assume dc: "?c*?d < 0" + + from dc one_plus_one_pos[where ?'a='a] have dc': "(1 + 1)*?c *?d < 0" + by (simp add: mult_less_0_iff field_simps add_neg_neg add_pos_pos) + hence c:"?c \ 0" and d: "?d\ 0" by auto + from add_frac_eq[OF c d, of "- ?t" "- ?s"] + have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)" + by (simp add: ring_simps) + have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0" + by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"]) + + also have "\ \ ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) >= 0" + + using dc' order_less_not_sym[OF dc'] mult_le_cancel_left[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp + also have "\ \ ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r <= 0" + using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using dc c d nc nd + apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + by (simp add: ring_simps order_less_not_sym[OF dc]) } + moreover + {assume c: "?c > 0" and d: "?d=0" + from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff) + from c have c': "(1 + 1)*?c \ 0" by simp + from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps) + have "?rhs \ Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"]) + also have "\ \ (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) <= 0" + using c mult_le_cancel_left[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp + also have "\ \ - ?a*?t+ (1 + 1)*?c *?r <= 0" + using nonzero_mult_divide_cancel_left[OF c'] c + by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib) + finally have ?thesis using c d nc nd + apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + using c order_less_not_sym[OF c] less_imp_neq[OF c] + by (simp add: ring_simps ) } + moreover + {assume c: "?c < 0" and d: "?d=0" hence c': "(1 + 1)*?c \ 0" by simp + from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff) + from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps) + have "?rhs \ Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"]) + also have "\ \ (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) >= 0" + using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_le_cancel_left[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp + also have "\ \ ?a*?t - (1 + 1)*?c *?r <= 0" + using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using c d nc nd + apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + using c order_less_not_sym[OF c] less_imp_neq[OF c] + by (simp add: ring_simps ) } + moreover + moreover + {assume c: "?c = 0" and d: "?d>0" + from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff) + from d have d': "(1 + 1)*?d \ 0" by simp + from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps) + have "?rhs \ Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"]) + also have "\ \ (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) <= 0" + using d mult_le_cancel_left[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp + also have "\ \ - ?a*?s+ (1 + 1)*?d *?r <= 0" + using nonzero_mult_divide_cancel_left[OF d'] d + by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib) + finally have ?thesis using c d nc nd + apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + using d order_less_not_sym[OF d] less_imp_neq[OF d] + by (simp add: ring_simps ) } + moreover + {assume c: "?c = 0" and d: "?d<0" hence d': "(1 + 1)*?d \ 0" by simp + from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff) + from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps) + have "?rhs \ Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th) + also have "\ \ ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"]) + also have "\ \ (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) >= 0" + using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_le_cancel_left[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp + also have "\ \ ?a*?s - (1 + 1)*?d *?r <= 0" + using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' + by (simp add: ring_simps diff_divide_distrib del: left_distrib) + finally have ?thesis using c d nc nd + apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm) + apply (simp only: one_add_one_is_two[symmetric] of_int_add) + using d order_less_not_sym[OF d] less_imp_neq[OF d] + by (simp add: ring_simps ) } +ultimately show ?thesis by blast +qed + + +fun msubst :: "fm \ (poly \ tm) \ (poly \ tm) \ fm" where + "msubst (And p q) ((c,t), (d,s)) = conj (msubst p ((c,t),(d,s))) (msubst q ((c,t),(d,s)))" +| "msubst (Or p q) ((c,t), (d,s)) = disj (msubst p ((c,t),(d,s))) (msubst q ((c,t), (d,s)))" +| "msubst (Eq (CNP 0 a r)) ((c,t),(d,s)) = msubsteq c t d s a r" +| "msubst (NEq (CNP 0 a r)) ((c,t),(d,s)) = msubstneq c t d s a r" +| "msubst (Lt (CNP 0 a r)) ((c,t),(d,s)) = msubstlt c t d s a r" +| "msubst (Le (CNP 0 a r)) ((c,t),(d,s)) = msubstle c t d s a r" +| "msubst p ((c,t),(d,s)) = p" + +lemma msubst_I: assumes lp: "islin p" and nc: "isnpoly c" and nd: "isnpoly d" + shows "Ifm vs (x#bs) (msubst p ((c,t),(d,s))) = Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) p" + using lp +by (induct p rule: islin.induct, auto simp add: tmbound0_I[where b="(- (Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>)) /(1 + 1)" and b'=x and bs = bs and vs=vs] bound0_I[where b="(- (Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>)) /(1 + 1)" and b'=x and bs = bs and vs=vs] msubsteq msubstneq msubstlt[OF nc nd] msubstle[OF nc nd]) + +lemma msubst_nb: assumes lp: "islin p" and t: "tmbound0 t" and s: "tmbound0 s" + shows "bound0 (msubst p ((c,t),(d,s)))" + using lp t s + by (induct p rule: islin.induct, auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb) + +lemma fr_eq_msubst: + assumes lp: "islin p" + shows "(\ x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \ (Ifm vs (x#bs) (plusinf p)) \ (\ (c,t) \ set (uset p). \ (d,s) \ set (uset p). Ifm vs (x#bs) (msubst p ((c,t),(d,s)))))" + (is "(\ x. ?I x p) = (?M \ ?P \ ?F)" is "?E = ?D") +proof- +from uset_l[OF lp] have th: "\(c, s)\set (uset p). isnpoly c \ tmbound0 s" by blast +{fix c t d s assume ctU: "(c,t) \set (uset p)" and dsU: "(d,s) \set (uset p)" + and pts: "Ifm vs ((- Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p" + from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d" by simp_all + from msubst_I[OF lp norm, of vs x bs t s] pts + have "Ifm vs (x # bs) (msubst p ((c, t), d, s))" ..} +moreover +{fix c t d s assume ctU: "(c,t) \set (uset p)" and dsU: "(d,s) \set (uset p)" + and pts: "Ifm vs (x # bs) (msubst p ((c, t), d, s))" + from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d" by simp_all + from msubst_I[OF lp norm, of vs x bs t s] pts + have "Ifm vs ((- Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p" ..} +ultimately have th': "(\ (c,t) \ set (uset p). \ (d,s) \ set (uset p). Ifm vs ((- Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p) \ ?F" by blast +from fr_eq[OF lp, of vs bs x, simplified th'] show ?thesis . +qed + +text {* Rest of the implementation *} + +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 \ set xs. \y\ set xs. 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 x y 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) + +lemma simpfm_lin: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" + shows "qfree p \ islin (simpfm p)" + by (induct p rule: simpfm.induct, auto simp add: conj_lin disj_lin) + +definition + "ferrack p \ let q = simpfm p ; mp = minusinf q ; pp = plusinf q + in if (mp = T \ pp = T) then T + else (let U = alluopairs (remdps (uset q)) + in decr0 (disj mp (disj pp (evaldjf (simpfm o (msubst q)) U ))))" + +lemma ferrack: + assumes qf: "qfree p" + shows "qfree (ferrack p) \ ((Ifm vs bs (ferrack p)) = (Ifm vs bs (E p)))" + (is "_ \ (?rhs = ?lhs)") +proof- + let ?I = "\ x p. Ifm vs (x#bs) p" + let ?N = "\ t. Ipoly vs t" + let ?Nt = "\x t. Itm vs (x#bs) t" + let ?q = "simpfm p" + let ?U = "remdps(uset ?q)" + let ?Up = "alluopairs ?U" + let ?mp = "minusinf ?q" + let ?pp = "plusinf ?q" + let ?I = "\p. Ifm vs (x#bs) p" + from simpfm_lin[OF qf] simpfm_qf[OF qf] have lq: "islin ?q" and q_qf: "qfree ?q" . + from minusinf_nb[OF lq] plusinf_nb[OF lq] have mp_nb: "bound0 ?mp" and pp_nb: "bound0 ?pp" . + from bound0_qf[OF mp_nb] bound0_qf[OF pp_nb] have mp_qf: "qfree ?mp" and pp_qf: "qfree ?pp" . + from uset_l[OF lq] have U_l: "\(c, s)\set ?U. isnpoly c \ c \ 0\<^sub>p \ tmbound0 s \ allpolys isnpoly s" + by simp + {fix c t d s assume ctU: "(c,t) \ set ?U" and dsU: "(d,s) \ set ?U" + from U_l ctU dsU have norm: "isnpoly c" "isnpoly d" by auto + from msubst_I[OF lq norm, of vs x bs t s] msubst_I[OF lq norm(2,1), of vs x bs s t] + have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))" by (simp add: ring_simps)} + hence th0: "\x \ set ?U. \y \ set ?U. ?I (msubst ?q (x, y)) \ ?I (msubst ?q (y, x))" by clarsimp + {fix x assume xUp: "x \ set ?Up" + then obtain c t d s where ctU: "(c,t) \ set ?U" and dsU: "(d,s) \ set ?U" + and x: "x = ((c,t),(d,s))" using alluopairs_set1[of ?U] by auto + from U_l[rule_format, OF ctU] U_l[rule_format, OF dsU] + have nbs: "tmbound0 t" "tmbound0 s" by simp_all + from simpfm_bound0[OF msubst_nb[OF lq nbs, of c d]] + have "bound0 ((simpfm o (msubst (simpfm p))) x)" using x by simp} + with evaldjf_bound0[of ?Up "(simpfm o (msubst (simpfm p)))"] + have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast + with mp_nb pp_nb + have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by (simp add: disj_nb) + from decr0_qf[OF th1] have thqf: "qfree (ferrack p)" by (simp add: ferrack_def Let_def) + have "?lhs \ (\x. Ifm vs (x#bs) ?q)" by simp + also have "\ \ ?I ?mp \ ?I ?pp \ (\(c, t)\set ?U. \(d, s)\set ?U. ?I (msubst (simpfm p) ((c, t), d, s)))" using fr_eq_msubst[OF lq, of vs bs x] by simp + also have "\ \ ?I ?mp \ ?I ?pp \ (\ (x,y) \ set ?Up. ?I ((simpfm o (msubst ?q)) (x,y)))" using alluopairs_ex[OF th0] by simp + also have "\ \ ?I ?mp \ ?I ?pp \ ?I (evaldjf (simpfm o (msubst ?q)) ?Up)" + by (simp add: evaldjf_ex) + also have "\ \ ?I (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up)))" by simp + also have "\ \ ?rhs" using decr0[OF th1, of vs x bs] + apply (simp add: ferrack_def Let_def) + by (cases "?mp = T \ ?pp = T", auto) + finally show ?thesis using thqf by blast +qed + +definition "frpar p = simpfm (qelim p ferrack)" +lemma frpar: "qfree (frpar p) \ (Ifm vs bs (frpar p) \ Ifm vs bs p)" +proof- + from ferrack have th: "\bs p. qfree p \ qfree (ferrack p) \ Ifm vs bs (ferrack p) = Ifm vs bs (E p)" by blast + from qelim[OF th, of p bs] show ?thesis unfolding frpar_def by auto +qed + +declare polyadd.simps[code] +lemma [simp,code]: "polyadd (CN c n p, CN c' n' p') = + (if n < n' then CN (polyadd(c,CN c' n' p')) n p + else if n'p then cc' else CN cc' n pp')))" + by (simp add: Let_def stupid) + + + +(* +lemmas [code func] = polysub_def +lemmas [code func del] = Zero_nat_def +code_gen "frpar" in SML to FRParTest +*) + +section{* Second implemenation: Case splits not local *} + +lemma fr_eq2: assumes lp: "islin p" + shows "(\ x. Ifm vs (x#bs) p) \ + ((Ifm vs (x#bs) (minusinf p)) \ (Ifm vs (x#bs) (plusinf p)) \ + (Ifm vs (0#bs) p) \ + (\ (n,t) \ set (uset p). Ipoly vs n \ 0 \ Ifm vs ((- Itm vs (x#bs) t / (Ipoly vs n * (1 + 1)))#bs) p) \ + (\ (n,t) \ set (uset p). \ (m,s) \ set (uset p). Ipoly vs n \ 0 \ Ipoly vs m \ 0 \ Ifm vs (((- Itm vs (x#bs) t / Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /(1 + 1))#bs) p))" + (is "(\ x. ?I x p) = (?M \ ?P \ ?Z \ ?U \ ?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 inf_uset[OF lp nmi npi, OF px] + obtain c t d s where ct: "(c,t) \ set (uset p)" "(d,s) \ set (uset p)" "?I ((- Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>) / ((1\'a) + (1\'a))) p" + by auto + let ?c = "\c\\<^sub>p\<^bsup>vs\<^esup>" + let ?d = "\d\\<^sub>p\<^bsup>vs\<^esup>" + let ?s = "Itm vs (x # bs) s" + let ?t = "Itm vs (x # bs) t" + have eq2: "\(x::'a). x + x = (1 + 1) * x" + by (simp add: ring_simps) + {assume "?c = 0 \ ?d = 0" + with ct have ?D by simp} + moreover + {assume z: "?c = 0" "?d \ 0" + from z have ?D using ct by auto} + moreover + {assume z: "?c \ 0" "?d = 0" + with ct have ?D by auto } + moreover + {assume z: "?c \ 0" "?d \ 0" + from z have ?F using ct + apply - apply (rule bexI[where x = "(c,t)"], simp_all) + by (rule bexI[where x = "(d,s)"], simp_all) + hence ?D by blast} + ultimately have ?D by auto} + ultimately show "?D" by blast +next + assume "?D" + moreover {assume m:"?M" from minusinf_ex[OF lp m] have "?E" .} + moreover {assume p: "?P" from plusinf_ex[OF lp p] have "?E" . } + moreover {assume f:"?F" hence "?E" by blast} + ultimately show "?E" by blast +qed + +definition "msubsteq2 c t a b = Eq (Add (Mul a t) (Mul c b))" +definition "msubstltpos c t a b = Lt (Add (Mul a t) (Mul c b))" +definition "msubstlepos c t a b = Le (Add (Mul a t) (Mul c b))" +definition "msubstltneg c t a b = Lt (Neg (Add (Mul a t) (Mul c b)))" +definition "msubstleneg c t a b = Le (Neg (Add (Mul a t) (Mul c b)))" + +lemma msubsteq2: + assumes nz: "Ipoly vs c \ 0" and l: "islin (Eq (CNP 0 a b))" + shows "Ifm vs (x#bs) (msubsteq2 c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Eq (CNP 0 a b))" (is "?lhs = ?rhs") + using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" , symmetric] + by (simp add: msubsteq2_def field_simps) + +lemma msubstltpos: + assumes nz: "Ipoly vs c > 0" and l: "islin (Lt (CNP 0 a b))" + shows "Ifm vs (x#bs) (msubstltpos c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Lt (CNP 0 a b))" (is "?lhs = ?rhs") + using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" , symmetric] + by (simp add: msubstltpos_def field_simps) + +lemma msubstlepos: + assumes nz: "Ipoly vs c > 0" and l: "islin (Le (CNP 0 a b))" + shows "Ifm vs (x#bs) (msubstlepos c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Le (CNP 0 a b))" (is "?lhs = ?rhs") + using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" , symmetric] + by (simp add: msubstlepos_def field_simps) + +lemma msubstltneg: + assumes nz: "Ipoly vs c < 0" and l: "islin (Lt (CNP 0 a b))" + shows "Ifm vs (x#bs) (msubstltneg c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Lt (CNP 0 a b))" (is "?lhs = ?rhs") + using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" , symmetric] + by (simp add: msubstltneg_def field_simps del: minus_add_distrib) + +lemma msubstleneg: + assumes nz: "Ipoly vs c < 0" and l: "islin (Le (CNP 0 a b))" + shows "Ifm vs (x#bs) (msubstleneg c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Le (CNP 0 a b))" (is "?lhs = ?rhs") + using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" , symmetric] + by (simp add: msubstleneg_def field_simps del: minus_add_distrib) + +fun msubstpos :: "fm \ poly \ tm \ fm" where + "msubstpos (And p q) c t = And (msubstpos p c t) (msubstpos q c t)" +| "msubstpos (Or p q) c t = Or (msubstpos p c t) (msubstpos q c t)" +| "msubstpos (Eq (CNP 0 a r)) c t = msubsteq2 c t a r" +| "msubstpos (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)" +| "msubstpos (Lt (CNP 0 a r)) c t = msubstltpos c t a r" +| "msubstpos (Le (CNP 0 a r)) c t = msubstlepos c t a r" +| "msubstpos p c t = p" + +lemma msubstpos_I: + assumes lp: "islin p" and pos: "Ipoly vs c > 0" + shows "Ifm vs (x#bs) (msubstpos p c t) = Ifm vs (Itm vs (x#bs) t / Ipoly vs c #bs) p" + using lp pos + by (induct p rule: islin.induct, auto simp add: msubsteq2 msubstltpos[OF pos] msubstlepos[OF pos] tmbound0_I[of _ vs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" bs x] bound0_I[of _ vs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps) + +fun msubstneg :: "fm \ poly \ tm \ fm" where + "msubstneg (And p q) c t = And (msubstneg p c t) (msubstneg q c t)" +| "msubstneg (Or p q) c t = Or (msubstneg p c t) (msubstneg q c t)" +| "msubstneg (Eq (CNP 0 a r)) c t = msubsteq2 c t a r" +| "msubstneg (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)" +| "msubstneg (Lt (CNP 0 a r)) c t = msubstltneg c t a r" +| "msubstneg (Le (CNP 0 a r)) c t = msubstleneg c t a r" +| "msubstneg p c t = p" + +lemma msubstneg_I: + assumes lp: "islin p" and pos: "Ipoly vs c < 0" + shows "Ifm vs (x#bs) (msubstneg p c t) = Ifm vs (Itm vs (x#bs) t / Ipoly vs c #bs) p" + using lp pos + by (induct p rule: islin.induct, auto simp add: msubsteq2 msubstltneg[OF pos] msubstleneg[OF pos] tmbound0_I[of _ vs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" bs x] bound0_I[of _ vs "Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps) + + +definition "msubst2 p c t = disj (conj (lt (CP (polyneg c))) (simpfm (msubstpos p c t))) (conj (lt (CP c)) (simpfm (msubstneg p c t)))" + +lemma msubst2: assumes lp: "islin p" and nc: "isnpoly c" and nz: "Ipoly vs c \ 0" + shows "Ifm vs (x#bs) (msubst2 p c t) = Ifm vs (Itm vs (x#bs) t / Ipoly vs c #bs) p" +proof- + let ?c = "Ipoly vs c" + from nc have anc: "allpolys isnpoly (CP c)" "allpolys isnpoly (CP (~\<^sub>p c))" + by (simp_all add: polyneg_norm) + from nz have "?c > 0 \ ?c < 0" by arith + moreover + {assume c: "?c < 0" + from c msubstneg_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"] + have ?thesis by (auto simp add: msubst2_def)} + moreover + {assume c: "?c > 0" + from c msubstpos_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"] + have ?thesis by (auto simp add: msubst2_def)} + ultimately show ?thesis by blast +qed + +term msubsteq2 +lemma msubsteq2_nb: "tmbound0 t \ islin (Eq (CNP 0 a r)) \ bound0 (msubsteq2 c t a r)" + by (simp add: msubsteq2_def) + +lemma msubstltpos_nb: "tmbound0 t \ islin (Lt (CNP 0 a r)) \ bound0 (msubstltpos c t a r)" + by (simp add: msubstltpos_def) +lemma msubstltneg_nb: "tmbound0 t \ islin (Lt (CNP 0 a r)) \ bound0 (msubstltneg c t a r)" + by (simp add: msubstltneg_def) + +lemma msubstlepos_nb: "tmbound0 t \ islin (Le (CNP 0 a r)) \ bound0 (msubstlepos c t a r)" + by (simp add: msubstlepos_def) +lemma msubstleneg_nb: "tmbound0 t \ islin (Le (CNP 0 a r)) \ bound0 (msubstleneg c t a r)" + by (simp add: msubstleneg_def) + +lemma msubstpos_nb: assumes lp: "islin p" and tnb: "tmbound0 t" + shows "bound0 (msubstpos p c t)" +using lp tnb +by (induct p c t rule: msubstpos.induct, auto simp add: msubsteq2_nb msubstltpos_nb msubstlepos_nb) + +lemma msubstneg_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t" + shows "bound0 (msubstneg p c t)" +using lp tnb +by (induct p c t rule: msubstneg.induct, auto simp add: msubsteq2_nb msubstltneg_nb msubstleneg_nb) + +lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t" + shows "bound0 (msubst2 p c t)" +using lp tnb +by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0) + +lemma of_int2: "of_int 2 = 1 + 1" +proof- + have "(2::int) = 1 + 1" by simp + hence "of_int 2 = of_int (1 + 1)" by simp + thus ?thesis unfolding of_int_add by simp +qed + +lemma of_int_minus2: "of_int (-2) = - (1 + 1)" +proof- + have th: "(-2::int) = - 2" by simp + show ?thesis unfolding th by (simp only: of_int_minus of_int2) +qed + + +lemma islin_qf: "islin p \ qfree p" + by (induct p rule: islin.induct, auto simp add: bound0_qf) +lemma fr_eq_msubst2: + assumes lp: "islin p" + shows "(\ x. Ifm vs (x#bs) p) \ ((Ifm vs (x#bs) (minusinf p)) \ (Ifm vs (x#bs) (plusinf p)) \ Ifm vs (x#bs) (subst0 (CP 0\<^sub>p) p) \ (\(n, t)\set (uset p). Ifm vs (x# bs) (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \ (\ (c,t) \ set (uset p). \ (d,s) \ set (uset p). Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))))" + (is "(\ x. ?I x p) = (?M \ ?P \ ?Pz \ ?PU \ ?F)" is "?E = ?D") +proof- + from uset_l[OF lp] have th: "\(c, s)\set (uset p). isnpoly c \ tmbound0 s" by blast + let ?I = "\p. Ifm vs (x#bs) p" + have n2: "isnpoly (C (-2,1))" by (simp add: isnpoly_def) + note eq0 = subst0[OF islin_qf[OF lp], of vs x bs "CP 0\<^sub>p", simplified] + + have eq1: "(\(n, t)\set (uset p). ?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \ (\(n, t)\set (uset p). \n\\<^sub>p\<^bsup>vs\<^esup> \ 0 \ Ifm vs (- Itm vs (x # bs) t / (\n\\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p)" + proof- + {fix n t assume H: "(n, t)\set (uset p)" "?I(msubst2 p (n *\<^sub>p C (-2, 1)) t)" + from H(1) th have "isnpoly n" by blast + hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" by (simp_all add: polymul_norm n2) + have nn': "allpolys isnpoly (CP (~\<^sub>p (n *\<^sub>p C (-2, 1))))" + by (simp add: polyneg_norm nn) + hence nn2: "\n *\<^sub>p(C (-2,1)) \\<^sub>p\<^bsup>vs\<^esup> \ 0" "\n \\<^sub>p\<^bsup>vs\<^esup> \ 0" using H(2) nn' nn + by (auto simp add: msubst2_def lt zero_less_mult_iff mult_less_0_iff) + from msubst2[OF lp nn nn2(1), of x bs t] + have "\n\\<^sub>p\<^bsup>vs\<^esup> \ 0 \ Ifm vs (- Itm vs (x # bs) t / (\n\\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p" + using H(2) nn2 by (simp add: of_int_minus2 del: minus_add_distrib)} + moreover + {fix n t assume H: "(n, t)\set (uset p)" "\n\\<^sub>p\<^bsup>vs\<^esup> \ 0" "Ifm vs (- Itm vs (x # bs) t / (\n\\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p" + from H(1) th have "isnpoly n" by blast + hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" "\n *\<^sub>p(C (-2,1)) \\<^sub>p\<^bsup>vs\<^esup> \ 0" + using H(2) by (simp_all add: polymul_norm n2) + from msubst2[OF lp nn, of x bs t] have "?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)" using H(2,3) by (simp add: of_int_minus2 del: minus_add_distrib)} + ultimately show ?thesis by blast + qed + have eq2: "(\ (c,t) \ set (uset p). \ (d,s) \ set (uset p). Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))) \ (\(n, t)\set (uset p). + \(m, s)\set (uset p). \n\\<^sub>p\<^bsup>vs\<^esup> \ 0 \ \m\\<^sub>p\<^bsup>vs\<^esup> \ 0 \ Ifm vs ((- Itm vs (x # bs) t / \n\\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \m\\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p)" + proof- + {fix c t d s assume H: "(c,t) \ set (uset p)" "(d,s) \ set (uset p)" + "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))" + from H(1,2) th have "isnpoly c" "isnpoly d" by blast+ + hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" + by (simp_all add: polymul_norm n2) + have stupid: "allpolys isnpoly (CP (~\<^sub>p (C (-2, 1) *\<^sub>p c *\<^sub>p d)))" "allpolys isnpoly (CP ((C (-2, 1) *\<^sub>p c *\<^sub>p d)))" + by (simp_all add: polyneg_norm nn) + have nn': "\(C (-2, 1) *\<^sub>p c*\<^sub>p d)\\<^sub>p\<^bsup>vs\<^esup> \ 0" "\c\\<^sub>p\<^bsup>vs\<^esup> \ 0" "\d\\<^sub>p\<^bsup>vs\<^esup> \ 0" + using H(3) by (auto simp add: msubst2_def lt[OF stupid(1)] lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff) + from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn' + have "\c\\<^sub>p\<^bsup>vs\<^esup> \ 0 \ \d\\<^sub>p\<^bsup>vs\<^esup> \ 0 \ Ifm vs ((- Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p" + apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib) + by (simp add: mult_commute)} + moreover + {fix c t d s assume H: "(c,t) \ set (uset p)" "(d,s) \ set (uset p)" + "\c\\<^sub>p\<^bsup>vs\<^esup> \ 0" "\d\\<^sub>p\<^bsup>vs\<^esup> \ 0" "Ifm vs ((- Itm vs (x # bs) t / \c\\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \d\\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p" + from H(1,2) th have "isnpoly c" "isnpoly d" by blast+ + hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\(C (-2, 1) *\<^sub>p c*\<^sub>p d)\\<^sub>p\<^bsup>vs\<^esup> \ 0" + using H(3,4) by (simp_all add: polymul_norm n2) + from msubst2[OF lp nn, of x bs ] H(3,4,5) + have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))" apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib)by (simp add: mult_commute)} + ultimately show ?thesis by blast + qed + from fr_eq2[OF lp, of vs bs x] show ?thesis + unfolding eq0 eq1 eq2 by blast +qed + +definition +"ferrack2 p \ let q = simpfm p ; mp = minusinf q ; pp = plusinf q + in if (mp = T \ pp = T) then T + else (let U = remdps (uset q) + in decr0 (list_disj [mp, pp, simpfm (subst0 (CP 0\<^sub>p) q), evaldjf (\(c,t). msubst2 q (c *\<^sub>p C (-2, 1)) t) U, + evaldjf (\((b,a),(d,c)). msubst2 q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs U)]))" + +definition "frpar2 p = simpfm (qelim (prep p) ferrack2)" + +lemma ferrack2: assumes qf: "qfree p" + shows "qfree (ferrack2 p) \ ((Ifm vs bs (ferrack2 p)) = (Ifm vs bs (E p)))" + (is "_ \ (?rhs = ?lhs)") +proof- + let ?J = "\ x p. Ifm vs (x#bs) p" + let ?N = "\ t. Ipoly vs t" + let ?Nt = "\x t. Itm vs (x#bs) t" + let ?q = "simpfm p" + let ?qz = "subst0 (CP 0\<^sub>p) ?q" + let ?U = "remdps(uset ?q)" + let ?Up = "alluopairs ?U" + let ?mp = "minusinf ?q" + let ?pp = "plusinf ?q" + let ?I = "\p. Ifm vs (x#bs) p" + from simpfm_lin[OF qf] simpfm_qf[OF qf] have lq: "islin ?q" and q_qf: "qfree ?q" . + from minusinf_nb[OF lq] plusinf_nb[OF lq] have mp_nb: "bound0 ?mp" and pp_nb: "bound0 ?pp" . + from bound0_qf[OF mp_nb] bound0_qf[OF pp_nb] have mp_qf: "qfree ?mp" and pp_qf: "qfree ?pp" . + from uset_l[OF lq] have U_l: "\(c, s)\set ?U. isnpoly c \ c \ 0\<^sub>p \ tmbound0 s \ allpolys isnpoly s" + by simp + have bnd0: "\x \ set ?U. bound0 ((\(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) x)" + proof- + {fix c t assume ct: "(c,t) \ set ?U" + hence tnb: "tmbound0 t" using U_l by blast + from msubst2_nb[OF lq tnb] + have "bound0 ((\(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))" by simp} + thus ?thesis by auto + qed + have bnd1: "\x \ set ?Up. bound0 ((\((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) x)" + proof- + {fix b a d c assume badc: "((b,a),(d,c)) \ set ?Up" + from badc U_l alluopairs_set1[of ?U] + have nb: "tmbound0 (Add (Mul d a) (Mul b c))" by auto + from msubst2_nb[OF lq nb] have "bound0 ((\((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) ((b,a),(d,c)))" by simp} + thus ?thesis by auto + qed + have stupid: "bound0 F" by simp + let ?R = "list_disj [?mp, ?pp, simpfm (subst0 (CP 0\<^sub>p) ?q), evaldjf (\(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) ?U, + evaldjf (\((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs ?U)]" + from subst0_nb[of "CP 0\<^sub>p" ?q] q_qf evaldjf_bound0[OF bnd1] evaldjf_bound0[OF bnd0] mp_nb pp_nb stupid + have nb: "bound0 ?R " + by (simp add: list_disj_def disj_nb0 simpfm_bound0) + let ?s = "\((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))" + + {fix b a d c assume baU: "(b,a) \ set ?U" and dcU: "(d,c) \ set ?U" + from U_l baU dcU have norm: "isnpoly b" "isnpoly d" "isnpoly (C (-2, 1))" + by auto (simp add: isnpoly_def) + have norm2: "isnpoly (C (-2, 1) *\<^sub>p b*\<^sub>p d)" "isnpoly (C (-2, 1) *\<^sub>p d*\<^sub>p b)" + using norm by (simp_all add: polymul_norm) + have stupid: "allpolys isnpoly (CP (C (-2, 1) *\<^sub>p b*\<^sub>p d))" "allpolys isnpoly (CP (C (-2, 1) *\<^sub>p d*\<^sub>p b))" "allpolys isnpoly (CP (~\<^sub>p(C (-2, 1) *\<^sub>p b*\<^sub>p d)))" "allpolys isnpoly (CP (~\<^sub>p(C (-2, 1) *\<^sub>p d*\<^sub>p b)))" + by (simp_all add: polyneg_norm norm2) + have "?I (msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) = ?I (msubst2 ?q (C (-2, 1) *\<^sub>p d*\<^sub>p b) (Add (Mul b c) (Mul d a)))" (is "?lhs \ ?rhs") + proof + assume H: ?lhs + hence z: "\C (-2, 1) *\<^sub>p b *\<^sub>p d\\<^sub>p\<^bsup>vs\<^esup> \ 0" "\C (-2, 1) *\<^sub>p d *\<^sub>p b\\<^sub>p\<^bsup>vs\<^esup> \ 0" + by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)] mult_less_0_iff zero_less_mult_iff) + from msubst2[OF lq norm2(1) z(1), of x bs] + msubst2[OF lq norm2(2) z(2), of x bs] H + show ?rhs by (simp add: ring_simps) + next + assume H: ?rhs + hence z: "\C (-2, 1) *\<^sub>p b *\<^sub>p d\\<^sub>p\<^bsup>vs\<^esup> \ 0" "\C (-2, 1) *\<^sub>p d *\<^sub>p b\\<^sub>p\<^bsup>vs\<^esup> \ 0" + by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)] mult_less_0_iff zero_less_mult_iff) + from msubst2[OF lq norm2(1) z(1), of x bs] + msubst2[OF lq norm2(2) z(2), of x bs] H + show ?lhs by (simp add: ring_simps) + qed} + hence th0: "\x \ set ?U. \y \ set ?U. ?I (?s (x, y)) \ ?I (?s (y, x))" + by clarsimp + + have "?lhs \ (\x. Ifm vs (x#bs) ?q)" by simp + also have "\ \ ?I ?mp \ ?I ?pp \ ?I (subst0 (CP 0\<^sub>p) ?q) \ (\(n,t) \ set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \ (\(b, a)\set ?U. \(d, c)\set ?U. ?I (msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))))" + using fr_eq_msubst2[OF lq, of vs bs x] by simp + also have "\ \ ?I ?mp \ ?I ?pp \ ?I (subst0 (CP 0\<^sub>p) ?q) \ (\(n,t) \ set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \ (\ x\set ?U. \ y \set ?U. ?I (?s (x,y)))" + by (simp add: split_def) + also have "\ \ ?I ?mp \ ?I ?pp \ ?I (subst0 (CP 0\<^sub>p) ?q) \ (\(n,t) \ set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \ (\ (x,y) \ set ?Up. ?I (?s (x,y)))" + using alluopairs_ex[OF th0] by simp + also have "\ \ ?I ?R" + by (simp add: list_disj_def evaldjf_ex split_def) + also have "\ \ ?rhs" + unfolding ferrack2_def + apply (cases "?mp = T") + apply (simp add: list_disj_def) + apply (cases "?pp = T") + apply (simp add: list_disj_def) + by (simp_all add: Let_def decr0[OF nb]) + finally show ?thesis using decr0_qf[OF nb] + by (simp add: ferrack2_def Let_def) +qed + +lemma frpar2: "qfree (frpar2 p) \ (Ifm vs bs (frpar2 p) \ Ifm vs bs p)" +proof- + from ferrack2 have th: "\bs p. qfree p \ qfree (ferrack2 p) \ Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)" by blast + from qelim[OF th, of "prep p" bs] +show ?thesis unfolding frpar2_def by (auto simp add: prep) +qed + +code_module FRPar + contains + frpar = "frpar" + frpar2 = "frpar2" + test = "%x . frpar (E(Lt (Mul 1\<^sub>p (Bound 0))))" + +ML{* + +structure ReflectedFRPar = +struct + +val bT = HOLogic.boolT; +fun num rT x = HOLogic.mk_number rT x; +fun rrelT rT = [rT,rT] ---> rT; +fun rrT rT = [rT, rT] ---> bT; +fun divt rT = Const(@{const_name "HOL.divide"},rrelT rT); +fun timest rT = Const(@{const_name "HOL.times"},rrelT rT); +fun plust rT = Const(@{const_name "HOL.plus"},rrelT rT); +fun minust rT = Const(@{const_name "HOL.minus"},rrelT rT); +fun uminust rT = Const(@{const_name "HOL.uminus"}, rT --> rT); +fun powt rT = Const(@{const_name "power"}, [rT,@{typ "nat"}] ---> rT); +val brT = [bT, bT] ---> bT; +val nott = @{term "Not"}; +val conjt = @{term "op &"}; +val disjt = @{term "op |"}; +val impt = @{term "op -->"}; +val ifft = @{term "op = :: bool => _"} +fun llt rT = Const(@{const_name "HOL.less"},rrT rT); +fun lle rT = Const(@{const_name "HOL.less"},rrT rT); +fun eqt rT = Const("op =",rrT rT); +fun rz rT = Const(@{const_name "HOL.zero"},rT); + +fun dest_nat t = case t of + Const ("Suc",_)$t' => 1 + dest_nat t' +| _ => (snd o HOLogic.dest_number) t; + +fun num_of_term m t = + case t of + Const(@{const_name "uminus"},_)$t => FRPar.Neg (num_of_term m t) + | Const(@{const_name "HOL.plus"},_)$a$b => FRPar.Add (num_of_term m a, num_of_term m b) + | Const(@{const_name "HOL.minus"},_)$a$b => FRPar.Sub (num_of_term m a, num_of_term m b) + | Const(@{const_name "HOL.times"},_)$a$b => FRPar.Mul (num_of_term m a, num_of_term m b) + | Const(@{const_name "power"},_)$a$n => FRPar.Pw (num_of_term m a, dest_nat n) + | Const(@{const_name "HOL.divide"},_)$a$b => FRPar.C (HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd) + | _ => (FRPar.C (HOLogic.dest_number t |> snd,1) + handle TERM _ => FRPar.Bound (AList.lookup (op aconv) m t |> valOf)); + +fun tm_of_term m m' t = + case t of + Const(@{const_name "uminus"},_)$t => FRPar.tm_Neg (tm_of_term m m' t) + | Const(@{const_name "HOL.plus"},_)$a$b => FRPar.tm_Add (tm_of_term m m' a, tm_of_term m m' b) + | Const(@{const_name "HOL.minus"},_)$a$b => FRPar.tm_Sub (tm_of_term m m' a, tm_of_term m m' b) + | Const(@{const_name "HOL.times"},_)$a$b => FRPar.tm_Mul (num_of_term m' a, tm_of_term m m' b) + | _ => (FRPar.CP (num_of_term m' t) + handle TERM _ => FRPar.tm_Bound (AList.lookup (op aconv) m t |> valOf) + | Option => FRPar.tm_Bound (AList.lookup (op aconv) m t |> valOf)); + +fun term_of_num T m t = + case t of + FRPar.C (a,b) => (if b = 1 then num T a else if b=0 then (rz T) + else (divt T) $ num T a $ num T b) +| FRPar.Bound i => AList.lookup (op = : int*int -> bool) m i |> valOf +| FRPar.Add(a,b) => (plust T)$(term_of_num T m a)$(term_of_num T m b) +| FRPar.Mul(a,b) => (timest T)$(term_of_num T m a)$(term_of_num T m b) +| FRPar.Sub(a,b) => (minust T)$(term_of_num T m a)$(term_of_num T m b) +| FRPar.Neg a => (uminust T)$(term_of_num T m a) +| FRPar.Pw(a,n) => (powt T)$(term_of_num T m t)$(HOLogic.mk_number HOLogic.natT n) +| FRPar.CN(c,n,p) => term_of_num T m (FRPar.Add(c,FRPar.Mul(FRPar.Bound n, p))) +| _ => error "term_of_num: Unknown term"; + +fun term_of_tm T m m' t = + case t of + FRPar.CP p => term_of_num T m' p +| FRPar.tm_Bound i => AList.lookup (op = : int*int -> bool) m i |> valOf +| FRPar.tm_Add(a,b) => (plust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b) +| FRPar.tm_Mul(a,b) => (timest T)$(term_of_num T m' a)$(term_of_tm T m m' b) +| FRPar.tm_Sub(a,b) => (minust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b) +| FRPar.tm_Neg a => (uminust T)$(term_of_tm T m m' a) +| FRPar.CNP(n,c,p) => term_of_tm T m m' (FRPar.tm_Add(FRPar.tm_Mul(c, FRPar.tm_Bound n), p)) +| _ => error "term_of_tm: Unknown term"; + +fun fm_of_term m m' fm = + case fm of + Const("True",_) => FRPar.T + | Const("False",_) => FRPar.F + | Const("Not",_)$p => FRPar.NOT (fm_of_term m m' p) + | Const("op &",_)$p$q => FRPar.And(fm_of_term m m' p, fm_of_term m m' q) + | Const("op |",_)$p$q => FRPar.Or(fm_of_term m m' p, fm_of_term m m' q) + | Const("op -->",_)$p$q => FRPar.Imp(fm_of_term m m' p, fm_of_term m m' q) + | Const("op =",ty)$p$q => + if domain_type ty = bT then FRPar.Iff(fm_of_term m m' p, fm_of_term m m' q) + else FRPar.Eq (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q)) + | Const(@{const_name "HOL.less"},_)$p$q => + FRPar.Lt (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q)) + | Const(@{const_name "HOL.less_eq"},_)$p$q => + FRPar.Le (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q)) + | Const("Ex",_)$Abs(xn,xT,p) => + let val (xn', p') = variant_abs (xn,xT,p) + val x = Free(xn',xT) + fun incr i = i + 1 + val m0 = (x,0):: (map (apsnd incr) m) + in FRPar.E (fm_of_term m0 m' p') end + | Const("All",_)$Abs(xn,xT,p) => + let val (xn', p') = variant_abs (xn,xT,p) + val x = Free(xn',xT) + fun incr i = i + 1 + val m0 = (x,0):: (map (apsnd incr) m) + in FRPar.A (fm_of_term m0 m' p') end + | _ => error "fm_of_term"; + + +fun term_of_fm T m m' t = + case t of + FRPar.T => Const("True",bT) + | FRPar.F => Const("False",bT) + | FRPar.NOT p => nott $ (term_of_fm T m m' p) + | FRPar.And (p,q) => conjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q) + | FRPar.Or (p,q) => disjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q) + | FRPar.Imp (p,q) => impt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q) + | FRPar.Iff (p,q) => ifft $ (term_of_fm T m m' p) $ (term_of_fm T m m' q) + | FRPar.Lt p => (llt T) $ (term_of_tm T m m' p) $ (rz T) + | FRPar.Le p => (lle T) $ (term_of_tm T m m' p) $ (rz T) + | FRPar.Eq p => (eqt T) $ (term_of_tm T m m' p) $ (rz T) + | FRPar.NEq p => nott $ ((eqt T) $ (term_of_tm T m m' p) $ (rz T)) + | _ => error "term_of_fm: quantifiers!!!!???"; + +fun frpar_oracle (T,m, m', fm) = + let + val t = HOLogic.dest_Trueprop fm + val im = 0 upto (length m - 1) + val im' = 0 upto (length m' - 1) + in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm T (im ~~ m) (im' ~~ m') + (FRPar.frpar (fm_of_term (m ~~ im) (m' ~~ im') t)))) + end; + +fun frpar_oracle2 (T,m, m', fm) = + let + val t = HOLogic.dest_Trueprop fm + val im = 0 upto (length m - 1) + val im' = 0 upto (length m' - 1) + in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm T (im ~~ m) (im' ~~ m') + (FRPar.frpar2 (fm_of_term (m ~~ im) (m' ~~ im') t)))) + end; + +end; + + +*} + +oracle frpar_oracle = {* fn (ty, ts, ts', ct) => + let + val thy = Thm.theory_of_cterm ct + in cterm_of thy (ReflectedFRPar.frpar_oracle (ty,ts, ts', term_of ct)) + end *} + +oracle frpar_oracle2 = {* fn (ty, ts, ts', ct) => + let + val thy = Thm.theory_of_cterm ct + in cterm_of thy (ReflectedFRPar.frpar_oracle2 (ty,ts, ts', term_of ct)) + end *} + +ML{* +structure FRParTac = +struct + +fun frpar_tac T ps ctxt i = + (ObjectLogic.full_atomize_tac i) + THEN (fn st => + let + val g = List.nth (cprems_of st, i - 1) + val thy = ProofContext.theory_of ctxt + val fs = subtract (op aconv) (map Free (Term.add_frees (term_of g) [])) ps + val th = frpar_oracle (T, fs,ps, (* Pattern.eta_long [] *)g) + in rtac (th RS iffD2) i st end); + +fun frpar2_tac T ps ctxt i = + (ObjectLogic.full_atomize_tac i) + THEN (fn st => + let + val g = List.nth (cprems_of st, i - 1) + val thy = ProofContext.theory_of ctxt + val fs = subtract (op aconv) (map Free (Term.add_frees (term_of g) [])) ps + val th = frpar_oracle2 (T, fs,ps, (* Pattern.eta_long [] *)g) + in rtac (th RS iffD2) i st end); + +end; + +*} + +method_setup frpar = {* +let + fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () + fun simple_keyword k = Scan.lift (Args.$$$ k) >> K () + val parsN = "pars" + val typN = "type" + val any_keyword = keyword parsN || keyword typN + val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat + val cterms = thms >> map Drule.dest_term; + val terms = Scan.repeat (Scan.unless any_keyword Args.term) + val typ = Scan.unless any_keyword Args.typ +in + (keyword typN |-- typ) -- (keyword parsN |-- terms) >> + (fn (T,ps) => fn ctxt => SIMPLE_METHOD' (FRParTac.frpar_tac T ps ctxt)) +end +*} "Parametric QE for linear Arithmetic over fields, Version 1" + +method_setup frpar2 = {* +let + fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K () + fun simple_keyword k = Scan.lift (Args.$$$ k) >> K () + val parsN = "pars" + val typN = "type" + val any_keyword = keyword parsN || keyword typN + val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat + val cterms = thms >> map Drule.dest_term; + val terms = Scan.repeat (Scan.unless any_keyword Args.term) + val typ = Scan.unless any_keyword Args.typ +in + (keyword typN |-- typ) -- (keyword parsN |-- terms) >> + (fn (T,ps) => fn ctxt => SIMPLE_METHOD' (FRParTac.frpar2_tac T ps ctxt)) +end +*} "Parametric QE for linear Arithmetic over fields, Version 2" + + +lemma "\(x::'a::{division_by_zero,ordered_field,number_ring}). y \ -1 \ (y + 1)*x < 0" + apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "y::'a::{division_by_zero,ordered_field,number_ring}") + apply (simp add: ring_simps) + apply (rule spec[where x=y]) + apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "z::'a::{division_by_zero,ordered_field,number_ring}") + by simp + +text{* Collins/Jones Problem *} +(* +lemma "\(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \ r < 1 \ 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \ (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0" +proof- + have "(\(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \ r < 1 \ 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \ (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \ (\(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \ r < 1 \ 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \ 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \ ?rhs") +by (simp add: ring_simps) +have "?rhs" + + apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "a::'a::{division_by_zero,ordered_field,number_ring}" "b::'a::{division_by_zero,ordered_field,number_ring}") + apply (simp add: ring_simps) +oops +*) +(* +lemma "ALL (x::'a::{division_by_zero,ordered_field,number_ring}) y. (1 - t)*x \ (1+t)*y \ (1 - t)*y \ (1+t)*x --> 0 \ y" +apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "t::'a::{division_by_zero,ordered_field,number_ring}") +oops +*) + +lemma "\(x::'a::{division_by_zero,ordered_field,number_ring}). y \ -1 \ (y + 1)*x < 0" + apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "y::'a::{division_by_zero,ordered_field,number_ring}") + apply (simp add: ring_simps) + apply (rule spec[where x=y]) + apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "z::'a::{division_by_zero,ordered_field,number_ring}") + by simp + +text{* Collins/Jones Problem *} + +(* +lemma "\(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \ r < 1 \ 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \ (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0" +proof- + have "(\(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \ r < 1 \ 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \ (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \ (\(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \ r < 1 \ 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \ 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \ ?rhs") +by (simp add: ring_simps) +have "?rhs" + apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "a::'a::{division_by_zero,ordered_field,number_ring}" "b::'a::{division_by_zero,ordered_field,number_ring}") + apply simp +oops +*) + +(* +lemma "ALL (x::'a::{division_by_zero,ordered_field,number_ring}) y. (1 - t)*x \ (1+t)*y \ (1 - t)*y \ (1+t)*x --> 0 \ y" +apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "t::'a::{division_by_zero,ordered_field,number_ring}") +apply (simp add: field_simps linorder_neq_iff[symmetric]) +apply ferrack +oops +*) +end \ No newline at end of file