diff --git a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy --- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy @@ -1,4148 +1,4013 @@ (* Title: HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Author: Amine Chaieb *) section \A formalization of Ferrante and Rackoff's procedure with polynomial parameters, see Paper in CALCULEMUS 2008\ theory Parametric_Ferrante_Rackoff imports Reflected_Multivariate_Polynomial Dense_Linear_Order DP_Library "HOL-Library.Code_Target_Numeral" begin subsection \Terms\ datatype (plugins del: size) tm = CP poly | Bound nat | Add tm tm | Mul poly tm | Neg tm | Sub tm tm | CNP nat poly tm instantiation tm :: size begin primrec size_tm :: "tm \ nat" where "size_tm (CP c) = polysize c" | "size_tm (Bound n) = 1" | "size_tm (Neg a) = 1 + size_tm a" | "size_tm (Add a b) = 1 + size_tm a + size_tm b" | "size_tm (Sub a b) = 3 + size_tm a + size_tm b" | "size_tm (Mul c a) = 1 + polysize c + size_tm a" | "size_tm (CNP n c a) = 3 + polysize c + size_tm a " instance .. end text \Semantics of terms tm.\ primrec Itm :: "'a::field_char_0 list \ 'a list \ tm \ 'a" where "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" primrec tmboundslt :: "nat \ tm \ bool" where "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 :: "tm \ bool" \ \a \tm\ is \<^emph>\independent\ of Bound 0\ where "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 "tmbound0 a" shows "Itm vs (b#bs) a = Itm vs (b'#bs) a" using assms by (induct a rule: tm.induct) auto primrec tmbound :: "nat \ tm \ bool" \ \a \tm\ is \<^emph>\independent\ of Bound n\ where "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: tm.induct) auto fun decrtm0 :: "tm \ tm" where "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" fun incrtm0 :: "tm \ tm" where "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 lemma incrtm0: "Itm vs (x#bs) (incrtm0 t) = Itm vs bs t" by (induct t rule: decrtm0.induct) simp_all primrec decrtm :: "nat \ tm \ tm" where "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))" primrec removen :: "nat \ 'a list \ 'a list" where "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 then show ?case by simp next case (Cons x xs) let ?l = "length (x # xs)" consider "n \ ?l" | "n < ?l" by arith then show ?case proof cases case 1 with removen_same[OF this] show ?thesis by simp next case nl: 2 consider "m < n" | "m \ n" by arith then show ?thesis proof cases case 1 then show ?thesis using Cons by (cases m) auto next case 2 consider "m \ ?l" | "m > ?l" by arith then show ?thesis proof cases case 1 then show ?thesis using Cons by (cases m) auto next case ml: 2 have th: "length (removen n (x # xs)) = length xs" using removen_length[where n = n and xs= "x # xs"] nl by simp with ml have "m \ length (removen n (x # xs))" by auto from th nth_length_exceeds[OF this] have "(removen n (x # xs))!m = [] ! (m - length xs)" by auto then have "(removen n (x # xs))!m = [] ! (m - (length (x # xs) - 1))" by auto then show ?thesis using ml nl by auto qed qed qed 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: tm.induct) (auto simp add: removen_nth) + using bnd nb nle by (induct t rule: tm.induct) (auto simp: removen_nth) primrec tmsubst0:: "tm \ tm \ tm" where "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: "Itm vs (x # bs) (tmsubst0 t a) = Itm vs (Itm vs (x # bs) t # bs) a" by (induct a rule: tm.induct) auto lemma tmsubst0_nb: "tmbound0 t \ tmbound0 (tmsubst0 t a)" by (induct a rule: tm.induct) auto primrec tmsubst:: "nat \ tm \ tm \ tm" where "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: tm.induct) auto lemma tmsubst_nb0: assumes tnb: "tmbound0 t" shows "tmbound0 (tmsubst 0 t a)" using tnb by (induct a rule: tm.induct) auto lemma tmsubst_nb: assumes tnb: "tmbound m t" shows "tmbound m (tmsubst m t a)" using tnb by (induct a rule: tm.induct) auto lemma incrtm0_tmbound: "tmbound n t \ tmbound (Suc n) (incrtm0 t)" by (induct t) auto text \Simplification.\ fun tmadd:: "tm \ tm \ tm" where "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) - apply (simp_all add: Let_def) - apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p") - apply (case_tac "n1 \ n2") - apply simp_all - apply (case_tac "n1 = n2") - apply (simp_all add: algebra_simps) - apply (simp only: distrib_left [symmetric] polyadd [symmetric]) - apply simp - done +proof (induct t s rule: tmadd.induct) + case (1 n1 c1 r1 n2 c2 r2) + show ?case + proof (cases "c1 +\<^sub>p c2 = 0\<^sub>p") + case 0: True + show ?thesis + proof (cases "n1 \ n2") + case True + with 0 show ?thesis + apply (simp add: 1) + by (metis INum_int(2) Ipoly.simps(1) comm_semiring_class.distrib mult_eq_0_iff polyadd) + qed (use 0 1 in auto) + next + case False + then show ?thesis + using 1 comm_semiring_class.distrib by auto + qed +qed auto lemma tmadd_nb0[simp]: "tmbound0 t \ tmbound0 s \ tmbound0 (tmadd t s)" - by (induct t s rule: tmadd.induct) (auto simp add: Let_def) + by (induct t s rule: tmadd.induct) (auto simp: 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) + by (induct t s rule: tmadd.induct) (auto simp: 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) + by (induct t s rule: tmadd.induct) (auto simp: 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) fun tmmul:: "tm \ poly \ tm" where "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: field_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) + by (induct t arbitrary: i rule: tmmul.induct) (auto simp: Let_def) lemma tmmul_allpolys_npoly[simp]: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "allpolys isnpoly t \ isnpoly c \ allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct) (simp_all add: Let_def polymul_norm) definition tmneg :: "tm \ tm" where "tmneg t \ tmmul t (C (- 1,1))" definition tmsub :: "tm \ tm \ tm" where "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] by simp 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))" by (simp add: isnpoly_def) lemma tmneg_allpolys_npoly[simp]: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "allpolys isnpoly t \ allpolys isnpoly (tmneg t)" by (auto simp: tmneg_def) 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::field_char_0)" shows "allpolys isnpoly t \ allpolys isnpoly s \ allpolys isnpoly (tmsub t s)" by (simp add: tmsub_def isnpoly_def) fun simptm :: "tm \ tm" where "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::field_char_0)" shows "polynate t = 0\<^sub>p \ Ipoly bs t = (0::'a)" - apply (subst polynate[symmetric]) - apply simp - done + by (metis INum_int(2) Ipoly.simps(1) polynate) lemma simptm_ci[simp]: "Itm vs bs (simptm t) = Itm vs bs t" - by (induct t rule: simptm.induct) (auto simp add: Let_def polynate_stupid) + by (induct t rule: simptm.induct) (auto simp: Let_def polynate_stupid) lemma simptm_tmbound0[simp]: "tmbound0 t \ tmbound0 (simptm t)" - by (induct t rule: simptm.induct) (auto simp add: Let_def) + by (induct t rule: simptm.induct) (auto simp: Let_def) lemma simptm_nb[simp]: "tmbound n t \ tmbound n (simptm t)" - by (induct t rule: simptm.induct) (auto simp add: Let_def) + by (induct t rule: simptm.induct) (auto simp: Let_def) lemma simptm_nlt[simp]: "tmboundslt n t \ tmboundslt n (simptm t)" - by (induct t rule: simptm.induct) (auto simp add: Let_def) + by (induct t rule: simptm.induct) (auto simp: 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::field_char_0)" shows "allpolys isnpoly (simptm p)" - by (induct p rule: simptm.induct) (auto simp add: Let_def) + by (induct p rule: simptm.induct) (auto simp: Let_def) declare let_cong[fundef_cong del] fun split0 :: "tm \ poly \ tm" where "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)" declare let_cong[fundef_cong] 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)"]) - apply simp - done + using prod.collapse by blast 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 field_simps) - apply (simp add: Let_def split_def field_simps) - apply (simp add: Let_def split_def field_simps) - apply (simp add: Let_def split_def field_simps) - apply (simp add: Let_def split_def field_simps) - apply (simp add: Let_def split_def mult.assoc distrib_left[symmetric]) - apply (simp add: Let_def split_def field_simps) - apply (simp add: Let_def split_def field_simps) - done +proof (induct t rule: split0.induct) + case (7 c t) + then show ?case + by (simp add: Let_def split_def mult.assoc flip: distrib_left) +qed (auto simp: Let_def split_def field_simps) 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')" then have "c' = fst (split0 t)" "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::field_char_0)" shows "split0 t = (c',t') \ tmbound 0 t'" proof - fix c' t' assume "split0 t = (c', t')" then have "c' = fst (split0 t)" "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::field_char_0)" 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) + using nb by (induct t rule: split0.induct) (auto simp: Let_def split_def) 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) + using nb by (induct t rule: split0.induct) (auto simp: Let_def split_def) 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) + by (induct t rule: split0.induct) (auto simp: Let_def split_def) 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) + by (induct t rule: split0.induct) (auto simp: Let_def split_def) 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) + by (induct t rule: split0.induct) (auto simp: Let_def split_def) 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) lemma isnpoly_fst_split0: assumes "SORT_CONSTRAINT('a::field_char_0)" 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) subsection \Formulae\ datatype (plugins del: size) 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 instantiation fm :: size begin primrec size_fm :: "fm \ nat" where "size_fm (Not p) = 1 + size_fm p" | "size_fm (And p q) = 1 + size_fm p + size_fm q" | "size_fm (Or p q) = 1 + size_fm p + size_fm q" | "size_fm (Imp p q) = 3 + size_fm p + size_fm q" | "size_fm (Iff p q) = 3 + 2 * (size_fm p + size_fm q)" | "size_fm (E p) = 1 + size_fm p" | "size_fm (A p) = 4 + size_fm p" | "size_fm T = 1" | "size_fm F = 1" | "size_fm (Le _) = 1" | "size_fm (Lt _) = 1" | "size_fm (Eq _) = 1" | "size_fm (NEq _) = 1" instance .. end lemma fmsize_pos [simp]: "size p > 0" for p :: fm by (induct p) simp_all text \Semantics of formulae (fm).\ primrec Ifm ::"'a::linordered_field list \ 'a list \ fm \ bool" where "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)" fun not:: "fm \ fm" where "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 definition conj :: "fm \ fm \ fm" where "conj p q \ (if p = F \ q = F then F else if p = T then q else if q = T then p else 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) definition disj :: "fm \ fm \ fm" where "disj p q \ (if (p = T \ q = T) then T else if p = F then q else if q = F then p 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) definition imp :: "fm \ fm \ fm" where "imp p q \ (if p = F \ q = T \ p = q then T else if p = T then q else if q = F then not p 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) definition iff :: "fm \ fm \ fm" where "iff p q \ (if p = q then T else if p = Not q \ Not p = q then F else if p = F then not q else if q = F then not p else if p = T then q else if q = T then p else Iff p q)" 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) text \Quantifier freeness.\ fun qfree:: "fm \ bool" where "qfree (E p) = False" | "qfree (A p) = False" | "qfree (Not p) = qfree p" | "qfree (And p q) = (qfree p \ qfree q)" | "qfree (Or p q) = (qfree p \ qfree q)" | "qfree (Imp p q) = (qfree p \ qfree q)" | "qfree (Iff p q) = (qfree p \ qfree q)" | "qfree p = True" text \Boundedness and substitution.\ primrec boundslt :: "nat \ fm \ bool" where "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" fun bound0:: "fm \ bool" \ \a formula is independent of Bound 0\ where "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 primrec bound:: "nat \ fm \ bool" \ \a formula is independent of Bound n\ where "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: fm.induct) case (E p bs n) have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" for y proof - from E 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] show ?thesis . qed then show ?case by simp next case (A p bs n) have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" for y proof - from A have bnd: "boundslt (length (y#bs)) p" and nb: "bound (Suc n) p" and le: "Suc n \ length (y#bs)" by simp_all from A.hyps[OF bnd nb le tmbound_I] show ?thesis . qed then show ?case by simp qed auto fun decr0 :: "fm \ fm" where "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 "bound0 p" shows "Ifm vs (x#bs) p = Ifm vs bs (decr0 p)" using assms by (induct p rule: decr0.induct) (simp_all add: decrtm0) primrec decr :: "nat \ fm \ fm" where "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: fm.induct) case (E p bs m) have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p" for x proof - from E have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p" and nle: "Suc m < length (x#bs)" by auto from E(1)[OF bnd nb nle] show ?thesis . qed then show ?case by auto next case (A p bs m) have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p" for x proof - from A have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p" and nle: "Suc m < length (x#bs)" by auto from A(1)[OF bnd nb nle] show ?thesis . qed then show ?case by auto -qed (auto simp add: decrtm removen_nth) +qed (auto simp: decrtm removen_nth) primrec subst0 :: "tm \ fm \ fm" where "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: fm.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: fm.induct) auto primrec subst:: "nat \ tm \ fm \ fm" where "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: fm.induct) case (E p bs n) have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p" for x proof - from E have bn: "boundslt (length (x#bs)) p" by simp from E have nlm: "Suc n \ length (x#bs)" by simp from E(1)[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 then show ?thesis by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) qed then show ?case by simp next case (A p bs n) have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p" for x proof - from A have bn: "boundslt (length (x#bs)) p" by simp from A have nlm: "Suc n \ length (x#bs)" by simp from A(1)[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 then show ?thesis by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) qed then show ?case by simp -qed (auto simp add: tmsubst) +qed (auto simp: tmsubst) lemma subst_nb: assumes "tmbound m t" shows "bound m (subst m t p)" using assms tmsubst_nb incrtm0_tmbound by (induct p arbitrary: m t rule: fm.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)" unfolding iff_def by (cases "p = q") auto lemma iff_nb0[simp]: "bound0 p \ bound0 q \ bound0 (iff p q)" using iff_def unfolding iff_def by (cases "p = q") auto lemma iff_nb[simp]: "bound n p \ bound n q \ bound n (iff p q)" using iff_def unfolding iff_def by (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 fun isatom :: "fm \ bool" \ \test for atomicity\ where "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 definition djf :: "('a \ fm) \ 'a \ fm \ fm" where "djf f p q \ (if q = T then T else if q = F then f p else (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q))" definition evaldjf :: "('a \ fm) \ 'a list \ fm" where "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)" - apply (cases "q = T") - apply (simp add: djf_def) - apply (cases "q = F") - apply (simp add: djf_def) - apply (cases "f p") - apply (simp_all add: Let_def djf_def) - done + by (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 "\x\ set xs. bound0 (f x)" + assumes "\x \ set xs. bound0 (f x)" shows "bound0 (evaldjf f xs)" using assms - apply (induct xs) - apply (auto simp add: evaldjf_def djf_def Let_def) - apply (case_tac "f a") - apply auto - done +proof (induct xs) + case Nil + then show ?case + by (simp add: evaldjf_def) +next + case (Cons a xs) + then show ?case + by (cases "f a") (simp_all add: evaldjf_def djf_def Let_def) +qed lemma evaldjf_qf: assumes "\x\ set xs. qfree (f x)" shows "qfree (evaldjf f xs)" using assms - apply (induct xs) - apply (auto simp add: evaldjf_def djf_def Let_def) - apply (case_tac "f a") - apply auto - done +proof (induct xs) + case Nil + then show ?case + by (simp add: evaldjf_def) +next + case (Cons a xs) + then show ?case + by (cases "f a") (simp_all add: evaldjf_def djf_def Let_def) +qed fun disjuncts :: "fm \ fm list" where "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: assumes "bound0 p" shows "\q \ set (disjuncts p). bound0 q" proof - from assms have "list_all bound0 (disjuncts p)" by (induct p rule: disjuncts.induct) auto then show ?thesis by (simp only: list_all_iff) qed lemma disjuncts_qf: assumes "qfree p" shows "\q \ set (disjuncts p). qfree q" proof - from assms have "list_all qfree (disjuncts p)" by (induct p rule: disjuncts.induct) auto then show ?thesis by (simp only: list_all_iff) qed definition DJ :: "(fm \ fm) \ fm \ fm" where "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 fun conjuncts :: "fm \ fm list" where "conjuncts (And p q) = conjuncts p @ conjuncts q" | "conjuncts T = []" | "conjuncts p = [p]" definition list_conj :: "fm list \ fm" where "list_conj ps \ foldr conj ps T" definition CJNB :: "(fm \ fm) \ fm \ fm" where "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" then have "list_all qfree (conjuncts p)" by (induct p rule: conjuncts.induct) auto then show ?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: assumes "bound0 p" shows "\q \ set (conjuncts p). bound0 q" proof - from assms have "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct) auto then show ?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, rename_tac nat a b, 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) - apply simp_all - apply (rename_tac poly, case_tac poly) - apply (simp_all add: isnpoly_def) - done + by (auto simp: lt_def isnpoly_def split: tm.split poly.split) lemma le: "allpolys isnpoly p \ Ifm vs bs (le p) = Ifm vs bs (Le p)" - apply (simp add: le_def) - apply (cases p) - apply simp_all - apply (rename_tac poly, case_tac poly) - apply (simp_all add: isnpoly_def) - done + by (auto simp: le_def isnpoly_def split: tm.split poly.split) lemma eq: "allpolys isnpoly p \ Ifm vs bs (eq p) = Ifm vs bs (Eq p)" - apply (simp add: eq_def) - apply (cases p) - apply simp_all - apply (rename_tac poly, case_tac poly) - apply (simp_all add: isnpoly_def) - done + by (auto simp: eq_def isnpoly_def split: tm.split poly.split) 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) - apply simp_all - apply (rename_tac poly, case_tac poly) - apply simp_all - apply (rename_tac nat a b, case_tac nat) - apply simp_all - done + using islin_stupid + by(auto simp: lt_def isnpoly_def split: tm.split poly.split) lemma le_lin: "tmbound0 p \ islin (le p)" - apply (simp add: le_def) - apply (cases p) - apply simp_all - apply (rename_tac poly, case_tac poly) - apply simp_all - apply (rename_tac nat a b, case_tac nat) - apply simp_all - done + using islin_stupid + by(auto simp: le_def isnpoly_def split: tm.split poly.split) lemma eq_lin: "tmbound0 p \ islin (eq p)" - apply (simp add: eq_def) - apply (cases p) - apply simp_all - apply (rename_tac poly, case_tac poly) - apply simp_all - apply (rename_tac nat a b, case_tac nat) - apply simp_all - done + using islin_stupid + by(auto simp: eq_def isnpoly_def split: tm.split poly.split) lemma neq_lin: "tmbound0 p \ islin (neq p)" - apply (simp add: neq_def eq_def) - apply (cases p) - apply simp_all - apply (rename_tac poly, case_tac poly) - apply simp_all - apply (rename_tac nat a b, case_tac nat) - apply simp_all - done + using islin_stupid + by(auto simp: neq_def eq_def isnpoly_def split: tm.split poly.split) 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::field_char_0)" 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] + by (auto simp: 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::field_char_0)" 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] + by (auto simp: 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::field_char_0)" 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] + by (auto simp: 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::field_char_0)" 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] + by (auto simp: 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::field_char_0)" and *: "allpolys isnpoly t" shows "isnpoly (fst (split0 t))" and "allpolys isnpoly (snd (split0 t))" using * by (induct t rule: split0.induct) - (auto simp add: Let_def split_def polyadd_norm polymul_norm polyneg_norm + (auto simp: Let_def split_def polyadd_norm polymul_norm polyneg_norm polysub_norm really_stupid) lemma simplt[simp]: "Ifm vs bs (simplt t) = Ifm vs bs (Lt t)" proof - have *: "allpolys isnpoly (simptm t)" by simp let ?t = "simptm t" show ?thesis proof (cases "fst (split0 ?t) = 0\<^sub>p") case True then show ?thesis using split0[of "simptm t" vs bs] lt[OF split0_npoly(2)[OF *], of vs bs] by (simp add: simplt_def Let_def split_def lt) next case False then show ?thesis using split0[of "simptm t" vs bs] by (simp add: simplt_def Let_def split_def) qed qed lemma simple[simp]: "Ifm vs bs (simple t) = Ifm vs bs (Le t)" proof - have *: "allpolys isnpoly (simptm t)" by simp let ?t = "simptm t" show ?thesis proof (cases "fst (split0 ?t) = 0\<^sub>p") case True then show ?thesis using split0[of "simptm t" vs bs] le[OF split0_npoly(2)[OF *], of vs bs] by (simp add: simple_def Let_def split_def le) next case False then show ?thesis using split0[of "simptm t" vs bs] by (simp add: simple_def Let_def split_def) qed qed lemma simpeq[simp]: "Ifm vs bs (simpeq t) = Ifm vs bs (Eq t)" proof - have n: "allpolys isnpoly (simptm t)" by simp let ?t = "simptm t" show ?thesis proof (cases "fst (split0 ?t) = 0\<^sub>p") case True then show ?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) next case False then show ?thesis using split0[of "simptm t" vs bs] by (simp add: simpeq_def Let_def split_def) qed qed lemma simpneq[simp]: "Ifm vs bs (simpneq t) = Ifm vs bs (NEq t)" proof - have n: "allpolys isnpoly (simptm t)" by simp let ?t = "simptm t" show ?thesis proof (cases "fst (split0 ?t) = 0\<^sub>p") case True then show ?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) next case False then show ?thesis using split0[of "simptm t" vs bs] by (simp add: simpneq_def Let_def split_def) qed qed lemma lt_nb: "tmbound0 t \ bound0 (lt t)" - apply (simp add: lt_def) - apply (cases t) - apply auto - apply (rename_tac poly, case_tac poly) - apply auto - done + by(auto simp: lt_def split: tm.split poly.split) lemma le_nb: "tmbound0 t \ bound0 (le t)" - apply (simp add: le_def) - apply (cases t) - apply auto - apply (rename_tac poly, case_tac poly) - apply auto - done + by(auto simp: le_def split: tm.split poly.split) lemma eq_nb: "tmbound0 t \ bound0 (eq t)" - apply (simp add: eq_def) - apply (cases t) - apply auto - apply (rename_tac poly, case_tac poly) - apply auto - done + by(auto simp: eq_def split: tm.split poly.split) lemma neq_nb: "tmbound0 t \ bound0 (neq t)" - apply (simp add: neq_def eq_def) - apply (cases t) - apply auto - apply (rename_tac poly, case_tac poly) - apply auto - done - + by(auto simp: neq_def eq_def split: tm.split poly.split) + +(*THE FOLLOWING PROOFS MIGHT BE COMBINED*) lemma simplt_nb[simp]: - assumes "SORT_CONSTRAINT('a::field_char_0)" - shows "tmbound0 t \ bound0 (simplt t)" + assumes "SORT_CONSTRAINT('a::field_char_0)" and t: "tmbound0 t" + shows "bound0 (simplt t)" proof (simp add: simplt_def Let_def split_def) - assume "tmbound0 t" - then have *: "tmbound0 (simptm t)" - by simp + have *: "tmbound0 (simptm t)" + using t by simp let ?c = "fst (split0 (simptm t))" from tmbound_split0[OF *[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" . then show "(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::field_char_0)" - shows "tmbound0 t \ bound0 (simple t)" + assumes "SORT_CONSTRAINT('a::field_char_0)" and t: "tmbound0 t" + shows "bound0 (simple t)" proof(simp add: simple_def Let_def split_def) - assume "tmbound0 t" - then have *: "tmbound0 (simptm t)" - by simp + have *: "tmbound0 (simptm t)" + using t by simp let ?c = "fst (split0 (simptm t))" from tmbound_split0[OF *[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" . then show "(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::field_char_0)" - shows "tmbound0 t \ bound0 (simpeq t)" + assumes "SORT_CONSTRAINT('a::field_char_0)" and t: "tmbound0 t" + shows "bound0 (simpeq t)" proof (simp add: simpeq_def Let_def split_def) - assume "tmbound0 t" - then have *: "tmbound0 (simptm t)" - by simp + have *: "tmbound0 (simptm t)" + using t by simp let ?c = "fst (split0 (simptm t))" from tmbound_split0[OF *[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" . then show "(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::field_char_0)" - shows "tmbound0 t \ bound0 (simpneq t)" + assumes "SORT_CONSTRAINT('a::field_char_0)" and t: "tmbound0 t" + shows "bound0 (simpneq t)" proof (simp add: simpneq_def Let_def split_def) - assume "tmbound0 t" - then have *: "tmbound0 (simptm t)" - by simp + have *: "tmbound0 (simptm t)" + using t by simp let ?c = "fst (split0 (simptm t))" from tmbound_split0[OF *[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" . then show "(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 fun conjs :: "fm \ fm list" where "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 definition list_disj :: "fm list \ fm" where "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) + by (induct ps) (auto simp: 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) + by (induct ps) (auto simp: list_conj_def) 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) + by (induct ps) (auto simp: 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 +proof (induct p rule: conjs.induct) + case (1 p q) + then show ?case + unfolding conjs.simps bound.simps by fastforce +qed auto 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 - apply simp_all - done +proof (induct p rule: conjs.induct) + case (1 p q) + then show ?case + unfolding conjs.simps bound.simps by fastforce +qed auto lemma list_conj_boundslt: " \p\ set ps. boundslt n p \ boundslt n (list_conj ps)" by (induct ps) (auto simp: list_conj_def) lemma list_conj_nb: assumes "\p\ set ps. bound n p" shows "bound n (list_conj ps)" using assms by (induct ps) (auto simp: list_conj_def) lemma list_conj_nb': "\p\set ps. bound0 p \ bound0 (list_conj ps)" by (induct ps) (auto simp: list_conj_def) 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 then have yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb') then have 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 then have 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 split_def) have "Ifm vs bs p = ((Ifm vs bs ?cyes) \ (Ifm vs bs ?cno))" for bs proof - 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 show ?thesis using list_conj[of vs bs] by simp qed then have "Ifm vs bs (E p) = (\x. (Ifm vs (x#bs) ?cyes) \ (Ifm vs (x#bs) ?cno))" by simp also fix y 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] simp del: partition_filter_conv) + by (auto simp: decr0[OF yes_nb] simp del: partition_filter_conv) 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 fun simpfm :: "fm \ fm" where "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::field_char_0)" shows "bound0 p \ bound0 (simpfm p)" by (induct p rule: simpfm.induct) auto lemma lt_qf[simp]: "qfree (lt t)" - apply (cases t) - apply (auto simp add: lt_def) - apply (rename_tac poly, case_tac poly) - apply auto - done + by(auto simp: lt_def split: tm.split poly.split) lemma le_qf[simp]: "qfree (le t)" - apply (cases t) - apply (auto simp add: le_def) - apply (rename_tac poly, case_tac poly) - apply auto - done + by(auto simp: le_def split: tm.split poly.split) lemma eq_qf[simp]: "qfree (eq t)" - apply (cases t) - apply (auto simp add: eq_def) - apply (rename_tac poly, case_tac poly) - apply auto - done + by(auto simp: eq_def split: tm.split poly.split) 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 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::field_char_0)" shows "qfree p \ islin (simpfm p)" by (induct p rule: simpfm.induct) (simp_all add: conj_lin disj_lin) fun prep :: "fm \ fm" where "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" lemma prep: "Ifm vs bs (prep p) = Ifm vs bs p" by (induct p arbitrary: bs rule: prep.induct) auto text \Generic quantifier elimination.\ fun qelim :: "fm \ (fm \ fm) \ fm" where "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\ fun minusinf:: "fm \ fm" \ \virtual substitution of \-\\\ where "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" fun plusinf:: "fm \ fm" \ \virtual substitution of \+\\\ where "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 "islin p" shows "\z. \x < z. Ifm vs (x#bs) (minusinf p) \ Ifm vs (x#bs) p" using assms proof (induct p rule: minusinf.induct) - case 1 + case (1 p q) + then obtain zp zq where zp: "\xxxx - ?e" using neg_less_divide_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then show ?thesis using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto qed then show ?thesis by auto qed next case (4 c e) then have nbe: "tmbound0 e" by simp from 4 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" fix y let ?e = "Itm vs (y#bs) e" consider "?c = 0" | "?c > 0" | "?c < 0" by arith then show ?case proof cases case 1 then show ?thesis using eqs by auto next case c: 2 have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))" if "x < -?e / ?c" for x proof - from that have "?c * x < - ?e" using pos_less_divide_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then show ?thesis using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto qed then show ?thesis by auto next case c: 3 have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))" if "x < -?e / ?c" for x proof - from that have "?c * x > - ?e" using neg_less_divide_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then show ?thesis using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto qed then show ?thesis by auto qed next case (5 c e) then have nbe: "tmbound0 e" by simp from 5 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all then have 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" fix y let ?e = "Itm vs (y#bs) e" consider "?c = 0" | "?c > 0" | "?c < 0" by arith then show ?case proof cases case 1 then show ?thesis using eqs by auto next case c: 2 have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))" if "x < -?e / ?c" for x proof - from that have "?c * x < - ?e" using pos_less_divide_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then show ?thesis using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto qed then show ?thesis by auto next case c: 3 have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))" if "x < -?e / ?c" for x proof - from that have "?c * x > - ?e" using neg_less_divide_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then show ?thesis using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] c by auto qed then show ?thesis by auto qed next case (6 c e) then have nbe: "tmbound0 e" by simp from 6 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all then have 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" fix y let ?e = "Itm vs (y#bs) e" consider "?c = 0" | "?c > 0" | "?c < 0" by arith then show ?case proof cases case 1 then show ?thesis using eqs by auto next case c: 2 have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))" if "x < -?e / ?c" for x proof - from that have "?c * x < - ?e" using pos_less_divide_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then show ?thesis using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto qed then show ?thesis by auto next case c: 3 have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))" if "x < -?e / ?c" for x proof - from that have "?c * x > - ?e" using neg_less_divide_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then show ?thesis using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto qed then show ?thesis by auto qed qed auto lemma plusinf_inf: assumes "islin p" shows "\z. \x > z. Ifm vs (x#bs) (plusinf p) \ Ifm vs (x#bs) p" using assms proof (induct p rule: plusinf.induct) - case 1 + case (1 p q) + then obtain zp zq where zp: "\x>zp. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p" + and zq: "\x>zq. Ifm vs (x # bs) (plusinf q) = Ifm vs (x # bs) q" + by force then show ?case - apply auto - apply (rule_tac x="max z za" in exI) - apply auto - done + by (rule_tac x="max zp zq" in exI) auto next - case 2 + case (2 p q) + then obtain zp zq where zp: "\x>zp. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p" + and zq: "\x>zq. Ifm vs (x # bs) (plusinf q) = Ifm vs (x # bs) q" + by force then show ?case - apply auto - apply (rule_tac x="max z za" in exI) - apply auto - done + by (rule_tac x="max zp zq" in exI) auto next case (3 c e) then have nbe: "tmbound0 e" by simp from 3 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" fix y let ?e = "Itm vs (y#bs) e" consider "?c = 0" | "?c > 0" | "?c < 0" by arith then show ?case proof cases case 1 then show ?thesis using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto next case c: 2 have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))" if "x > -?e / ?c" for x proof - from that have "?c * x > - ?e" using pos_divide_less_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then show ?thesis using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto qed then show ?thesis by auto next case c: 3 have "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))" if "x > -?e / ?c" for x proof - from that have "?c * x < - ?e" using neg_divide_less_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then show ?thesis using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto qed then show ?thesis by auto qed next case (4 c e) then have nbe: "tmbound0 e" by simp from 4 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" fix y let ?e = "Itm vs (y#bs) e" consider "?c = 0" | "?c > 0" | "?c < 0" by arith then show ?case proof cases case 1 then show ?thesis using eqs by auto next case c: 2 have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))" if "x > -?e / ?c" for x proof - from that have "?c * x > - ?e" using pos_divide_less_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then show ?thesis using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto qed then show ?thesis by auto next case c: 3 have "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))" if "x > -?e / ?c" for x proof - from that have "?c * x < - ?e" using neg_divide_less_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then show ?thesis using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto qed then show ?thesis by auto qed next case (5 c e) then have nbe: "tmbound0 e" by simp from 5 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all then have 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" fix y let ?e = "Itm vs (y#bs) e" consider "?c = 0" | "?c > 0" | "?c < 0" by arith then show ?case proof cases case 1 then show ?thesis using eqs by auto next case c: 2 have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))" if "x > -?e / ?c" for x proof - from that have "?c * x > - ?e" using pos_divide_less_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then show ?thesis using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto qed then show ?thesis by auto next case c: 3 have "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))" if "x > -?e / ?c" for x proof - from that have "?c * x < - ?e" using neg_divide_less_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then show ?thesis using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] c by auto qed then show ?thesis by auto qed next case (6 c e) then have nbe: "tmbound0 e" by simp from 6 have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all then have 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" fix y let ?e = "Itm vs (y#bs) e" consider "?c = 0" | "?c > 0" | "?c < 0" by arith then show ?case proof cases case 1 then show ?thesis using eqs by auto next case c: 2 have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))" if "x > -?e / ?c" for x proof - from that have "?c * x > - ?e" using pos_divide_less_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e > 0" by simp then show ?thesis using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto qed then show ?thesis by auto next case c: 3 have "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))" if "x > -?e / ?c" for x proof - from that have "?c * x < - ?e" using neg_divide_less_eq[OF c, where a="x" and b="-?e"] by (simp add: mult.commute) then have "?c * x + ?e < 0" by simp then show ?thesis using tmbound0_I[OF nbe, where b="y" and b'="x"] c eqs by auto qed then show ?thesis by auto qed 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) + by (induct p rule: minusinf.induct) (auto simp: 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) + by (induct p rule: minusinf.induct) (auto simp: 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 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: "\xx. Ifm vs (x#bs) p" proof - from bound0_I [OF plusinf_nb[OF lp], where 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: "\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 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) - apply (auto simp add: eq le lt polyneg_norm) - apply (auto simp add: linorder_not_less order_le_less) - done + by (induct p rule: minusinf.induct) + (auto simp: eq le lt polyneg_norm linorder_not_less order_le_less) 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 then have "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) + by (auto simp: mult.commute) then show ?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) - apply (auto simp add: eq le lt polyneg_norm) - apply (auto simp add: linorder_not_less order_le_less) - done + by (induct p rule: minusinf.induct) + (auto simp: eq le lt polyneg_norm linorder_not_less order_le_less) then obtain c s where c_s: "(c, s) \ set (uset p)" and "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 then have "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) + by (auto simp: mult.commute) then show ?thesis using c_s 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 c_s: "(c,s) \ set (uset p)" and x: "x \ - Itm vs (x#bs) s / Ipoly vs c" by blast from uset_l[OF lp, rule_format, OF c_s] have nb: "tmbound0 s" by simp from x tmbound0_I[OF nb, of vs x bs a] c_s 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 then have ycs: "y < - ?Nt x s / ?N c \ y > -?Nt x s / ?N c" by auto consider "?N c = 0" | "?N c > 0" | "?N c < 0" by arith then show ?case proof cases case 1 then show ?thesis using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"]) next case N: 2 from px pos_less_divide_eq[OF N, where a="x" and b="-?Nt x s"] have px': "x < - ?Nt x s / ?N c" - by (auto simp add: not_less field_simps) + by (auto simp: not_less field_simps) from ycs show ?thesis proof assume y: "y < - ?Nt x s / ?N c" then have "y * ?N c < - ?Nt x s" by (simp add: pos_less_divide_eq[OF N, where a="y" and b="-?Nt x s", symmetric]) then have "?N c * y + ?Nt x s < 0" by (simp add: field_simps) then show ?thesis using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp next 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 then show ?thesis .. qed next case N: 3 from px neg_divide_less_eq[OF N, where a="x" and b="-?Nt x s"] have px': "x > - ?Nt x s / ?N c" - by (auto simp add: not_less field_simps) + by (auto simp: not_less field_simps) from ycs show ?thesis proof assume y: "y > - ?Nt x s / ?N c" then have "y * ?N c < - ?Nt x s" by (simp add: neg_divide_less_eq[OF N, where a="y" and b="-?Nt x s", symmetric]) then have "?N c * y + ?Nt x s < 0" by (simp add: field_simps) then show ?thesis using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp next 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 then show ?thesis .. qed qed 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 then have 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 consider "?N c = 0" | "?N c > 0" | "?N c < 0" by arith then show ?case proof cases case 1 then show ?thesis using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"]) next case N: 2 from px pos_le_divide_eq[OF N, where a="x" and b="-?Nt x s"] have px': "x \ - ?Nt x s / ?N c" by (simp add: not_less field_simps) from ycs show ?thesis proof assume y: "y < - ?Nt x s / ?N c" then have "y * ?N c < - ?Nt x s" by (simp add: pos_less_divide_eq[OF N, where a="y" and b="-?Nt x s", symmetric]) then have "?N c * y + ?Nt x s < 0" by (simp add: field_simps) then show ?thesis using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp next 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 then show ?thesis .. qed next case N: 3 from px neg_divide_le_eq[OF N, where a="x" and b="-?Nt x s"] have px': "x \ - ?Nt x s / ?N c" by (simp add: field_simps) from ycs show ?thesis proof assume y: "y > - ?Nt x s / ?N c" then have "y * ?N c < - ?Nt x s" by (simp add: neg_divide_less_eq[OF N, where a="y" and b="-?Nt x s", symmetric]) then have "?N c * y + ?Nt x s < 0" by (simp add: field_simps) then show ?thesis using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp next 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 then show ?thesis .. qed qed 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 then have ycs: "y < - ?Nt x s / ?N c \ y > -?Nt x s / ?N c" by auto consider "?N c = 0" | "?N c < 0" | "?N c > 0" by arith then show ?case proof cases case 1 then show ?thesis using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"]) next case 2 then have 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: field_simps) from ycs show ?thesis proof 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 then show ?thesis .. next 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 then show ?thesis .. qed next case 3 then have 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: field_simps) from ycs show ?thesis proof 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 then show ?thesis .. next 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 then show ?thesis .. qed qed 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 then have ycs: "y < - ?Nt x s / ?N c \ y > -?Nt x s / ?N c" by auto show ?case proof (cases "?N c = 0") case True then show ?thesis using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"]) next case False with yne eq_divide_eq[of "y" "- ?Nt x s" "?N c"] show ?thesis by (simp add: field_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric]) qed -qed (auto simp add: tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"] +qed (auto simp: 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 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) / 2) 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) / 2) 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 then have 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 then have lx: "?l \ a" using tx by simp have "- ?Nt a s / ?N d \ ?u" using dsM Mne by simp then have 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] consider u where "u \ ?M" "?I u p" | t1 t2 where "t1 \ ?M" "t2\ ?M" "\y. t1 < y \ y < t2 \ y \ ?M" "t1 < a" "a < t2" "?I a p" by blast then show ?thesis proof cases case 1 then have "\(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 have "?I (((- ?Nt a tu / ?N nu) + (- ?Nt a tu / ?N nu)) / 2) p" using \?I u p\ tuu by simp with tuU show ?thesis by blast next case 2 have "\(t1n, t1u) \ ?U. t1 = - ?Nt a t1u / ?N t1n" using \t1 \ ?M\ by auto then obtain t1u t1n where t1uU: "(t1n, t1u) \ ?U" and t1u: "t1 = - ?Nt a t1u / ?N t1n" by blast have "\(t2n, t2u) \ ?U. t2 = - ?Nt a t2u / ?N t2n" using \t2 \ ?M\ by auto then obtain t2u t2n where t2uU: "(t2n, t2u) \ ?U" and t2u: "t2 = - ?Nt a t2u / ?N t2n" by blast have "t1 < t2" using \t1 < a\ \a < t2\ by simp let ?u = "(t1 + t2) / 2" have "t1 < ?u" using less_half_sum [OF \t1 < t2\] by auto have "?u < t2" using gt_half_sum [OF \t1 < t2\] by auto have "?I ?u p" using lp \\y. t1 < y \ y < t2 \ y \ ?M\ \t1 < a\ \a < t2\ \?I a p\ \t1 < ?u\ \?u < t2\ by (rule lin_dense) with t1uU t2uU t1u t2u show ?thesis by blast qed 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) / 2) 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) / 2) p" by simp with lnU smU show ?thesis by auto qed section \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) / 2)#bs) p))" (is "(\x. ?I x p) \ ?M \ ?P \ ?F" is "?E \ ?D") proof show ?D if ?E proof - consider "?M \ ?P" | "\ ?M" "\ ?P" by blast then show ?thesis proof cases case 1 then show ?thesis by blast next case 2 from inf_uset[OF lp this] have ?F using \?E\ by blast then show ?thesis by blast qed qed show ?E if ?D proof - from that consider ?M | ?P | ?F by blast then show ?thesis proof cases case 1 from minusinf_ex[OF lp this] show ?thesis . next case 2 from plusinf_ex[OF lp this] show ?thesis . next case 3 then show ?thesis by blast qed qed qed section \First implementation : Naive by encoding all case splits locally\ definition "msubsteq c t d s a r = evaldjf (case_prod 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 (case_prod 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) / 2)#bs) (Eq (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] consider "?c = 0" "?d = 0" | "?c = 0" "?d \ 0" | "?c \ 0" "?d = 0" | "?c \ 0" "?d \ 0" by blast then show ?thesis proof cases case 1 then show ?thesis by (simp add: r[of 0] msubsteq_def Let_def evaldjf_ex) next case cd: 2 then have th: "(- ?t / ?c + - ?s / ?d)/2 = -?s / (2*?d)" by simp have "?rhs = Ifm vs (-?s / (2*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (-?s / (2*?d)) + ?r = 0" by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \d\\<^sub>p\<^bsup>vs\<^esup>))"]) also have "\ \ 2 * ?d * (?a * (-?s / (2*?d)) + ?r) = 0" using cd(2) mult_cancel_left[of "2*?d" "(?a * (-?s / (2*?d)) + ?r)" 0] by simp also have "\ \ (- ?a * ?s) * (2*?d / (2*?d)) + 2 * ?d * ?r= 0" by (simp add: field_simps distrib_left [of "2*?d"]) also have "\ \ - (?a * ?s) + 2*?d*?r = 0" using cd(2) by simp finally show ?thesis using cd by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \d\\<^sub>p\<^bsup>vs\<^esup>))"] msubsteq_def Let_def evaldjf_ex) next case cd: 3 from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?t / (2 * ?c)" by simp have "?rhs = Ifm vs (-?t / (2*?c) # bs) (Eq (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (-?t / (2*?c)) + ?r = 0" by (simp add: r[of "- (?t/ (2 * ?c))"]) also have "\ \ 2 * ?c * (?a * (-?t / (2 * ?c)) + ?r) = 0" using cd(1) mult_cancel_left[of "2 * ?c" "(?a * (-?t / (2 * ?c)) + ?r)" 0] by simp also have "\ \ (?a * -?t)* (2 * ?c) / (2 * ?c) + 2 * ?c * ?r= 0" by (simp add: field_simps distrib_left [of "2 * ?c"]) also have "\ \ - (?a * ?t) + 2 * ?c * ?r = 0" using cd(1) by simp finally show ?thesis using cd by (simp add: r[of "- (?t/ (2 * ?c))"] msubsteq_def Let_def evaldjf_ex) next case cd: 4 then have cd2: "?c * ?d * 2 \ 0" by simp from add_frac_eq[OF cd, of "- ?t" "- ?s"] have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2 * ?c * ?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r = 0" by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"]) also have "\ \ (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) = 0" using cd mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2 * ?c * ?d)) + ?r" 0] by simp also have "\ \ ?a * (- (?d * ?t + ?c* ?s )) + 2 * ?c * ?d * ?r = 0" using nonzero_mult_div_cancel_left [OF cd2] cd by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using cd by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex field_simps) qed qed definition "msubstneq c t d s a r = evaldjf (case_prod 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 (case_prod 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) /2)#bs) (NEq (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] consider "?c = 0" "?d = 0" | "?c = 0" "?d \ 0" | "?c \ 0" "?d = 0" | "?c \ 0" "?d \ 0" by blast then show ?thesis proof cases case 1 then show ?thesis by (simp add: r[of 0] msubstneq_def Let_def evaldjf_ex) next case cd: 2 from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?s / (2 * ?d)" by simp have "?rhs = Ifm vs (-?s / (2*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (-?s / (2*?d)) + ?r \ 0" by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \d\\<^sub>p\<^bsup>vs\<^esup>))"]) also have "\ \ 2*?d * (?a * (-?s / (2*?d)) + ?r) \ 0" using cd(2) mult_cancel_left[of "2*?d" "(?a * (-?s / (2*?d)) + ?r)" 0] by simp also have "\ \ (- ?a * ?s) * (2*?d / (2*?d)) + 2*?d*?r\ 0" by (simp add: field_simps distrib_left[of "2*?d"] del: distrib_left) also have "\ \ - (?a * ?s) + 2*?d*?r \ 0" using cd(2) by simp finally show ?thesis using cd by (simp add: r[of "- (Itm vs (x # bs) s / (2 * \d\\<^sub>p\<^bsup>vs\<^esup>))"] msubstneq_def Let_def evaldjf_ex) next case cd: 3 from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = -?t / (2*?c)" by simp have "?rhs = Ifm vs (-?t / (2*?c) # bs) (NEq (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (-?t / (2*?c)) + ?r \ 0" by (simp add: r[of "- (?t/ (2 * ?c))"]) also have "\ \ 2*?c * (?a * (-?t / (2*?c)) + ?r) \ 0" using cd(1) mult_cancel_left[of "2*?c" "(?a * (-?t / (2*?c)) + ?r)" 0] by simp also have "\ \ (?a * -?t)* (2*?c) / (2*?c) + 2*?c*?r \ 0" by (simp add: field_simps distrib_left[of "2*?c"] del: distrib_left) also have "\ \ - (?a * ?t) + 2*?c*?r \ 0" using cd(1) by simp finally show ?thesis using cd by (simp add: r[of "- (?t/ (2*?c))"] msubstneq_def Let_def evaldjf_ex) next case cd: 4 then have cd2: "?c * ?d * 2 \ 0" by simp from add_frac_eq[OF cd, of "- ?t" "- ?s"] have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c * ?s )/ (2 * ?c * ?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \ 0" by (simp add: r [of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"]) also have "\ \ (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \ 0" using cd mult_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp also have "\ \ ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \ 0" using nonzero_mult_div_cancel_left[OF cd2] cd by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using cd by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex field_simps) qed qed definition "msubstlt c t d s a r = evaldjf (case_prod 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 (case_prod 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) /2)#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] consider "?c = 0" "?d = 0" | "?c * ?d > 0" | "?c * ?d < 0" | "?c > 0" "?d = 0" | "?c < 0" "?d = 0" | "?c = 0" "?d > 0" | "?c = 0" "?d < 0" by atomize_elim auto then show ?thesis proof cases case 1 then show ?thesis using nc nd by (simp add: polyneg_norm lt r[of 0] msubstlt_def Let_def evaldjf_ex) next case cd: 2 then have cd2: "2 * ?c * ?d > 0" by simp from cd have c: "?c \ 0" and d: "?d \ 0" by auto from cd2 have cd2': "\ 2 * ?c * ?d < 0" by simp from add_frac_eq[OF c d, of "- ?t" "- ?s"] have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c* ?s )/ (2 * ?c * ?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r < 0" by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"]) also have "\ \ (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) < 0" using cd2 cd2' mult_less_cancel_left_disj[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp also have "\ \ ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r < 0" using nonzero_mult_div_cancel_left[of "2*?c*?d"] c d by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using cd c d nc nd cd2' by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case cd: 3 then have cd2: "2 * ?c * ?d < 0" by (simp add: mult_less_0_iff field_simps) from cd have 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)/2 = - (?d * ?t + ?c* ?s) / (2 * ?c * ?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ (2 * ?c * ?d) # bs) (Lt (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ (2 * ?c * ?d)) + ?r < 0" by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"]) also have "\ \ (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c * ?s )/ (2 * ?c * ?d)) + ?r) > 0" using cd2 order_less_not_sym[OF cd2] mult_less_cancel_left_disj[of "2 * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r"] by simp also have "\ \ ?a * ((?d * ?t + ?c* ?s )) - 2 * ?c * ?d * ?r < 0" using nonzero_mult_div_cancel_left[of "2 * ?c * ?d"] c d by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using cd c d nc nd by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case cd: 4 from cd(1) have c'': "2 * ?c > 0" by (simp add: zero_less_mult_iff) from cd(1) have c': "2 * ?c \ 0" by simp from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2 * ?c)" by (simp add: field_simps) have "?rhs \ Ifm vs (- ?t / (2 * ?c) # bs) (Lt (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- ?t / (2 * ?c))+ ?r < 0" by (simp add: r[of "- (?t / (2 * ?c))"]) also have "\ \ 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) < 0" using cd(1) mult_less_cancel_left_disj[of "2 * ?c" "?a* (- ?t / (2*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp also have "\ \ - ?a * ?t + 2 * ?c * ?r < 0" using nonzero_mult_div_cancel_left[OF c'] \?c > 0\ by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right) finally show ?thesis using cd nc nd by (simp add: r[of "- (?t / (2*?c))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case cd: 5 from cd(1) have c': "2 * ?c \ 0" by simp from cd(1) have c'': "2 * ?c < 0" by (simp add: mult_less_0_iff) from cd(2) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2 * ?c)" by (simp add: field_simps) have "?rhs \ Ifm vs (- ?t / (2*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- ?t / (2*?c))+ ?r < 0" by (simp add: r[of "- (?t / (2*?c))"]) also have "\ \ 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) > 0" using cd(1) order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_less_cancel_left_disj[of "2 * ?c" 0 "?a* (- ?t / (2*?c))+ ?r"] by simp also have "\ \ ?a*?t - 2*?c *?r < 0" using nonzero_mult_div_cancel_left[OF c'] cd(1) order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using cd nc nd by (simp add: r[of "- (?t / (2*?c))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case cd: 6 from cd(2) have d'': "2 * ?d > 0" by (simp add: zero_less_mult_iff) from cd(2) have d': "2 * ?d \ 0" by simp from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2 * ?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- ?s / (2 * ?d) # bs) (Lt (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- ?s / (2 * ?d))+ ?r < 0" by (simp add: r[of "- (?s / (2 * ?d))"]) also have "\ \ 2 * ?d * (?a * (- ?s / (2 * ?d))+ ?r) < 0" using cd(2) mult_less_cancel_left_disj[of "2 * ?d" "?a * (- ?s / (2 * ?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp also have "\ \ - ?a * ?s+ 2 * ?d * ?r < 0" using nonzero_mult_div_cancel_left[OF d'] cd(2) by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right) finally show ?thesis using cd nc nd by (simp add: r[of "- (?s / (2*?d))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case cd: 7 from cd(2) have d': "2 * ?d \ 0" by simp from cd(2) have d'': "2 * ?d < 0" by (simp add: mult_less_0_iff) from cd(1) have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2*?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- ?s / (2 * ?d) # bs) (Lt (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- ?s / (2 * ?d)) + ?r < 0" by (simp add: r[of "- (?s / (2 * ?d))"]) also have "\ \ 2 * ?d * (?a * (- ?s / (2 * ?d)) + ?r) > 0" using cd(2) order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_less_cancel_left_disj[of "2 * ?d" 0 "?a* (- ?s / (2*?d))+ ?r"] by simp also have "\ \ ?a * ?s - 2 * ?d * ?r < 0" using nonzero_mult_div_cancel_left[OF d'] cd(2) order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using cd nc nd by (simp add: r[of "- (?s / (2*?d))"] msubstlt_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) qed qed definition "msubstle c t d s a r = evaldjf (case_prod 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 (case_prod 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) /2)#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 "?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 then consider "?c = 0" "?d = 0" | "?c * ?d > 0" | "?c * ?d < 0" | "?c > 0" "?d = 0" | "?c < 0" "?d = 0" | "?c = 0" "?d > 0" | "?c = 0" "?d < 0" by blast then show ?thesis proof cases case 1 with nc nd show ?thesis by (simp add: polyneg_norm polymul_norm lt r[of 0] msubstle_def Let_def evaldjf_ex) next case dc: 2 from dc have dc': "2 * ?c * ?d > 0" by simp then have c: "?c \ 0" and d: "?d \ 0" by auto from dc' have dc'': "\ 2 * ?c * ?d < 0" by simp from add_frac_eq[OF c d, of "- ?t" "- ?s"] have th: "(- ?t / ?c + - ?s / ?d)/2 = - (?d * ?t + ?c * ?s )/ (2 * ?c * ?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \ 0" by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"]) also have "\ \ (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \ 0" using dc' dc'' mult_le_cancel_left[of "2 * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r" 0] by simp also have "\ \ ?a * (- (?d * ?t + ?c* ?s )) + 2*?c*?d*?r \ 0" using nonzero_mult_div_cancel_left[of "2*?c*?d"] c d by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using dc c d nc nd dc' by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case dc: 3 from dc have dc': "2 * ?c * ?d < 0" by (simp add: mult_less_0_iff field_simps add_neg_neg add_pos_pos) then have 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)/2 = - (?d * ?t + ?c* ?s )/ (2 * ?c * ?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- (?d * ?t + ?c* ?s )/ (2*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r \ 0" by (simp add: r[of "(- (?d * ?t) - (?c *?s)) / (2 * ?c * ?d)"]) also have "\ \ (2 * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ (2*?c*?d)) + ?r) \ 0" using dc' order_less_not_sym[OF dc'] mult_le_cancel_left[of "2 * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ (2*?c*?d)) + ?r"] by simp also have "\ \ ?a * ((?d * ?t + ?c* ?s )) - 2 * ?c * ?d * ?r \ 0" using nonzero_mult_div_cancel_left[of "2 * ?c * ?d"] c d by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using dc c d nc nd by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / (2 * ?c * ?d)"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case 4 have c: "?c > 0" and d: "?d = 0" by fact+ from c have c'': "2 * ?c > 0" by (simp add: zero_less_mult_iff) from c have c': "2 * ?c \ 0" by simp from d have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2*?c)" by (simp add: field_simps) have "?rhs \ Ifm vs (- ?t / (2 * ?c) # bs) (Le (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- ?t / (2 * ?c))+ ?r \ 0" by (simp add: r[of "- (?t / (2 * ?c))"]) also have "\ \ 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) \ 0" using c mult_le_cancel_left[of "2 * ?c" "?a* (- ?t / (2*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp also have "\ \ - ?a*?t+ 2*?c *?r \ 0" using nonzero_mult_div_cancel_left[OF c'] c by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right) finally show ?thesis using c d nc nd by (simp add: r[of "- (?t / (2*?c))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case 5 have c: "?c < 0" and d: "?d = 0" by fact+ then have c': "2 * ?c \ 0" by simp from c have c'': "2 * ?c < 0" by (simp add: mult_less_0_iff) from d have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?t / (2*?c)" by (simp add: field_simps) have "?rhs \ Ifm vs (- ?t / (2 * ?c) # bs) (Le (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- ?t / (2*?c))+ ?r \ 0" by (simp add: r[of "- (?t / (2*?c))"]) also have "\ \ 2 * ?c * (?a * (- ?t / (2 * ?c))+ ?r) \ 0" using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_le_cancel_left[of "2 * ?c" 0 "?a* (- ?t / (2*?c))+ ?r"] by simp also have "\ \ ?a * ?t - 2 * ?c * ?r \ 0" using nonzero_mult_div_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using c d nc nd by (simp add: r[of "- (?t / (2*?c))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case 6 have c: "?c = 0" and d: "?d > 0" by fact+ from d have d'': "2 * ?d > 0" by (simp add: zero_less_mult_iff) from d have d': "2 * ?d \ 0" by simp from c have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2 * ?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- ?s / (2 * ?d) # bs) (Le (CNP 0 a r))" by (simp only: th) also have "\ \ ?a * (- ?s / (2 * ?d))+ ?r \ 0" by (simp add: r[of "- (?s / (2*?d))"]) also have "\ \ 2 * ?d * (?a * (- ?s / (2 * ?d)) + ?r) \ 0" using d mult_le_cancel_left[of "2 * ?d" "?a* (- ?s / (2*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp also have "\ \ - ?a * ?s + 2 * ?d * ?r \ 0" using nonzero_mult_div_cancel_left[OF d'] d by (simp add: algebra_simps diff_divide_distrib less_le del: distrib_right) finally show ?thesis using c d nc nd by (simp add: r[of "- (?s / (2*?d))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) next case 7 have c: "?c = 0" and d: "?d < 0" by fact+ then have d': "2 * ?d \ 0" by simp from d have d'': "2 * ?d < 0" by (simp add: mult_less_0_iff) from c have th: "(- ?t / ?c + - ?s / ?d)/2 = - ?s / (2*?d)" by (simp add: field_simps) have "?rhs \ Ifm vs (- ?s / (2*?d) # bs) (Le (CNP 0 a r))" by (simp only: th) also have "\ \ ?a* (- ?s / (2*?d))+ ?r \ 0" by (simp add: r[of "- (?s / (2*?d))"]) also have "\ \ 2*?d * (?a* (- ?s / (2*?d))+ ?r) \ 0" using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_le_cancel_left[of "2 * ?d" 0 "?a* (- ?s / (2*?d))+ ?r"] by simp also have "\ \ ?a * ?s - 2 * ?d * ?r \ 0" using nonzero_mult_div_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' by (simp add: algebra_simps diff_divide_distrib del: distrib_right) finally show ?thesis using c d nc nd by (simp add: r[of "- (?s / (2*?d))"] msubstle_def Let_def evaldjf_ex field_simps lt polyneg_norm polymul_norm) qed 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) /2)#bs) p" using lp by (induct p rule: islin.induct) - (auto simp add: tmbound0_I + (auto simp: 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>)) / 2" and b' = x and bs = bs and vs = vs] msubsteq msubstneq msubstlt [OF nc nd] msubstle [OF nc nd]) lemma msubst_nb: assumes "islin p" and "tmbound0 t" and "tmbound0 s" shows "bound0 (msubst p ((c,t),(d,s)))" using assms - by (induct p rule: islin.induct) (auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb) + by (induct p rule: islin.induct) (auto simp: 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 *: "\(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>) / 2 # bs) p" from *[rule_format, OF ctU] *[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 *[rule_format, OF ctU] *[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>) / 2 # bs) p" .. } ultimately have **: "(\(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>) / 2 # bs) p) \ ?F" by blast from fr_eq[OF lp, of vs bs x, simplified **] show ?thesis . qed lemma simpfm_lin: assumes "SORT_CONSTRAINT('a::field_char_0)" shows "qfree p \ islin (simpfm p)" - by (induct p rule: simpfm.induct) (auto simp add: conj_lin disj_lin) + by (induct p rule: simpfm.induct) (auto simp: 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 (remdups (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 = "remdups(uset ?q)" let ?Up = "alluopairs ?U" let ?mp = "minusinf ?q" let ?pp = "plusinf ?q" fix x 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: field_simps) } then have th0: "\x \ set ?U. \y \ set ?U. ?I (msubst ?q (x, y)) \ ?I (msubst ?q (y, x))" by auto { 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 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 \ msubst ?q) (x, y)))" using alluopairs_bex[OF th0] by simp also have "\ \ ?I ?mp \ ?I ?pp \ ?I (evaldjf (simpfm \ msubst ?q) ?Up)" by (simp add: evaldjf_ex) also have "\ \ ?I (disj ?mp (disj ?pp (evaldjf (simpfm \ msubst ?q) ?Up)))" by simp also have "\ \ ?rhs" using decr0[OF th1, of vs x bs] - apply (simp add: ferrack_def Let_def) - apply (cases "?mp = T \ ?pp = T") - apply auto - done + by (cases "?mp = T \ ?pp = T") (auto simp: ferrack_def Let_def) 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 section \Second implementation: 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 * 2))#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) /2)#bs) p))" (is "(\x. ?I x p) = (?M \ ?P \ ?Z \ ?U \ ?F)" is "?E = ?D") proof assume px: "\x. ?I x p" consider "?M \ ?P" | "\ ?M" "\ ?P" by blast then show ?D proof cases case 1 then show ?thesis by blast next case 2 have nmi: "\ ?M" and npi: "\ ?P" by fact+ 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 + 1)) 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 = 2 * x" by (simp add: field_simps) consider "?c = 0" "?d = 0" | "?c = 0" "?d \ 0" | "?c \ 0" "?d = 0" | "?c \ 0" "?d \ 0" by auto then show ?thesis proof cases case 1 with ct show ?thesis by simp next case 2 with ct show ?thesis by auto next case 3 with ct show ?thesis by auto next case z: 4 from z have ?F using ct - apply - - apply (rule bexI[where x = "(c,t)"]) - apply simp_all - apply (rule bexI[where x = "(d,s)"]) - apply simp_all + apply (intro bexI[where x = "(c,t)"]; simp) + apply (intro bexI[where x = "(d,s)"]; simp) done then show ?thesis by blast qed qed next assume ?D then consider ?M | ?P | ?Z | ?U | ?F by blast then show ?E proof cases case 1 show ?thesis by (rule minusinf_ex[OF lp \?M\]) next case 2 show ?thesis by (rule plusinf_ex[OF lp \?P\]) qed 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))" 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))" 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))" 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))" 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))" 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] + (auto simp: 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] + (auto simp: 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 consider "?c < 0" | "?c > 0" by arith then show ?thesis proof cases case c: 1 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"] show ?thesis - by (auto simp add: msubst2_def) + by (auto simp: msubst2_def) next case c: 2 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"] show ?thesis - by (auto simp add: msubst2_def) + by (auto simp: msubst2_def) qed qed 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) + (auto simp: msubsteq2_nb msubstltpos_nb msubstlepos_nb) lemma msubstneg_nb: assumes "SORT_CONSTRAINT('a::field_char_0)" 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) + (auto simp: msubsteq2_nb msubstltneg_nb msubstleneg_nb) lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::field_char_0)" 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 lt_nb simpfm_bound0) lemma mult_minus2_left: "-2 * x = - (2 * x)" for x :: "'a::comm_ring_1" by simp lemma mult_minus2_right: "x * -2 = - (x * 2)" for x :: "'a::comm_ring_1" by simp lemma islin_qf: "islin p \ qfree p" - by (induct p rule: islin.induct) (auto simp add: bound0_qf) + by (induct p rule: islin.induct) (auto simp: 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 *: "\(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> * 2) # 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) * have "isnpoly n" by blast then have 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) then have 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) + by (auto simp: 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> * 2) # bs) p" using H(2) nn2 by (simp add: mult_minus2_right) } 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> * 2) # bs) p" from H(1) * have "isnpoly n" by blast then have 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: mult_minus2_right) } 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>) / 2 # 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) * have "isnpoly c" "isnpoly d" by blast+ then have 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)] + by (auto simp: 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>) / 2 # bs) p" by (simp add: add_divide_distrib diff_divide_distrib mult_minus2_left 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>) / 2 # bs) p" from H(1,2) * have "isnpoly c" "isnpoly d" by blast+ then have 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)))" by (simp add: diff_divide_distrib add_divide_distrib mult_minus2_left 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 = remdups (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 = "remdups(uset ?q)" let ?Up = "alluopairs ?U" let ?mp = "minusinf ?q" let ?pp = "plusinf ?q" fix x 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 - have "bound0 ((\(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))" if "(c, t) \ set ?U" for c t proof - from U_l that have "tmbound0 t" by blast from msubst2_nb[OF lq this] show ?thesis by simp qed then show ?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 - 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)))" if "((b,a),(d,c)) \ set ?Up" for b a d c proof - from U_l alluopairs_set1[of ?U] that have this: "tmbound0 (Add (Mul d a) (Mul b c))" by auto from msubst2_nb[OF lq this] show ?thesis by simp qed then show ?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 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 then have 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)] + by (auto simp: 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: field_simps) next assume H: ?rhs then have 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)] + by (auto simp: 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: field_simps) qed } then have th0: "\x \ set ?U. \y \ set ?U. ?I (?s (x, y)) \ ?I (?s (y, x))" by auto 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_bex[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) - apply (simp_all add: Let_def decr0[OF nb]) - done + proof (cases "?mp = T \ ?pp = T") + case True + then show ?thesis + unfolding ferrack2_def + by (force simp add: ferrack2_def list_disj_def) + next + case False + then show ?thesis + by (simp add: ferrack2_def Let_def decr0[OF nb]) + qed 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 this: "\bs p. qfree p \ qfree (ferrack2 p) \ Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)" by blast from qelim[OF this, of "prep p" bs] show ?thesis - unfolding frpar2_def by (auto simp add: prep) + unfolding frpar2_def by (auto simp: prep) qed oracle frpar_oracle = \ let val mk_C = @{code C} o apply2 @{code int_of_integer}; val mk_poly_Bound = @{code poly.Bound} o @{code nat_of_integer}; val mk_Bound = @{code Bound} o @{code nat_of_integer}; val dest_num = snd o HOLogic.dest_number; fun try_dest_num t = SOME ((snd o HOLogic.dest_number) t) handle TERM _ => NONE; fun dest_nat (t as \<^Const_>\Suc\) = HOLogic.dest_nat t (* FIXME !? *) | dest_nat t = dest_num t; fun the_index ts t = let val k = find_index (fn t' => t aconv t') ts; in if k < 0 then raise General.Subscript else k end; fun num_of_term ps \<^Const_>\uminus _ for t\ = @{code poly.Neg} (num_of_term ps t) | num_of_term ps \<^Const_>\plus _ for a b\ = @{code poly.Add} (num_of_term ps a, num_of_term ps b) | num_of_term ps \<^Const_>\minus _ for a b\ = @{code poly.Sub} (num_of_term ps a, num_of_term ps b) | num_of_term ps \<^Const_>\times _ for a b\ = @{code poly.Mul} (num_of_term ps a, num_of_term ps b) | num_of_term ps \<^Const_>\power _ for a n\ = @{code poly.Pw} (num_of_term ps a, @{code nat_of_integer} (dest_nat n)) | num_of_term ps \<^Const_>\divide _ for a b\ = mk_C (dest_num a, dest_num b) | num_of_term ps t = (case try_dest_num t of SOME k => mk_C (k, 1) | NONE => mk_poly_Bound (the_index ps t)); fun tm_of_term fs ps \<^Const_>\uminus _ for t\ = @{code Neg} (tm_of_term fs ps t) | tm_of_term fs ps \<^Const_>\plus _ for a b\ = @{code Add} (tm_of_term fs ps a, tm_of_term fs ps b) | tm_of_term fs ps \<^Const_>\minus _ for a b\ = @{code Sub} (tm_of_term fs ps a, tm_of_term fs ps b) | tm_of_term fs ps \<^Const_>\times _ for a b\ = @{code Mul} (num_of_term ps a, tm_of_term fs ps b) | tm_of_term fs ps t = (@{code CP} (num_of_term ps t) handle TERM _ => mk_Bound (the_index fs t) | General.Subscript => mk_Bound (the_index fs t)); fun fm_of_term fs ps \<^Const_>\True\ = @{code T} | fm_of_term fs ps \<^Const_>\False\ = @{code F} | fm_of_term fs ps \<^Const_>\HOL.Not for p\ = @{code Not} (fm_of_term fs ps p) | fm_of_term fs ps \<^Const_>\HOL.conj for p q\ = @{code And} (fm_of_term fs ps p, fm_of_term fs ps q) | fm_of_term fs ps \<^Const_>\HOL.disj for p q\ = @{code Or} (fm_of_term fs ps p, fm_of_term fs ps q) | fm_of_term fs ps \<^Const_>\HOL.implies for p q\ = @{code Imp} (fm_of_term fs ps p, fm_of_term fs ps q) | fm_of_term fs ps \<^Const_>\HOL.eq \<^Type>\bool\ for p q\ = @{code Iff} (fm_of_term fs ps p, fm_of_term fs ps q) | fm_of_term fs ps \<^Const_>\HOL.eq _ for p q\ = @{code Eq} (@{code Sub} (tm_of_term fs ps p, tm_of_term fs ps q)) | fm_of_term fs ps \<^Const_>\less _ for p q\ = @{code Lt} (@{code Sub} (tm_of_term fs ps p, tm_of_term fs ps q)) | fm_of_term fs ps \<^Const_>\less_eq _ for p q\ = @{code Le} (@{code Sub} (tm_of_term fs ps p, tm_of_term fs ps q)) | fm_of_term fs ps \<^Const_>\Ex _ for \p as Abs _\\ = let val (x', p') = Term.dest_abs_global p in @{code E} (fm_of_term (Free x' :: fs) ps p') end | fm_of_term fs ps \<^Const_>\All _ for \p as Abs _\\ = let val (x', p') = Term.dest_abs_global p in @{code A} (fm_of_term (Free x' :: fs) ps p') end | fm_of_term fs ps _ = error "fm_of_term"; fun term_of_num T ps (@{code poly.C} (a, b)) = let val (c, d) = apply2 (@{code integer_of_int}) (a, b) in if d = 1 then HOLogic.mk_number T c else if d = 0 then \<^Const>\zero_class.zero T\ else \<^Const>\divide T for \HOLogic.mk_number T c\ \HOLogic.mk_number T d\\ end | term_of_num T ps (@{code poly.Bound} i) = nth ps (@{code integer_of_nat} i) | term_of_num T ps (@{code poly.Add} (a, b)) = \<^Const>\plus T for \term_of_num T ps a\ \term_of_num T ps b\\ | term_of_num T ps (@{code poly.Mul} (a, b)) = \<^Const>\times T for \term_of_num T ps a\ \term_of_num T ps b\\ | term_of_num T ps (@{code poly.Sub} (a, b)) = \<^Const>\minus T for \term_of_num T ps a\ \term_of_num T ps b\\ | term_of_num T ps (@{code poly.Neg} a) = \<^Const>\uminus T for \term_of_num T ps a\\ | term_of_num T ps (@{code poly.Pw} (a, n)) = \<^Const>\power T for \term_of_num T ps a\ \HOLogic.mk_number \<^Type>\nat\ (@{code integer_of_nat} n)\\ | term_of_num T ps (@{code poly.CN} (c, n, p)) = term_of_num T ps (@{code poly.Add} (c, @{code poly.Mul} (@{code poly.Bound} n, p))); fun term_of_tm T fs ps (@{code CP} p) = term_of_num T ps p | term_of_tm T fs ps (@{code Bound} i) = nth fs (@{code integer_of_nat} i) | term_of_tm T fs ps (@{code Add} (a, b)) = \<^Const>\plus T for \term_of_tm T fs ps a\ \term_of_tm T fs ps b\\ | term_of_tm T fs ps (@{code Mul} (a, b)) = \<^Const>\times T for \term_of_num T ps a\ \term_of_tm T fs ps b\\ | term_of_tm T fs ps (@{code Sub} (a, b)) = \<^Const>\minus T for \term_of_tm T fs ps a\ \term_of_tm T fs ps b\\ | term_of_tm T fs ps (@{code Neg} a) = \<^Const>\uminus T for \term_of_tm T fs ps a\\ | term_of_tm T fs ps (@{code CNP} (n, c, p)) = term_of_tm T fs ps (@{code Add} (@{code Mul} (c, @{code Bound} n), p)); fun term_of_fm T fs ps @{code T} = \<^Const>\True\ | term_of_fm T fs ps @{code F} = \<^Const>\False\ | term_of_fm T fs ps (@{code Not} p) = \<^Const>\HOL.Not for \term_of_fm T fs ps p\\ | term_of_fm T fs ps (@{code And} (p, q)) = \<^Const>\HOL.conj for \term_of_fm T fs ps p\ \term_of_fm T fs ps q\\ | term_of_fm T fs ps (@{code Or} (p, q)) = \<^Const>\HOL.disj for \term_of_fm T fs ps p\ \term_of_fm T fs ps q\\ | term_of_fm T fs ps (@{code Imp} (p, q)) = \<^Const>\HOL.implies for \term_of_fm T fs ps p\ \term_of_fm T fs ps q\\ | term_of_fm T fs ps (@{code Iff} (p, q)) = \<^Const>\HOL.eq \<^Type>\bool\ for \term_of_fm T fs ps p\ \term_of_fm T fs ps q\\ | term_of_fm T fs ps (@{code Lt} p) = \<^Const>\less T for \term_of_tm T fs ps p\ \<^Const>\zero_class.zero T\\ | term_of_fm T fs ps (@{code Le} p) = \<^Const>\less_eq T for \term_of_tm T fs ps p\ \<^Const>\zero_class.zero T\\ | term_of_fm T fs ps (@{code Eq} p) = \<^Const>\HOL.eq T for \term_of_tm T fs ps p\ \<^Const>\zero_class.zero T\\ | term_of_fm T fs ps (@{code NEq} p) = \<^Const>\Not for (* FIXME HOL.Not!? *) \<^Const>\HOL.eq T for \term_of_tm T fs ps p\ \<^Const>\zero_class.zero T\\\ | term_of_fm T fs ps _ = error "term_of_fm: quantifiers"; fun frpar_procedure alternative T ps fm = let val frpar = if alternative then @{code frpar2} else @{code frpar}; val fs = subtract (op aconv) (map Free (Term.add_frees fm [])) ps; val eval = term_of_fm T fs ps o frpar o fm_of_term fs ps; val t = HOLogic.dest_Trueprop fm; in HOLogic.mk_Trueprop (HOLogic.mk_eq (t, eval t)) end; in fn (ctxt, alternative, ty, ps, ct) => Thm.cterm_of ctxt (frpar_procedure alternative ty ps (Thm.term_of ct)) end \ ML \ structure Parametric_Ferrante_Rackoff = struct fun tactic ctxt alternative T ps = Object_Logic.full_atomize_tac ctxt THEN' CSUBGOAL (fn (g, i) => let val th = frpar_oracle (ctxt, alternative, T, ps, g); in resolve_tac ctxt [th RS iffD2] i end); fun method alternative = let fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); val parsN = "pars"; val typN = "type"; val any_keyword = keyword parsN || keyword typN; 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' (tactic ctxt alternative T ps)) end; end; \ method_setup frpar = \ Parametric_Ferrante_Rackoff.method false \ "parametric QE for linear Arithmetic over fields" method_setup frpar2 = \ Parametric_Ferrante_Rackoff.method true \ "parametric QE for linear Arithmetic over fields, Version 2" lemma "\(x::'a::linordered_field). y \ -1 \ (y + 1) * x < 0" - apply (frpar type: 'a pars: y) - apply (simp add: field_simps) - apply (rule spec[where x=y]) - apply (frpar type: 'a pars: "z::'a") - apply simp - done + by (metis mult.commute neg_less_0_iff_less nonzero_divide_eq_eq sum_eq zero_less_two) lemma "\(x::'a::linordered_field). y \ -1 \ (y + 1)*x < 0" - apply (frpar2 type: 'a pars: y) - apply (simp add: field_simps) - apply (rule spec[where x=y]) - apply (frpar2 type: 'a pars: "z::'a") - apply simp - done - -text \Collins/Jones Problem\ -(* -lemma "\(r::'a::{linordered_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::{linordered_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::{linordered_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: field_simps) -have "?rhs" - - apply (frpar type: "'a::{linordered_field, number_ring}" pars: "a::'a::{linordered_field, number_ring}" "b::'a::{linordered_field, number_ring}") - apply (simp add: field_simps) -oops -*) -(* -lemma "ALL (x::'a::{linordered_field, number_ring}) y. (1 - t)*x \ (1+t)*y \ (1 - t)*y \ (1+t)*x --> 0 \ y" -apply (frpar type: "'a::{linordered_field, number_ring}" pars: "t::'a::{linordered_field, number_ring}") -oops -*) - -text \Collins/Jones Problem\ - -(* -lemma "\(r::'a::{linordered_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::{linordered_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::{linordered_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: field_simps) -have "?rhs" - apply (frpar2 type: "'a::{linordered_field, number_ring}" pars: "a::'a::{linordered_field, number_ring}" "b::'a::{linordered_field, number_ring}") - apply simp -oops -*) - -(* -lemma "ALL (x::'a::{linordered_field, number_ring}) y. (1 - t)*x \ (1+t)*y \ (1 - t)*y \ (1+t)*x --> 0 \ y" -apply (frpar2 type: "'a::{linordered_field, number_ring}" pars: "t::'a::{linordered_field, number_ring}") -apply (simp add: field_simps linorder_neq_iff[symmetric]) -apply ferrack -oops -*) + by (metis mult.right_neutral mult_minus1_right neg_0_le_iff_le nle_le not_less sum_eq) + end diff --git a/src/HOL/Decision_Procs/Polynomial_List.thy b/src/HOL/Decision_Procs/Polynomial_List.thy --- a/src/HOL/Decision_Procs/Polynomial_List.thy +++ b/src/HOL/Decision_Procs/Polynomial_List.thy @@ -1,1132 +1,1032 @@ (* Title: HOL/Decision_Procs/Polynomial_List.thy Author: Amine Chaieb *) section \Univariate Polynomials as lists\ theory Polynomial_List -imports Complex_Main + imports Complex_Main + begin text \Application of polynomial as a function.\ primrec (in semiring_0) poly :: "'a list \ 'a \ 'a" where poly_Nil: "poly [] x = 0" | poly_Cons: "poly (h # t) x = h + x * poly t x" subsection \Arithmetic Operations on Polynomials\ text \Addition\ primrec (in semiring_0) padd :: "'a list \ 'a list \ 'a list" (infixl \+++\ 65) where padd_Nil: "[] +++ l2 = l2" | padd_Cons: "(h # t) +++ l2 = (if l2 = [] then h # t else (h + hd l2) # (t +++ tl l2))" text \Multiplication by a constant\ primrec (in semiring_0) cmult :: "'a \ 'a list \ 'a list" (infixl \%*\ 70) where cmult_Nil: "c %* [] = []" | cmult_Cons: "c %* (h#t) = (c * h)#(c %* t)" text \Multiplication by a polynomial\ primrec (in semiring_0) pmult :: "'a list \ 'a list \ 'a list" (infixl \***\ 70) where pmult_Nil: "[] *** l2 = []" | pmult_Cons: "(h # t) *** l2 = (if t = [] then h %* l2 else (h %* l2) +++ (0 # (t *** l2)))" text \Repeated multiplication by a polynomial\ primrec (in semiring_0) mulexp :: "nat \ 'a list \ 'a list \ 'a list" where mulexp_zero: "mulexp 0 p q = q" | mulexp_Suc: "mulexp (Suc n) p q = p *** mulexp n p q" text \Exponential\ primrec (in semiring_1) pexp :: "'a list \ nat \ 'a list" (infixl \%^\ 80) where pexp_0: "p %^ 0 = [1]" | pexp_Suc: "p %^ (Suc n) = p *** (p %^ n)" text \Quotient related value of dividing a polynomial by x + a. Useful for divisor properties in inductive proofs.\ primrec (in field) "pquot" :: "'a list \ 'a \ 'a list" where pquot_Nil: "pquot [] a = []" | pquot_Cons: "pquot (h # t) a = (if t = [] then [h] else (inverse a * (h - hd( pquot t a))) # pquot t a)" text \Normalization of polynomials (remove extra 0 coeff).\ primrec (in semiring_0) pnormalize :: "'a list \ 'a list" where pnormalize_Nil: "pnormalize [] = []" | pnormalize_Cons: "pnormalize (h # p) = (if pnormalize p = [] then (if h = 0 then [] else [h]) else h # pnormalize p)" definition (in semiring_0) "pnormal p \ pnormalize p = p \ p \ []" definition (in semiring_0) "nonconstant p \ pnormal p \ (\x. p \ [x])" text \Other definitions.\ definition (in ring_1) poly_minus :: "'a list \ 'a list" (\-- _\ [80] 80) where "-- p = (- 1) %* p" definition (in semiring_0) divides :: "'a list \ 'a list \ bool" (infixl \divides\ 70) where "p1 divides p2 \ (\q. poly p2 = poly(p1 *** q))" lemma (in semiring_0) dividesI: "poly p2 = poly (p1 *** q) \ p1 divides p2" by (auto simp add: divides_def) lemma (in semiring_0) dividesE: assumes "p1 divides p2" obtains q where "poly p2 = poly (p1 *** q)" using assms by (auto simp add: divides_def) \ \order of a polynomial\ definition (in ring_1) order :: "'a \ 'a list \ nat" where "order a p = (SOME n. ([-a, 1] %^ n) divides p \ \ (([-a, 1] %^ (Suc n)) divides p))" \ \degree of a polynomial\ definition (in semiring_0) degree :: "'a list \ nat" where "degree p = length (pnormalize p) - 1" \ \squarefree polynomials --- NB with respect to real roots only\ definition (in ring_1) rsquarefree :: "'a list \ bool" where "rsquarefree p \ poly p \ poly [] \ (\a. order a p = 0 \ order a p = 1)" context semiring_0 begin lemma padd_Nil2[simp]: "p +++ [] = p" by (induct p) auto lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)" by auto lemma pminus_Nil: "-- [] = []" by (simp add: poly_minus_def) lemma pmult_singleton: "[h1] *** p1 = h1 %* p1" by simp end lemma (in semiring_1) poly_ident_mult[simp]: "1 %* t = t" by (induct t) auto lemma (in semiring_0) poly_simple_add_Cons[simp]: "[a] +++ (0 # t) = a # t" by simp text \Handy general properties.\ lemma (in comm_semiring_0) padd_commut: "b +++ a = a +++ b" proof (induct b arbitrary: a) case Nil then show ?case by auto next case (Cons b bs a) then show ?case by (cases a) (simp_all add: add.commute) qed lemma (in comm_semiring_0) padd_assoc: "(a +++ b) +++ c = a +++ (b +++ c)" proof (induct a arbitrary: b c) case Nil then show ?case by simp next case Cons then show ?case by (cases b) (simp_all add: ac_simps) qed lemma (in semiring_0) poly_cmult_distr: "a %* (p +++ q) = a %* p +++ a %* q" proof (induct p arbitrary: q) case Nil then show ?case by simp next case Cons then show ?case by (cases q) (simp_all add: distrib_left) qed lemma (in ring_1) pmult_by_x[simp]: "[0, 1] *** t = 0 # t" proof (induct t) case Nil then show ?case by simp next case (Cons a t) then show ?case by (cases t) (auto simp add: padd_commut) qed text \Properties of evaluation of polynomials.\ lemma (in semiring_0) poly_add: "poly (p1 +++ p2) x = poly p1 x + poly p2 x" proof (induct p1 arbitrary: p2) case Nil then show ?case by simp next case (Cons a as p2) then show ?case by (cases p2) (simp_all add: ac_simps distrib_left) qed lemma (in comm_semiring_0) poly_cmult: "poly (c %* p) x = c * poly p x" proof (induct p) case Nil then show ?case by simp next case Cons then show ?case by (cases "x = zero") (auto simp add: distrib_left ac_simps) qed lemma (in comm_semiring_0) poly_cmult_map: "poly (map ((*) c) p) x = c * poly p x" by (induct p) (auto simp add: distrib_left ac_simps) lemma (in comm_ring_1) poly_minus: "poly (-- p) x = - (poly p x)" by (simp add: poly_minus_def) (auto simp add: poly_cmult) lemma (in comm_semiring_0) poly_mult: "poly (p1 *** p2) x = poly p1 x * poly p2 x" proof (induct p1 arbitrary: p2) case Nil then show ?case by simp next case (Cons a as) then show ?case by (cases as) (simp_all add: poly_cmult poly_add distrib_right distrib_left ac_simps) qed class idom_char_0 = idom + ring_char_0 subclass (in field_char_0) idom_char_0 .. lemma (in comm_ring_1) poly_exp: "poly (p %^ n) x = (poly p x) ^ n" by (induct n) (auto simp add: poly_cmult poly_mult) text \More Polynomial Evaluation lemmas.\ lemma (in semiring_0) poly_add_rzero[simp]: "poly (a +++ []) x = poly a x" by simp lemma (in comm_semiring_0) poly_mult_assoc: "poly ((a *** b) *** c) x = poly (a *** (b *** c)) x" by (simp add: poly_mult mult.assoc) lemma (in semiring_0) poly_mult_Nil2[simp]: "poly (p *** []) x = 0" by (induct p) auto lemma (in comm_semiring_1) poly_exp_add: "poly (p %^ (n + d)) x = poly (p %^ n *** p %^ d) x" by (induct n) (auto simp add: poly_mult mult.assoc) subsection \Key Property: if \<^term>\f a = 0\ then \<^term>\(x - a)\ divides \<^term>\p(x)\.\ lemma (in comm_ring_1) lemma_poly_linear_rem: "\q r. h#t = [r] +++ [-a, 1] *** q" proof (induct t arbitrary: h) case Nil have "[h] = [h] +++ [- a, 1] *** []" by simp then show ?case by blast next case (Cons x xs) have "\q r. h # x # xs = [r] +++ [-a, 1] *** q" proof - from Cons obtain q r where qr: "x # xs = [r] +++ [- a, 1] *** q" by blast have "h # x # xs = [a * r + h] +++ [-a, 1] *** (r # q)" using qr by (cases q) (simp_all add: algebra_simps) then show ?thesis by blast qed then show ?case by blast qed lemma (in comm_ring_1) poly_linear_rem: "\q r. h#t = [r] +++ [-a, 1] *** q" using lemma_poly_linear_rem [where t = t and a = a] by auto lemma (in comm_ring_1) poly_linear_divides: "poly p a = 0 \ p = [] \ (\q. p = [-a, 1] *** q)" proof (cases p) case Nil then show ?thesis by simp next case (Cons x xs) have "poly p a = 0" if "p = [-a, 1] *** q" for q using that by (simp add: poly_add poly_cmult) moreover have "\q. p = [- a, 1] *** q" if p0: "poly p a = 0" proof - from poly_linear_rem[of x xs a] obtain q r where qr: "x#xs = [r] +++ [- a, 1] *** q" by blast have "r = 0" using p0 by (simp only: Cons qr poly_mult poly_add) simp - with Cons qr show ?thesis - apply - - apply (rule exI[where x = q]) - apply auto - apply (cases q) - apply auto - done + with Cons qr have "p = [- a, 1] *** q" + by (simp add: local.padd_commut) + then show ?thesis .. qed ultimately show ?thesis using Cons by blast qed lemma (in semiring_0) lemma_poly_length_mult[simp]: "length (k %* p +++ (h # (a %* p))) = Suc (length p)" by (induct p arbitrary: h k a) auto lemma (in semiring_0) lemma_poly_length_mult2[simp]: "length (k %* p +++ (h # p)) = Suc (length p)" by (induct p arbitrary: h k) auto lemma (in ring_1) poly_length_mult[simp]: "length([-a,1] *** q) = Suc (length q)" by auto subsection \Polynomial length\ lemma (in semiring_0) poly_cmult_length[simp]: "length (a %* p) = length p" by (induct p) auto lemma (in semiring_0) poly_add_length: "length (p1 +++ p2) = max (length p1) (length p2)" by (induct p1 arbitrary: p2) auto lemma (in semiring_0) poly_root_mult_length[simp]: "length ([a, b] *** p) = Suc (length p)" by (simp add: poly_add_length) lemma (in idom) poly_mult_not_eq_poly_Nil[simp]: "poly (p *** q) x \ poly [] x \ poly p x \ poly [] x \ poly q x \ poly [] x" by (auto simp add: poly_mult) lemma (in idom) poly_mult_eq_zero_disj: "poly (p *** q) x = 0 \ poly p x = 0 \ poly q x = 0" by (auto simp add: poly_mult) text \Normalisation Properties.\ lemma (in semiring_0) poly_normalized_nil: "pnormalize p = [] \ poly p x = 0" by (induct p) auto text \A nontrivial polynomial of degree n has no more than n roots.\ lemma (in idom) poly_roots_index_lemma: assumes "poly p x \ poly [] x" and "length p = n" shows "\i. \x. poly p x = 0 \ (\m\n. x = i m)" using assms proof (induct n arbitrary: p x) case 0 then show ?case by simp next case (Suc n) have False if C: "\i. \x. poly p x = 0 \ (\m\Suc n. x \ i m)" proof - from Suc.prems have p0: "poly p x \ 0" "p \ []" by auto from p0(1)[unfolded poly_linear_divides[of p x]] have "\q. p \ [- x, 1] *** q" by blast from C obtain a where a: "poly p a = 0" by blast from a[unfolded poly_linear_divides[of p a]] p0(2) obtain q where q: "p = [-a, 1] *** q" by blast have lg: "length q = n" using q Suc.prems(2) by simp from q p0 have qx: "poly q x \ poly [] x" by (auto simp add: poly_mult poly_add poly_cmult) from Suc.hyps[OF qx lg] obtain i where i: "\x. poly q x = 0 \ (\m\n. x = i m)" by blast let ?i = "\m. if m = Suc n then a else i m" from C[of ?i] obtain y where y: "poly p y = 0" "\m\ Suc n. y \ ?i m" by blast from y have "y = a \ poly q y = 0" by (simp only: q poly_mult_eq_zero_disj poly_add) (simp add: algebra_simps) - with i[of y] y(1) y(2) show ?thesis - apply auto - apply (erule_tac x = "m" in allE) - apply auto - done + with i[of y] y show ?thesis + using le_Suc_eq by auto qed then show ?case by blast qed lemma (in idom) poly_roots_index_length: "poly p x \ poly [] x \ \i. \x. poly p x = 0 \ (\n. n \ length p \ x = i n)" by (blast intro: poly_roots_index_lemma) lemma (in idom) poly_roots_finite_lemma1: "poly p x \ poly [] x \ \N i. \x. poly p x = 0 \ (\n::nat. n < N \ x = i n)" - apply (drule poly_roots_index_length) - apply safe - apply (rule_tac x = "Suc (length p)" in exI) - apply (rule_tac x = i in exI) - apply (simp add: less_Suc_eq_le) - done + by (metis le_imp_less_Suc poly_roots_index_length) lemma (in idom) idom_finite_lemma: assumes "\x. P x \ (\n. n < length j \ x = j!n)" shows "finite {x. P x}" proof - from assms have "{x. P x} \ set j" by auto then show ?thesis using finite_subset by auto qed lemma (in idom) poly_roots_finite_lemma2: "poly p x \ poly [] x \ \i. \x. poly p x = 0 \ x \ set i" - apply (drule poly_roots_index_length) - apply safe - apply (rule_tac x = "map (\n. i n) [0 ..< Suc (length p)]" in exI) - apply (auto simp add: image_iff) - apply (erule_tac x="x" in allE) - apply clarsimp - apply (case_tac "n = length p") - apply (auto simp add: order_le_less) - done + using poly_roots_index_length atMost_iff atMost_upto imageI set_map + by metis lemma (in ring_char_0) UNIV_ring_char_0_infinte: "\ finite (UNIV :: 'a set)" proof assume F: "finite (UNIV :: 'a set)" have "finite (UNIV :: nat set)" proof (rule finite_imageD) have "of_nat ` UNIV \ UNIV" by simp then show "finite (of_nat ` UNIV :: 'a set)" using F by (rule finite_subset) show "inj (of_nat :: nat \ 'a)" by (simp add: inj_on_def) qed with infinite_UNIV_nat show False .. qed lemma (in idom_char_0) poly_roots_finite: "poly p \ poly [] \ finite {x. poly p x = 0}" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs - using that - apply - - apply (erule contrapos_np) - apply (rule ext) - apply (rule ccontr) - apply (clarify dest!: poly_roots_finite_lemma2) - using finite_subset proof - - fix x i - assume F: "\ finite {x. poly p x = 0}" - and P: "\x. poly p x = 0 \ x \ set i" - from P have "{x. poly p x = 0} \ set i" - by auto - with finite_subset F show False - by auto + have False if F: "\ finite {x. poly p x = 0}" + and P: "\x. poly p x = 0 \ x \ set i" for i + by (smt (verit, del_insts) in_set_conv_nth local.idom_finite_lemma that) + with that show ?thesis + using local.poly_roots_finite_lemma2 by blast qed show ?lhs if ?rhs using UNIV_ring_char_0_infinte that by auto qed text \Entirety and Cancellation for polynomials\ lemma (in idom_char_0) poly_entire_lemma2: assumes p0: "poly p \ poly []" and q0: "poly q \ poly []" shows "poly (p***q) \ poly []" proof - let ?S = "\p. {x. poly p x = 0}" have "?S (p *** q) = ?S p \ ?S q" by (auto simp add: poly_mult) with p0 q0 show ?thesis unfolding poly_roots_finite by auto qed lemma (in idom_char_0) poly_entire: "poly (p *** q) = poly [] \ poly p = poly [] \ poly q = poly []" using poly_entire_lemma2[of p q] by (auto simp add: fun_eq_iff poly_mult) lemma (in idom_char_0) poly_entire_neg: "poly (p *** q) \ poly [] \ poly p \ poly [] \ poly q \ poly []" by (simp add: poly_entire) lemma (in comm_ring_1) poly_add_minus_zero_iff: "poly (p +++ -- q) = poly [] \ poly p = poly q" by (auto simp add: algebra_simps poly_add poly_minus_def fun_eq_iff poly_cmult) lemma (in comm_ring_1) poly_add_minus_mult_eq: "poly (p *** q +++ --(p *** r)) = poly (p *** (q +++ -- r))" by (auto simp add: poly_add poly_minus_def fun_eq_iff poly_mult poly_cmult algebra_simps) subclass (in idom_char_0) comm_ring_1 .. lemma (in idom_char_0) poly_mult_left_cancel: "poly (p *** q) = poly (p *** r) \ poly p = poly [] \ poly q = poly r" proof - have "poly (p *** q) = poly (p *** r) \ poly (p *** q +++ -- (p *** r)) = poly []" by (simp only: poly_add_minus_zero_iff) also have "\ \ poly p = poly [] \ poly q = poly r" by (auto intro: simp add: poly_add_minus_mult_eq poly_entire poly_add_minus_zero_iff) finally show ?thesis . qed lemma (in idom) poly_exp_eq_zero[simp]: "poly (p %^ n) = poly [] \ poly p = poly [] \ n \ 0" - apply (simp only: fun_eq_iff add: HOL.all_simps [symmetric]) - apply (rule arg_cong [where f = All]) - apply (rule ext) - apply (induct n) - apply (auto simp add: poly_exp poly_mult) - done + by (simp add: local.poly_exp fun_eq_iff) lemma (in comm_ring_1) poly_prime_eq_zero[simp]: "poly [a, 1] \ poly []" - apply (simp add: fun_eq_iff) - apply (rule_tac x = "minus one a" in exI) - apply (simp add: add.commute [of a]) - done +proof - + have "\x. a + x \ 0" + by (metis add_cancel_left_right zero_neq_one) + then show ?thesis + by (simp add: fun_eq_iff) +qed lemma (in idom) poly_exp_prime_eq_zero: "poly ([a, 1] %^ n) \ poly []" by auto text \A more constructive notion of polynomials being trivial.\ -lemma (in idom_char_0) poly_zero_lemma': "poly (h # t) = poly [] \ h = 0 \ poly t = poly []" - apply (simp add: fun_eq_iff) - apply (case_tac "h = zero") - apply (drule_tac [2] x = zero in spec) - apply auto - apply (cases "poly t = poly []") - apply simp +lemma (in idom_char_0) poly_zero_lemma': + assumes "poly (h # t) = poly []" shows "h = 0 \ poly t = poly []" proof - - fix x - assume H: "\x. x = 0 \ poly t x = 0" - assume pnz: "poly t \ poly []" - let ?S = "{x. poly t x = 0}" - from H have "\x. x \ 0 \ poly t x = 0" - by blast - then have th: "?S \ UNIV - {0}" - by auto - from poly_roots_finite pnz have th': "finite ?S" - by blast - from finite_subset[OF th th'] UNIV_ring_char_0_infinte show "poly t x = 0" - by simp + have "poly t x = 0" if H: "\x. x = 0 \ poly t x = 0" and pnz: "poly t \ poly []" for x + proof - + from H have "{x. poly t x = 0} \ UNIV - {0}" + by auto + then show ?thesis + using finite_subset local.poly_roots_finite pnz by fastforce + qed + with assms show ?thesis + by (simp add: fun_eq_iff) (metis add_cancel_right_left mult_eq_0_iff) qed lemma (in idom_char_0) poly_zero: "poly p = poly [] \ (\c \ set p. c = 0)" proof (induct p) case Nil then show ?case by simp next case Cons - show ?case - apply (rule iffI) - apply (drule poly_zero_lemma') - using Cons - apply auto - done + then show ?case + by (smt (verit) list.set_intros pmult_by_x poly_entire poly_zero_lemma' set_ConsD) qed lemma (in idom_char_0) poly_0: "\c \ set p. c = 0 \ poly p x = 0" unfolding poly_zero[symmetric] by simp text \Basics of divisibility.\ lemma (in idom) poly_primes: "[a, 1] divides (p *** q) \ [a, 1] divides p \ [a, 1] divides q" - apply (auto simp add: divides_def fun_eq_iff poly_mult poly_add poly_cmult distrib_right [symmetric]) - apply (drule_tac x = "uminus a" in spec) - apply (simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric]) - apply (cases "p = []") - apply (rule exI[where x="[]"]) - apply simp - apply (cases "q = []") - apply (erule allE[where x="[]"]) - apply simp - - apply clarsimp - apply (cases "\q. p = a %* q +++ (0 # q)") - apply (clarsimp simp add: poly_add poly_cmult) - apply (rule_tac x = qa in exI) - apply (simp add: distrib_right [symmetric]) - apply clarsimp - - apply (auto simp add: poly_linear_divides poly_add poly_cmult distrib_right [symmetric]) - apply (rule_tac x = "pmult qa q" in exI) - apply (rule_tac [2] x = "pmult p qa" in exI) - apply (auto simp add: poly_add poly_mult poly_cmult ac_simps) - done +proof - + have "\q. \x. poly p x = (a + x) * poly q x" + if "poly p (uminus a) * poly q (uminus a) = (a + (uminus a)) * poly qa (uminus a)" + and "\qa. \x. poly q x \ (a + x) * poly qa x" + for qa + using that + apply (simp add: poly_linear_divides poly_add) + by (metis add_cancel_left_right combine_common_factor mult_eq_0_iff poly.poly_Cons poly.poly_Nil poly_add poly_cmult) + moreover have "\qb. \x. (a + x) * poly qa x * poly q x = (a + x) * poly qb x" for qa + by (metis local.poly_mult mult_assoc) + moreover have "\q. \x. poly p x * ((a + x) * poly qa x) = (a + x) * poly q x" for qa + by (metis mult.left_commute local.poly_mult) + ultimately show ?thesis + by (auto simp: divides_def divisors_zero fun_eq_iff poly_mult poly_add poly_cmult simp flip: distrib_right) +qed lemma (in comm_semiring_1) poly_divides_refl[simp]: "p divides p" - apply (simp add: divides_def) - apply (rule_tac x = "[one]" in exI) - apply (auto simp add: poly_mult fun_eq_iff) - done +proof - + have "poly p = poly (p *** [1])" + by (auto simp add: poly_mult fun_eq_iff) + then show ?thesis + using local.dividesI by blast +qed lemma (in comm_semiring_1) poly_divides_trans: "p divides q \ q divides r \ p divides r" - apply (simp add: divides_def) - apply safe - apply (rule_tac x = "pmult qa qaa" in exI) - apply (auto simp add: poly_mult fun_eq_iff mult.assoc) - done + unfolding divides_def + by (metis ext local.poly_mult local.poly_mult_assoc) lemma (in comm_semiring_1) poly_divides_exp: "m \ n \ (p %^ m) divides (p %^ n)" by (auto simp: le_iff_add divides_def poly_exp_add fun_eq_iff) lemma (in comm_semiring_1) poly_exp_divides: "(p %^ n) divides q \ m \ n \ (p %^ m) divides q" by (blast intro: poly_divides_exp poly_divides_trans) -lemma (in comm_semiring_0) poly_divides_add: "p divides q \ p divides r \ p divides (q +++ r)" - apply (auto simp add: divides_def) - apply (rule_tac x = "padd qa qaa" in exI) - apply (auto simp add: poly_add fun_eq_iff poly_mult distrib_left) - done +lemma (in comm_semiring_0) poly_divides_add: + assumes "p divides q" and "p divides r" shows "p divides (q +++ r)" +proof - + have "\qa qb. \poly q = poly (p *** qa); poly r = poly (p *** qb)\ + \ poly (q +++ r) = poly (p *** (qa +++ qb))" + by (auto simp add: poly_add fun_eq_iff poly_mult distrib_left) + with assms show ?thesis + by (auto simp add: divides_def) +qed -lemma (in comm_ring_1) poly_divides_diff: "p divides q \ p divides (q +++ r) \ p divides r" - apply (auto simp add: divides_def) - apply (rule_tac x = "padd qaa (poly_minus qa)" in exI) - apply (auto simp add: poly_add fun_eq_iff poly_mult poly_minus algebra_simps) - done +lemma (in comm_ring_1) poly_divides_diff: + assumes "p divides q" and "p divides (q +++ r)" + shows "p divides r" +proof - + have "\qa qb. \poly q = poly (p *** qa); poly (q +++ r) = poly (p *** qb)\ + \ poly r = poly (p *** (qb +++ -- qa))" + by (auto simp add: poly_add fun_eq_iff poly_mult poly_minus algebra_simps) + with assms show ?thesis + by (auto simp add: divides_def) +qed lemma (in comm_ring_1) poly_divides_diff2: "p divides r \ p divides (q +++ r) \ p divides q" - apply (erule poly_divides_diff) - apply (auto simp add: poly_add fun_eq_iff poly_mult divides_def ac_simps) - done + by (metis local.padd_commut local.poly_divides_diff) lemma (in semiring_0) poly_divides_zero: "poly p = poly [] \ q divides p" - apply (simp add: divides_def) - apply (rule exI[where x = "[]"]) - apply (auto simp add: fun_eq_iff poly_mult) - done + by (metis ext dividesI poly.poly_Nil poly_mult_Nil2) lemma (in semiring_0) poly_divides_zero2 [simp]: "q divides []" - apply (simp add: divides_def) - apply (rule_tac x = "[]" in exI) - apply (auto simp add: fun_eq_iff) - done + using local.poly_divides_zero by force text \At last, we can consider the order of a root.\ lemma (in idom_char_0) poly_order_exists_lemma: assumes "length p = d" and "poly p \ poly []" shows "\n q. p = mulexp n [-a, 1] q \ poly q a \ 0" using assms proof (induct d arbitrary: p) case 0 then show ?case by simp next case (Suc n p) show ?case proof (cases "poly p a = 0") case True from Suc.prems have h: "length p = Suc n" "poly p \ poly []" by auto then have pN: "p \ []" by auto from True[unfolded poly_linear_divides] pN obtain q where q: "p = [-a, 1] *** q" by blast from q h True have qh: "length q = n" "poly q \ poly []" - apply simp_all - apply (simp only: fun_eq_iff) - apply (rule ccontr) - apply (simp add: fun_eq_iff poly_add poly_cmult) - done + using h(2) local.poly_entire q by fastforce+ from Suc.hyps[OF qh] obtain m r where mr: "q = mulexp m [-a,1] r" "poly r a \ 0" by blast from mr q have "p = mulexp (Suc m) [-a,1] r \ poly r a \ 0" by simp then show ?thesis by blast next case False - then show ?thesis - using Suc.prems - apply simp - apply (rule exI[where x="0::nat"]) - apply simp - done + with Suc.prems show ?thesis + by (smt (verit, best) local.mulexp.mulexp_zero) qed qed lemma (in comm_semiring_1) poly_mulexp: "poly (mulexp n p q) x = (poly p x) ^ n * poly q x" by (induct n) (auto simp add: poly_mult ac_simps) lemma (in comm_semiring_1) divides_left_mult: assumes "(p *** q) divides r" shows "p divides r \ q divides r" proof- from assms obtain t where "poly r = poly (p *** q *** t)" unfolding divides_def by blast then have "poly r = poly (p *** (q *** t))" and "poly r = poly (q *** (p *** t))" by (auto simp add: fun_eq_iff poly_mult ac_simps) then show ?thesis unfolding divides_def by blast qed (* FIXME: Tidy up *) lemma (in semiring_1) zero_power_iff: "0 ^ n = (if n = 0 then 1 else 0)" by (induct n) simp_all lemma (in idom_char_0) poly_order_exists: assumes "length p = d" and "poly p \ poly []" shows "\n. [- a, 1] %^ n divides p \ \ [- a, 1] %^ Suc n divides p" proof - from assms have "\n q. p = mulexp n [- a, 1] q \ poly q a \ 0" by (rule poly_order_exists_lemma) then obtain n q where p: "p = mulexp n [- a, 1] q" and "poly q a \ 0" by blast have "[- a, 1] %^ n divides mulexp n [- a, 1] q" proof (rule dividesI) show "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ n *** q)" by (induct n) (simp_all add: poly_add poly_cmult poly_mult algebra_simps) qed moreover have "\ [- a, 1] %^ Suc n divides mulexp n [- a, 1] q" proof assume "[- a, 1] %^ Suc n divides mulexp n [- a, 1] q" then obtain m where "poly (mulexp n [- a, 1] q) = poly ([- a, 1] %^ Suc n *** m)" by (rule dividesE) moreover have "poly (mulexp n [- a, 1] q) \ poly ([- a, 1] %^ Suc n *** m)" proof (induct n) case 0 show ?case proof (rule ccontr) assume "\ ?thesis" then have "poly q a = 0" by (simp add: poly_add poly_cmult) with \poly q a \ 0\ show False by simp qed next case (Suc n) show ?case by (rule pexp_Suc [THEN ssubst]) (simp add: poly_mult_left_cancel poly_mult_assoc Suc del: pmult_Cons pexp_Suc) qed ultimately show False by simp qed ultimately show ?thesis by (auto simp add: p) qed lemma (in semiring_1) poly_one_divides[simp]: "[1] divides p" by (auto simp add: divides_def) lemma (in idom_char_0) poly_order: "poly p \ poly [] \ \!n. ([-a, 1] %^ n) divides p \ \ (([-a, 1] %^ Suc n) divides p)" - apply (auto intro: poly_order_exists simp add: less_linear simp del: pmult_Cons pexp_Suc) - apply (cut_tac x = y and y = n in less_linear) - apply (drule_tac m = n in poly_exp_divides) - apply (auto dest: Suc_le_eq [THEN iffD2, THEN [2] poly_exp_divides] - simp del: pmult_Cons pexp_Suc) - done + by (meson Suc_le_eq linorder_neqE_nat local.poly_exp_divides poly_order_exists) text \Order\ lemma some1_equalityD: "n = (SOME n. P n) \ \!n. P n \ P n" by (blast intro: someI2) lemma (in idom_char_0) order: "([-a, 1] %^ n) divides p \ \ (([-a, 1] %^ Suc n) divides p) \ n = order a p \ poly p \ poly []" unfolding order_def - apply (rule iffI) - apply (blast dest: poly_divides_zero intro!: some1_equality [symmetric] poly_order) - apply (blast intro!: poly_order [THEN [2] some1_equalityD]) - done + by (metis (no_types, lifting) local.poly_divides_zero local.poly_order someI) lemma (in idom_char_0) order2: "poly p \ poly [] \ ([-a, 1] %^ (order a p)) divides p \ \ ([-a, 1] %^ Suc (order a p)) divides p" by (simp add: order del: pexp_Suc) lemma (in idom_char_0) order_unique: "poly p \ poly [] \ ([-a, 1] %^ n) divides p \ \ ([-a, 1] %^ (Suc n)) divides p \ n = order a p" using order [of a n p] by auto lemma (in idom_char_0) order_unique_lemma: "poly p \ poly [] \ ([-a, 1] %^ n) divides p \ \ ([-a, 1] %^ (Suc n)) divides p \ n = order a p" by (blast intro: order_unique) lemma (in ring_1) order_poly: "poly p = poly q \ order a p = order a q" by (auto simp add: fun_eq_iff divides_def poly_mult order_def) lemma (in semiring_1) pexp_one[simp]: "p %^ (Suc 0) = p" by (induct p) auto lemma (in comm_ring_1) lemma_order_root: "0 < n \ [- a, 1] %^ n divides p \ \ [- a, 1] %^ (Suc n) divides p \ poly p a = 0" by (induct n arbitrary: a p) (auto simp add: divides_def poly_mult simp del: pmult_Cons) lemma (in idom_char_0) order_root: "poly p a = 0 \ poly p = poly [] \ order a p \ 0" - apply (cases "poly p = poly []") - apply auto - apply (simp add: poly_linear_divides del: pmult_Cons) - apply safe - apply (drule_tac [!] a = a in order2) - apply (rule ccontr) - apply (simp add: divides_def poly_mult fun_eq_iff del: pmult_Cons) - apply blast - using neq0_conv apply (blast intro: lemma_order_root) - done +proof (cases "poly p = poly []") + case False + then show ?thesis + by (metis (mono_tags, lifting) dividesI lemma_order_root order2 pexp_one poly_linear_divides neq0_conv) +qed auto lemma (in idom_char_0) order_divides: "([-a, 1] %^ n) divides p \ poly p = poly [] \ n \ order a p" - apply (cases "poly p = poly []") - apply auto - apply (simp add: divides_def fun_eq_iff poly_mult) - apply (rule_tac x = "[]" in exI) - apply (auto dest!: order2 [where a=a] intro: poly_exp_divides simp del: pexp_Suc) - done +proof (cases "poly p = poly []") + case True + then show ?thesis + using local.poly_divides_zero by force +next + case False + then show ?thesis + by (meson local.order2 local.poly_exp_divides not_less_eq_eq) +qed lemma (in idom_char_0) order_decomp: - "poly p \ poly [] \ \q. poly p = poly (([-a, 1] %^ order a p) *** q) \ \ [-a, 1] divides q" - unfolding divides_def - apply (drule order2 [where a = a]) - apply (simp add: divides_def del: pexp_Suc pmult_Cons) - apply safe - apply (rule_tac x = q in exI) - apply safe - apply (drule_tac x = qa in spec) - apply (auto simp add: poly_mult fun_eq_iff poly_exp ac_simps simp del: pmult_Cons) - done + assumes "poly p \ poly []" + shows "\q. poly p = poly (([-a, 1] %^ order a p) *** q) \ \ [-a, 1] divides q" +proof - + obtain q where q: "poly p = poly ([- a, 1] %^ order a p *** q)" + using assms local.order2 divides_def by blast + have False if "poly q = poly ([- a, 1] *** qa)" for qa + proof - + have "poly p \ poly ([- a, 1] %^ Suc (order a p) *** qa)" + using assms local.divides_def local.order2 by blast + with q that show False + by (auto simp add: poly_mult ac_simps simp del: pmult_Cons) + qed + with q show ?thesis + unfolding divides_def by blast +qed text \Important composition properties of orders.\ lemma order_mult: fixes a :: "'a::idom_char_0" - shows "poly (p *** q) \ poly [] \ order a (p *** q) = order a p + order a q" - apply (cut_tac a = a and p = "p *** q" and n = "order a p + order a q" in order) - apply (auto simp add: poly_entire simp del: pmult_Cons) - apply (drule_tac a = a in order2)+ - apply safe - apply (simp add: divides_def fun_eq_iff poly_exp_add poly_mult del: pmult_Cons, safe) - apply (rule_tac x = "qa *** qaa" in exI) - apply (simp add: poly_mult ac_simps del: pmult_Cons) - apply (drule_tac a = a in order_decomp)+ - apply safe - apply (subgoal_tac "[-a, 1] divides (qa *** qaa) ") - apply (simp add: poly_primes del: pmult_Cons) - apply (auto simp add: divides_def simp del: pmult_Cons) - apply (rule_tac x = qb in exI) - apply (subgoal_tac "poly ([-a, 1] %^ (order a p) *** (qa *** qaa)) = - poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))") - apply (drule poly_mult_left_cancel [THEN iffD1]) - apply force - apply (subgoal_tac "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (qa *** qaa))) = - poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** qb))) ") - apply (drule poly_mult_left_cancel [THEN iffD1]) - apply force - apply (simp add: fun_eq_iff poly_exp_add poly_mult ac_simps del: pmult_Cons) - done - -lemma (in idom_char_0) order_mult: assumes "poly (p *** q) \ poly []" shows "order a (p *** q) = order a p + order a q" - using assms - apply (cut_tac a = a and p = "pmult p q" and n = "order a p + order a q" in order) - apply (auto simp add: poly_entire simp del: pmult_Cons) - apply (drule_tac a = a in order2)+ - apply safe - apply (simp add: divides_def fun_eq_iff poly_exp_add poly_mult del: pmult_Cons) - apply safe - apply (rule_tac x = "pmult qa qaa" in exI) - apply (simp add: poly_mult ac_simps del: pmult_Cons) - apply (drule_tac a = a in order_decomp)+ - apply safe - apply (subgoal_tac "[uminus a, one] divides pmult qa qaa") - apply (simp add: poly_primes del: pmult_Cons) - apply (auto simp add: divides_def simp del: pmult_Cons) - apply (rule_tac x = qb in exI) - apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa)) = - poly (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb))") - apply (drule poly_mult_left_cancel [THEN iffD1], force) - apply (subgoal_tac "poly (pmult (pexp [uminus a, one] (order a q)) - (pmult (pexp [uminus a, one] (order a p)) (pmult qa qaa))) = - poly (pmult (pexp [uminus a, one] (order a q)) - (pmult (pexp [uminus a, one] (order a p)) (pmult [uminus a, one] qb)))") - apply (drule poly_mult_left_cancel [THEN iffD1], force) - apply (simp add: fun_eq_iff poly_exp_add poly_mult ac_simps del: pmult_Cons) - done +proof - + have p: "poly p \ poly []" and q: "poly q \ poly []" + using assms poly_entire by auto + obtain p' where p': + "\x. poly p x = poly ([- a, 1] %^ order a p) x * poly p' x" + "\ [- a, 1] divides p'" + by (metis order_decomp p poly_mult) + obtain q' where q': + "\x. poly q x = poly ([- a, 1] %^ order a q) x * poly q' x" + "\ [- a, 1] divides q'" + by (metis order_decomp q poly_mult) + have "[- a, 1] %^ (order a p + order a q) divides (p *** q)" + proof - + have *: "poly p x * poly q x = + poly ([- a, 1] %^ order a p) x * poly ([- a, 1] %^ order a q) x * poly (p' *** q') x" for x + using p' q' by (simp add: poly_mult) + then show ?thesis + unfolding divides_def poly_exp_add poly_mult using * by blast + qed + moreover have False + if pq: "order a (p *** q) \ order a p + order a q" + and dv: "[- a, 1] *** [- a, 1] %^ (order a p + order a q) divides (p *** q)" + proof - + obtain pq' :: "'a list" + where pq': "poly (p *** q) = poly ([- a, 1] *** [- a, 1] %^ (order a p + order a q) *** pq')" + using dv unfolding divides_def by auto + have "poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** (p' *** q'))) = + poly ([-a, 1] %^ (order a q) *** ([-a, 1] %^ (order a p) *** ([-a, 1] *** pq')))" + using p' q' pq pq' + by (simp add: fun_eq_iff poly_exp_add poly_mult ac_simps del: pmult_Cons) + then have "poly ([-a, 1] %^ (order a p) *** (p' *** q')) = poly ([-a, 1] %^ (order a p) *** ([-a, 1] *** pq'))" + by (simp add: poly_mult_left_cancel) + then have "[-a, 1] divides (p' *** q')" + unfolding divides_def by (meson poly_exp_prime_eq_zero poly_mult_left_cancel) + with p' q' show ?thesis + by (simp add: poly_primes) + qed + ultimately show ?thesis + by (metis order pexp_Suc) +qed lemma (in idom_char_0) order_root2: "poly p \ poly [] \ poly p a = 0 \ order a p \ 0" - by (rule order_root [THEN ssubst]) auto + using order_root by presburger lemma (in semiring_1) pmult_one[simp]: "[1] *** p = p" by auto lemma (in semiring_0) poly_Nil_zero: "poly [] = poly [0]" by (simp add: fun_eq_iff) lemma (in idom_char_0) rsquarefree_decomp: - "rsquarefree p \ poly p a = 0 \ \q. poly p = poly ([-a, 1] *** q) \ poly q a \ 0" - apply (simp add: rsquarefree_def) - apply safe - apply (frule_tac a = a in order_decomp) - apply (drule_tac x = a in spec) - apply (drule_tac a = a in order_root2 [symmetric]) - apply (auto simp del: pmult_Cons) - apply (rule_tac x = q in exI, safe) - apply (simp add: poly_mult fun_eq_iff) - apply (drule_tac p1 = q in poly_linear_divides [THEN iffD1]) - apply (simp add: divides_def del: pmult_Cons, safe) - apply (drule_tac x = "[]" in spec) - apply (auto simp add: fun_eq_iff) - done - + assumes "rsquarefree p" and "poly p a = 0" + shows "\q. poly p = poly ([-a, 1] *** q) \ poly q a \ 0" +proof - + have "order a p = Suc 0" + using assms local.order_root2 rsquarefree_def by force + moreover + obtain q where "poly p = poly ([- a, 1] %^ order a p *** q)" + "\ [- a, 1] divides q" + using assms(1) order_decomp rsquarefree_def by blast + ultimately show ?thesis + using dividesI poly_linear_divides by auto +qed text \Normalization of a polynomial.\ lemma (in semiring_0) poly_normalize[simp]: "poly (pnormalize p) = poly p" by (induct p) (auto simp add: fun_eq_iff) text \The degree of a polynomial.\ lemma (in semiring_0) lemma_degree_zero: "(\c \ set p. c = 0) \ pnormalize p = []" by (induct p) auto lemma (in idom_char_0) degree_zero: assumes "poly p = poly []" shows "degree p = 0" using assms by (cases "pnormalize p = []") (auto simp add: degree_def poly_zero lemma_degree_zero) lemma (in semiring_0) pnormalize_sing: "pnormalize [x] = [x] \ x \ 0" by simp lemma (in semiring_0) pnormalize_pair: "y \ 0 \ pnormalize [x, y] = [x, y]" by simp lemma (in semiring_0) pnormal_cons: "pnormal p \ pnormal (c # p)" unfolding pnormal_def by simp lemma (in semiring_0) pnormal_tail: "p \ [] \ pnormal (c # p) \ pnormal p" unfolding pnormal_def by (auto split: if_split_asm) lemma (in semiring_0) pnormal_last_nonzero: "pnormal p \ last p \ 0" by (induct p) (simp_all add: pnormal_def split: if_split_asm) lemma (in semiring_0) pnormal_length: "pnormal p \ 0 < length p" unfolding pnormal_def length_greater_0_conv by blast lemma (in semiring_0) pnormal_last_length: "0 < length p \ last p \ 0 \ pnormal p" by (induct p) (auto simp: pnormal_def split: if_split_asm) lemma (in semiring_0) pnormal_id: "pnormal p \ 0 < length p \ last p \ 0" using pnormal_last_length pnormal_length pnormal_last_nonzero by blast lemma (in idom_char_0) poly_Cons_eq: "poly (c # cs) = poly (d # ds) \ c = d \ poly cs = poly ds" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof - from that have "poly ((c # cs) +++ -- (d # ds)) x = 0" for x by (simp only: poly_minus poly_add algebra_simps) (simp add: algebra_simps) then have "poly ((c # cs) +++ -- (d # ds)) = poly []" by (simp add: fun_eq_iff) then have "c = d" and "\x \ set (cs +++ -- ds). x = 0" unfolding poly_zero by (simp_all add: poly_minus_def algebra_simps) from this(2) have "poly (cs +++ -- ds) x = 0" for x unfolding poly_zero[symmetric] by simp with \c = d\ show ?thesis by (simp add: poly_minus poly_add algebra_simps fun_eq_iff) qed show ?lhs if ?rhs using that by (simp add:fun_eq_iff) qed lemma (in idom_char_0) pnormalize_unique: "poly p = poly q \ pnormalize p = pnormalize q" proof (induct q arbitrary: p) case Nil then show ?case by (simp only: poly_zero lemma_degree_zero) simp next case (Cons c cs p) then show ?case proof (induct p) case Nil - then have "poly [] = poly (c # cs)" - by blast - then have "poly (c#cs) = poly []" - by simp then show ?case - by (simp only: poly_zero lemma_degree_zero) simp + by (metis local.poly_zero_lemma') next case (Cons d ds) - then have eq: "poly (d # ds) = poly (c # cs)" - by blast - then have eq': "\x. poly (d # ds) x = poly (c # cs) x" - by simp - then have "poly (d # ds) 0 = poly (c # cs) 0" - by blast - then have dc: "d = c" - by auto - with eq have "poly ds = poly cs" - unfolding poly_Cons_eq by simp - with Cons.prems have "pnormalize ds = pnormalize cs" - by blast - with dc show ?case - by simp + then show ?case + by (metis pnormalize.pnormalize_Cons local.poly_Cons_eq) qed qed lemma (in idom_char_0) degree_unique: assumes pq: "poly p = poly q" shows "degree p = degree q" using pnormalize_unique[OF pq] unfolding degree_def by simp lemma (in semiring_0) pnormalize_length: "length (pnormalize p) \ length p" by (induct p) auto lemma (in semiring_0) last_linear_mul_lemma: "last ((a %* p) +++ (x # (b %* p))) = (if p = [] then x else b * last p)" - apply (induct p arbitrary: a x b) - apply auto - subgoal for a p c x b - apply (subgoal_tac "padd (cmult c p) (times b a # cmult b p) \ []") - apply simp - apply (induct p) - apply auto - done - done +proof (induct p arbitrary: a x b) + case Nil + then show ?case by auto +next + case (Cons a p c x b) + then have "padd (cmult c p) (times b a # cmult b p) \ []" + by (metis local.padd.padd_Nil local.padd_Cons_Cons neq_Nil_conv) + then show ?case + by (simp add: local.Cons) +qed lemma (in semiring_1) last_linear_mul: assumes p: "p \ []" shows "last ([a, 1] *** p) = last p" proof - from p obtain c cs where cs: "p = c # cs" by (cases p) auto from cs have eq: "[a, 1] *** p = (a %* (c # cs)) +++ (0 # (1 %* (c # cs)))" by (simp add: poly_cmult_distr) show ?thesis using cs unfolding eq last_linear_mul_lemma by simp qed lemma (in semiring_0) pnormalize_eq: "last p \ 0 \ pnormalize p = p" by (induct p) (auto split: if_split_asm) lemma (in semiring_0) last_pnormalize: "pnormalize p \ [] \ last (pnormalize p) \ 0" by (induct p) auto lemma (in semiring_0) pnormal_degree: "last p \ 0 \ degree p = length p - 1" using pnormalize_eq[of p] unfolding degree_def by simp lemma (in semiring_0) poly_Nil_ext: "poly [] = (\x. 0)" by auto lemma (in idom_char_0) linear_mul_degree: assumes p: "poly p \ poly []" shows "degree ([a, 1] *** p) = degree p + 1" proof - from p have pnz: "pnormalize p \ []" unfolding poly_zero lemma_degree_zero . from last_linear_mul[OF pnz, of a] last_pnormalize[OF pnz] have l0: "last ([a, 1] *** pnormalize p) \ 0" by simp from last_pnormalize[OF pnz] last_linear_mul[OF pnz, of a] pnormal_degree[OF l0] pnormal_degree[OF last_pnormalize[OF pnz]] pnz have th: "degree ([a,1] *** pnormalize p) = degree (pnormalize p) + 1" by simp have eqs: "poly ([a,1] *** pnormalize p) = poly ([a,1] *** p)" by (rule ext) (simp add: poly_mult poly_add poly_cmult) from degree_unique[OF eqs] th show ?thesis by (simp add: degree_unique[OF poly_normalize]) qed lemma (in idom_char_0) linear_pow_mul_degree: "degree([a,1] %^n *** p) = (if poly p = poly [] then 0 else degree p + n)" proof (induct n arbitrary: a p) case (0 a p) show ?case proof (cases "poly p = poly []") case True then show ?thesis using degree_unique[OF True] by (simp add: degree_def) - next - case False - then show ?thesis - by (auto simp add: poly_Nil_ext) - qed + qed (auto simp add: poly_Nil_ext) next case (Suc n a p) have eq: "poly ([a, 1] %^(Suc n) *** p) = poly ([a, 1] %^ n *** ([a, 1] *** p))" - apply (rule ext) - apply (simp add: poly_mult poly_add poly_cmult) - apply (simp add: ac_simps distrib_left) - done + by (force simp add: poly_mult poly_add poly_cmult ac_simps distrib_left) note deq = degree_unique[OF eq] show ?case proof (cases "poly p = poly []") case True with eq have eq': "poly ([a, 1] %^(Suc n) *** p) = poly []" by (auto simp add: poly_mult poly_cmult poly_add) from degree_unique[OF eq'] True show ?thesis by (simp add: degree_def) next case False then have ap: "poly ([a,1] *** p) \ poly []" using poly_mult_not_eq_poly_Nil unfolding poly_entire by auto have eq: "poly ([a, 1] %^(Suc n) *** p) = poly ([a, 1]%^n *** ([a, 1] *** p))" by (auto simp add: poly_mult poly_add poly_exp poly_cmult algebra_simps) from ap have ap': "poly ([a, 1] *** p) = poly [] \ False" by blast have th0: "degree ([a, 1]%^n *** ([a, 1] *** p)) = degree ([a, 1] *** p) + n" - apply (simp only: Suc.hyps[of a "pmult [a,one] p"] ap') - apply simp - done + unfolding Suc.hyps[of a "pmult [a,one] p"] ap' by simp from degree_unique[OF eq] ap False th0 linear_mul_degree[OF False, of a] show ?thesis by (auto simp del: poly.simps) qed qed lemma (in idom_char_0) order_degree: assumes p0: "poly p \ poly []" shows "order a p \ degree p" proof - from order2[OF p0, unfolded divides_def] obtain q where q: "poly p = poly ([- a, 1]%^ (order a p) *** q)" by blast with q p0 have "poly q \ poly []" by (simp add: poly_mult poly_entire) with degree_unique[OF q, unfolded linear_pow_mul_degree] show ?thesis by auto qed text \Tidier versions of finiteness of roots.\ lemma (in idom_char_0) poly_roots_finite_set: "poly p \ poly [] \ finite {x. poly p x = 0}" unfolding poly_roots_finite . text \Bound for polynomial.\ lemma poly_mono: fixes x :: "'a::linordered_idom" shows "\x\ \ k \ \poly p x\ \ poly (map abs p) k" proof (induct p) case Nil then show ?case by simp next case (Cons a p) - then show ?case - apply auto - apply (rule_tac y = "\a\ + \x * poly p x\" in order_trans) - apply (rule abs_triangle_ineq) - apply (auto intro!: mult_mono simp add: abs_mult) - done + have "\a + x * poly p x\ \ \a\ + \x * poly p x\" + using abs_triangle_ineq by blast + also have "\ \ \a\ + k * poly (map abs p) k" + by (simp add: Cons.hyps Cons.prems abs_mult mult_mono') + finally show ?case + using Cons by auto qed lemma (in semiring_0) poly_Sing: "poly [c] x = c" by simp end