--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,3227 @@
+(* Title: HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy
+ Author: Amine Chaieb
+*)
+
+header{* A formalization of Ferrante and Rackoff's procedure with polynomial parameters, see Paper in CALCULEMUS 2008 *}
+
+theory Parametric_Ferrante_Rackoff
+imports Reflected_Multivariate_Polynomial
+ "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
+begin
+
+
+subsection {* Terms *}
+
+datatype tm = CP poly | Bound nat | Add tm tm | Mul poly tm
+ | Neg tm | Sub tm tm | CNP nat poly tm
+ (* A size for poly to make inductive proofs simpler*)
+
+consts tmsize :: "tm \<Rightarrow> nat"
+primrec
+ "tmsize (CP c) = polysize c"
+ "tmsize (Bound n) = 1"
+ "tmsize (Neg a) = 1 + tmsize a"
+ "tmsize (Add a b) = 1 + tmsize a + tmsize b"
+ "tmsize (Sub a b) = 3 + tmsize a + tmsize b"
+ "tmsize (Mul c a) = 1 + polysize c + tmsize a"
+ "tmsize (CNP n c a) = 3 + polysize c + tmsize a "
+
+ (* Semantics of terms tm *)
+consts Itm :: "'a::{ring_char_0,division_by_zero,field} list \<Rightarrow> 'a list \<Rightarrow> tm \<Rightarrow> 'a"
+primrec
+ "Itm vs bs (CP c) = (Ipoly vs c)"
+ "Itm vs bs (Bound n) = bs!n"
+ "Itm vs bs (Neg a) = -(Itm vs bs a)"
+ "Itm vs bs (Add a b) = Itm vs bs a + Itm vs bs b"
+ "Itm vs bs (Sub a b) = Itm vs bs a - Itm vs bs b"
+ "Itm vs bs (Mul c a) = (Ipoly vs c) * Itm vs bs a"
+ "Itm vs bs (CNP n c t) = (Ipoly vs c)*(bs!n) + Itm vs bs t"
+
+
+fun allpolys:: "(poly \<Rightarrow> bool) \<Rightarrow> tm \<Rightarrow> bool" where
+ "allpolys P (CP c) = P c"
+| "allpolys P (CNP n c p) = (P c \<and> allpolys P p)"
+| "allpolys P (Mul c p) = (P c \<and> allpolys P p)"
+| "allpolys P (Neg p) = allpolys P p"
+| "allpolys P (Add p q) = (allpolys P p \<and> allpolys P q)"
+| "allpolys P (Sub p q) = (allpolys P p \<and> allpolys P q)"
+| "allpolys P p = True"
+
+consts
+ tmboundslt:: "nat \<Rightarrow> tm \<Rightarrow> bool"
+ tmbound0:: "tm \<Rightarrow> bool" (* a tm is INDEPENDENT of Bound 0 *)
+ tmbound:: "nat \<Rightarrow> tm \<Rightarrow> bool" (* a tm is INDEPENDENT of Bound n *)
+ incrtm0:: "tm \<Rightarrow> tm"
+ incrtm:: "nat \<Rightarrow> tm \<Rightarrow> tm"
+ decrtm0:: "tm \<Rightarrow> tm"
+ decrtm:: "nat \<Rightarrow> tm \<Rightarrow> tm"
+primrec
+ "tmboundslt n (CP c) = True"
+ "tmboundslt n (Bound m) = (m < n)"
+ "tmboundslt n (CNP m c a) = (m < n \<and> tmboundslt n a)"
+ "tmboundslt n (Neg a) = tmboundslt n a"
+ "tmboundslt n (Add a b) = (tmboundslt n a \<and> tmboundslt n b)"
+ "tmboundslt n (Sub a b) = (tmboundslt n a \<and> tmboundslt n b)"
+ "tmboundslt n (Mul i a) = tmboundslt n a"
+primrec
+ "tmbound0 (CP c) = True"
+ "tmbound0 (Bound n) = (n>0)"
+ "tmbound0 (CNP n c a) = (n\<noteq>0 \<and> tmbound0 a)"
+ "tmbound0 (Neg a) = tmbound0 a"
+ "tmbound0 (Add a b) = (tmbound0 a \<and> tmbound0 b)"
+ "tmbound0 (Sub a b) = (tmbound0 a \<and> tmbound0 b)"
+ "tmbound0 (Mul i a) = tmbound0 a"
+lemma tmbound0_I:
+ assumes nb: "tmbound0 a"
+ shows "Itm vs (b#bs) a = Itm vs (b'#bs) a"
+using nb
+by (induct a rule: tmbound0.induct,auto simp add: nth_pos2)
+
+primrec
+ "tmbound n (CP c) = True"
+ "tmbound n (Bound m) = (n \<noteq> m)"
+ "tmbound n (CNP m c a) = (n\<noteq>m \<and> tmbound n a)"
+ "tmbound n (Neg a) = tmbound n a"
+ "tmbound n (Add a b) = (tmbound n a \<and> tmbound n b)"
+ "tmbound n (Sub a b) = (tmbound n a \<and> 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 \<le> length bs"
+ shows "Itm vs (bs[n:=x]) t = Itm vs bs t"
+ using nb le bnd
+ by (induct t rule: tmbound.induct , auto)
+
+recdef decrtm0 "measure size"
+ "decrtm0 (Bound n) = Bound (n - 1)"
+ "decrtm0 (Neg a) = Neg (decrtm0 a)"
+ "decrtm0 (Add a b) = Add (decrtm0 a) (decrtm0 b)"
+ "decrtm0 (Sub a b) = Sub (decrtm0 a) (decrtm0 b)"
+ "decrtm0 (Mul c a) = Mul c (decrtm0 a)"
+ "decrtm0 (CNP n c a) = CNP (n - 1) c (decrtm0 a)"
+ "decrtm0 a = a"
+recdef incrtm0 "measure size"
+ "incrtm0 (Bound n) = Bound (n + 1)"
+ "incrtm0 (Neg a) = Neg (incrtm0 a)"
+ "incrtm0 (Add a b) = Add (incrtm0 a) (incrtm0 b)"
+ "incrtm0 (Sub a b) = Sub (incrtm0 a) (incrtm0 b)"
+ "incrtm0 (Mul c a) = Mul c (incrtm0 a)"
+ "incrtm0 (CNP n c a) = CNP (n + 1) c (incrtm0 a)"
+ "incrtm0 a = a"
+lemma decrtm0: assumes nb: "tmbound0 t"
+ shows "Itm vs (x#bs) t = Itm vs bs (decrtm0 t)"
+ using nb by (induct t rule: decrtm0.induct, simp_all add: nth_pos2)
+lemma incrtm0: "Itm vs (x#bs) (incrtm0 t) = Itm vs bs t"
+ by (induct t rule: decrtm0.induct, simp_all add: nth_pos2)
+
+primrec
+ "decrtm m (CP c) = (CP c)"
+ "decrtm m (Bound n) = (if n < m then Bound n else Bound (n - 1))"
+ "decrtm m (Neg a) = Neg (decrtm m a)"
+ "decrtm m (Add a b) = Add (decrtm m a) (decrtm m b)"
+ "decrtm m (Sub a b) = Sub (decrtm m a) (decrtm m b)"
+ "decrtm m (Mul c a) = Mul c (decrtm m a)"
+ "decrtm m (CNP n c a) = (if n < m then CNP n c (decrtm m a) else CNP (n - 1) c (decrtm m a))"
+
+consts removen:: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
+primrec
+ "removen n [] = []"
+ "removen n (x#xs) = (if n=0 then xs else (x#(removen (n - 1) xs)))"
+
+lemma removen_same: "n \<ge> length xs \<Longrightarrow> removen n xs = xs"
+ by (induct xs arbitrary: n, auto)
+
+lemma nth_length_exceeds: "n \<ge> length xs \<Longrightarrow> xs!n = []!(n - length xs)"
+ by (induct xs arbitrary: n, auto)
+
+lemma removen_length: "length (removen n xs) = (if n \<ge> length xs then length xs else length xs - 1)"
+ by (induct xs arbitrary: n, auto)
+lemma removen_nth: "(removen n xs)!m = (if n \<ge> length xs then xs!m
+ else if m < n then xs!m else if m \<le> length xs then xs!(Suc m) else []!(m - (length xs - 1)))"
+proof(induct xs arbitrary: n m)
+ case Nil thus ?case by simp
+next
+ case (Cons x xs n m)
+ {assume nxs: "n \<ge> length (x#xs)" hence ?case using removen_same[OF nxs] by simp}
+ moreover
+ {assume nxs: "\<not> (n \<ge> length (x#xs))"
+ {assume mln: "m < n" hence ?case using prems by (cases m, auto)}
+ moreover
+ {assume mln: "\<not> (m < n)"
+
+ {assume mxs: "m \<le> length (x#xs)" hence ?case using prems by (cases m, auto)}
+ moreover
+ {assume mxs: "\<not> (m \<le> length (x#xs))"
+ have th: "length (removen n (x#xs)) = length xs"
+ using removen_length[where n="n" and xs="x#xs"] nxs by simp
+ with mxs have mxs':"m \<ge> length (removen n (x#xs))" by auto
+ hence "(removen n (x#xs))!m = [] ! (m - length xs)"
+ using th nth_length_exceeds[OF mxs'] by auto
+ hence th: "(removen n (x#xs))!m = [] ! (m - (length (x#xs) - 1))"
+ by auto
+ hence ?case using nxs mln mxs by auto }
+ ultimately have ?case by blast
+ }
+ ultimately have ?case by blast
+
+ } ultimately show ?case by blast
+qed
+
+lemma decrtm: assumes bnd: "tmboundslt (length bs) t" and nb: "tmbound m t"
+ and nle: "m \<le> length bs"
+ shows "Itm vs (removen m bs) (decrtm m t) = Itm vs bs t"
+ using bnd nb nle
+ by (induct t rule: decrtm.induct, auto simp add: removen_nth)
+
+consts tmsubst0:: "tm \<Rightarrow> tm \<Rightarrow> tm"
+primrec
+ "tmsubst0 t (CP c) = CP c"
+ "tmsubst0 t (Bound n) = (if n=0 then t else Bound n)"
+ "tmsubst0 t (CNP n c a) = (if n=0 then Add (Mul c t) (tmsubst0 t a) else CNP n c (tmsubst0 t a))"
+ "tmsubst0 t (Neg a) = Neg (tmsubst0 t a)"
+ "tmsubst0 t (Add a b) = Add (tmsubst0 t a) (tmsubst0 t b)"
+ "tmsubst0 t (Sub a b) = Sub (tmsubst0 t a) (tmsubst0 t b)"
+ "tmsubst0 t (Mul i a) = Mul i (tmsubst0 t a)"
+lemma tmsubst0:
+ shows "Itm vs (x#bs) (tmsubst0 t a) = Itm vs ((Itm vs (x#bs) t)#bs) a"
+by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2)
+
+lemma tmsubst0_nb: "tmbound0 t \<Longrightarrow> tmbound0 (tmsubst0 t a)"
+by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2)
+
+consts tmsubst:: "nat \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm"
+
+primrec
+ "tmsubst n t (CP c) = CP c"
+ "tmsubst n t (Bound m) = (if n=m then t else Bound m)"
+ "tmsubst n t (CNP m c a) = (if n=m then Add (Mul c t) (tmsubst n t a)
+ else CNP m c (tmsubst n t a))"
+ "tmsubst n t (Neg a) = Neg (tmsubst n t a)"
+ "tmsubst n t (Add a b) = Add (tmsubst n t a) (tmsubst n t b)"
+ "tmsubst n t (Sub a b) = Sub (tmsubst n t a) (tmsubst n t b)"
+ "tmsubst n t (Mul i a) = Mul i (tmsubst n t a)"
+
+lemma tmsubst: assumes nb: "tmboundslt (length bs) a" and nlt: "n \<le> length bs"
+ shows "Itm vs bs (tmsubst n t a) = Itm vs (bs[n:= Itm vs bs t]) a"
+using nb nlt
+by (induct a rule: tmsubst0.induct,auto simp add: nth_pos2)
+
+lemma tmsubst_nb0: assumes tnb: "tmbound0 t"
+shows "tmbound0 (tmsubst 0 t a)"
+using tnb
+by (induct a rule: tmsubst.induct, auto)
+
+lemma tmsubst_nb: assumes tnb: "tmbound m t"
+shows "tmbound m (tmsubst m t a)"
+using tnb
+by (induct a rule: tmsubst.induct, auto)
+lemma incrtm0_tmbound: "tmbound n t \<Longrightarrow> tmbound (Suc n) (incrtm0 t)"
+ by (induct t, auto)
+ (* Simplification *)
+
+consts
+ simptm:: "tm \<Rightarrow> tm"
+ tmadd:: "tm \<times> tm \<Rightarrow> tm"
+ tmmul:: "tm \<Rightarrow> poly \<Rightarrow> tm"
+recdef tmadd "measure (\<lambda> (t,s). size t + size s)"
+ "tmadd (CNP n1 c1 r1,CNP n2 c2 r2) =
+ (if n1=n2 then
+ (let c = c1 +\<^sub>p c2
+ in if c = 0\<^sub>p then tmadd(r1,r2) else CNP n1 c (tmadd (r1,r2)))
+ else if n1 \<le> n2 then (CNP n1 c1 (tmadd (r1,CNP n2 c2 r2)))
+ else (CNP n2 c2 (tmadd (CNP n1 c1 r1,r2))))"
+ "tmadd (CNP n1 c1 r1,t) = CNP n1 c1 (tmadd (r1, t))"
+ "tmadd (t,CNP n2 c2 r2) = CNP n2 c2 (tmadd (t,r2))"
+ "tmadd (CP b1, CP b2) = CP (b1 +\<^sub>p b2)"
+ "tmadd (a,b) = Add a b"
+
+lemma tmadd[simp]: "Itm vs bs (tmadd (t,s)) = Itm vs bs (Add t s)"
+apply (induct t s rule: tmadd.induct, simp_all add: Let_def)
+apply (case_tac "c1 +\<^sub>p c2 = 0\<^sub>p",case_tac "n1 \<le> n2", simp_all)
+apply (case_tac "n1 = n2", simp_all add: ring_simps)
+apply (simp only: right_distrib[symmetric])
+by (auto simp del: polyadd simp add: polyadd[symmetric])
+
+lemma tmadd_nb0[simp]: "\<lbrakk> tmbound0 t ; tmbound0 s\<rbrakk> \<Longrightarrow> tmbound0 (tmadd (t,s))"
+by (induct t s rule: tmadd.induct, auto simp add: Let_def)
+
+lemma tmadd_nb[simp]: "\<lbrakk> tmbound n t ; tmbound n s\<rbrakk> \<Longrightarrow> tmbound n (tmadd (t,s))"
+by (induct t s rule: tmadd.induct, auto simp add: Let_def)
+lemma tmadd_blt[simp]: "\<lbrakk>tmboundslt n t ; tmboundslt n s\<rbrakk> \<Longrightarrow> tmboundslt n (tmadd (t,s))"
+by (induct t s rule: tmadd.induct, auto simp add: Let_def)
+
+lemma tmadd_allpolys_npoly[simp]: "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmadd(t,s))" by (induct t s rule: tmadd.induct, simp_all add: Let_def polyadd_norm)
+
+recdef tmmul "measure size"
+ "tmmul (CP j) = (\<lambda> i. CP (i *\<^sub>p j))"
+ "tmmul (CNP n c a) = (\<lambda> i. CNP n (i *\<^sub>p c) (tmmul a i))"
+ "tmmul t = (\<lambda> i. Mul i t)"
+
+lemma tmmul[simp]: "Itm vs bs (tmmul t i) = Itm vs bs (Mul i t)"
+by (induct t arbitrary: i rule: tmmul.induct, simp_all add: ring_simps)
+
+lemma tmmul_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmmul t i)"
+by (induct t arbitrary: i rule: tmmul.induct, auto )
+
+lemma tmmul_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmmul t i)"
+by (induct t arbitrary: n rule: tmmul.induct, auto )
+lemma tmmul_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmmul t i)"
+by (induct t arbitrary: i rule: tmmul.induct, auto simp add: Let_def)
+
+lemma tmmul_allpolys_npoly[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm)
+
+constdefs tmneg :: "tm \<Rightarrow> tm"
+ "tmneg t \<equiv> tmmul t (C (- 1,1))"
+
+constdefs tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm"
+ "tmsub s t \<equiv> (if s = t then CP 0\<^sub>p else tmadd (s,tmneg t))"
+
+lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)"
+using tmneg_def[of t]
+apply simp
+apply (subst number_of_Min)
+apply (simp only: of_int_minus)
+apply simp
+done
+
+lemma tmneg_nb0[simp]: "tmbound0 t \<Longrightarrow> tmbound0 (tmneg t)"
+using tmneg_def by simp
+
+lemma tmneg_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (tmneg t)"
+using tmneg_def by simp
+lemma tmneg_blt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (tmneg t)"
+using tmneg_def by simp
+lemma [simp]: "isnpoly (C (-1,1))" unfolding isnpoly_def by simp
+lemma tmneg_allpolys_npoly[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly (tmneg t)"
+ unfolding tmneg_def by auto
+
+lemma tmsub[simp]: "Itm vs bs (tmsub a b) = Itm vs bs (Sub a b)"
+using tmsub_def by simp
+
+lemma tmsub_nb0[simp]: "\<lbrakk> tmbound0 t ; tmbound0 s\<rbrakk> \<Longrightarrow> tmbound0 (tmsub t s)"
+using tmsub_def by simp
+lemma tmsub_nb[simp]: "\<lbrakk> tmbound n t ; tmbound n s\<rbrakk> \<Longrightarrow> tmbound n (tmsub t s)"
+using tmsub_def by simp
+lemma tmsub_blt[simp]: "\<lbrakk>tmboundslt n t ; tmboundslt n s\<rbrakk> \<Longrightarrow> tmboundslt n (tmsub t s )"
+using tmsub_def by simp
+lemma tmsub_allpolys_npoly[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ shows "allpolys isnpoly t \<Longrightarrow> allpolys isnpoly s \<Longrightarrow> allpolys isnpoly (tmsub t s)"
+ unfolding tmsub_def by (simp add: isnpoly_def)
+
+recdef simptm "measure size"
+ "simptm (CP j) = CP (polynate j)"
+ "simptm (Bound n) = CNP n 1\<^sub>p (CP 0\<^sub>p)"
+ "simptm (Neg t) = tmneg (simptm t)"
+ "simptm (Add t s) = tmadd (simptm t,simptm s)"
+ "simptm (Sub t s) = tmsub (simptm t) (simptm s)"
+ "simptm (Mul i t) = (let i' = polynate i in if i' = 0\<^sub>p then CP 0\<^sub>p else tmmul (simptm t) i')"
+ "simptm (CNP n c t) = (let c' = polynate c in if c' = 0\<^sub>p then simptm t else tmadd (CNP n c' (CP 0\<^sub>p ), simptm t))"
+
+lemma polynate_stupid:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ shows "polynate t = 0\<^sub>p \<Longrightarrow> Ipoly bs t = (0::'a::{ring_char_0,division_by_zero, field})"
+apply (subst polynate[symmetric])
+apply simp
+done
+
+lemma simptm_ci[simp]: "Itm vs bs (simptm t) = Itm vs bs t"
+by (induct t rule: simptm.induct, auto simp add: tmneg tmadd tmsub tmmul Let_def polynate_stupid)
+
+lemma simptm_tmbound0[simp]:
+ "tmbound0 t \<Longrightarrow> tmbound0 (simptm t)"
+by (induct t rule: simptm.induct, auto simp add: Let_def)
+
+lemma simptm_nb[simp]: "tmbound n t \<Longrightarrow> tmbound n (simptm t)"
+by (induct t rule: simptm.induct, auto simp add: Let_def)
+lemma simptm_nlt[simp]: "tmboundslt n t \<Longrightarrow> tmboundslt n (simptm t)"
+by (induct t rule: simptm.induct, auto simp add: Let_def)
+
+lemma [simp]: "isnpoly 0\<^sub>p" and [simp]: "isnpoly (C(1,1))"
+ by (simp_all add: isnpoly_def)
+lemma simptm_allpolys_npoly[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ shows "allpolys isnpoly (simptm p)"
+ by (induct p rule: simptm.induct, auto simp add: Let_def)
+
+consts split0 :: "tm \<Rightarrow> (poly \<times> tm)"
+recdef split0 "measure tmsize"
+ "split0 (Bound 0) = (1\<^sub>p, CP 0\<^sub>p)"
+ "split0 (CNP 0 c t) = (let (c',t') = split0 t in (c +\<^sub>p c',t'))"
+ "split0 (Neg t) = (let (c,t') = split0 t in (~\<^sub>p c,Neg t'))"
+ "split0 (CNP n c t) = (let (c',t') = split0 t in (c',CNP n c t'))"
+ "split0 (Add s t) = (let (c1,s') = split0 s ; (c2,t') = split0 t in (c1 +\<^sub>p c2, Add s' t'))"
+ "split0 (Sub s t) = (let (c1,s') = split0 s ; (c2,t') = split0 t in (c1 -\<^sub>p c2, Sub s' t'))"
+ "split0 (Mul c t) = (let (c',t') = split0 t in (c *\<^sub>p c', Mul c t'))"
+ "split0 t = (0\<^sub>p, t)"
+
+lemma split0_stupid[simp]: "\<exists>x y. (x,y) = split0 p"
+ apply (rule exI[where x="fst (split0 p)"])
+ apply (rule exI[where x="snd (split0 p)"])
+ by simp
+
+lemma split0:
+ "tmbound 0 (snd (split0 t)) \<and> (Itm vs bs (CNP 0 (fst (split0 t)) (snd (split0 t))) = Itm vs bs t)"
+ apply (induct t rule: split0.induct)
+ apply simp
+ apply (simp add: Let_def split_def ring_simps)
+ apply (simp add: Let_def split_def ring_simps)
+ apply (simp add: Let_def split_def ring_simps)
+ apply (simp add: Let_def split_def ring_simps)
+ apply (simp add: Let_def split_def ring_simps)
+ apply (simp add: Let_def split_def mult_assoc right_distrib[symmetric])
+ apply (simp add: Let_def split_def ring_simps)
+ apply (simp add: Let_def split_def ring_simps)
+ done
+
+lemma split0_ci: "split0 t = (c',t') \<Longrightarrow> Itm vs bs t = Itm vs bs (CNP 0 c' t')"
+proof-
+ fix c' t'
+ assume "split0 t = (c', t')" hence "c' = fst (split0 t)" and "t' = snd (split0 t)" by auto
+ with split0[where t="t" and bs="bs"] show "Itm vs bs t = Itm vs bs (CNP 0 c' t')" by simp
+qed
+
+lemma split0_nb0:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ shows "split0 t = (c',t') \<Longrightarrow> tmbound 0 t'"
+proof-
+ fix c' t'
+ assume "split0 t = (c', t')" hence "c' = fst (split0 t)" and "t' = snd (split0 t)" by auto
+ with conjunct1[OF split0[where t="t"]] show "tmbound 0 t'" by simp
+qed
+
+lemma split0_nb0'[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ shows "tmbound0 (snd (split0 t))"
+ using split0_nb0[of t "fst (split0 t)" "snd (split0 t)"] by (simp add: tmbound0_tmbound_iff)
+
+
+lemma split0_nb: assumes nb:"tmbound n t" shows "tmbound n (snd (split0 t))"
+ using nb by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma split0_blt: assumes nb:"tmboundslt n t" shows "tmboundslt n (snd (split0 t))"
+ using nb by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma tmbound_split0: "tmbound 0 t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0"
+ by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma tmboundslt_split0: "tmboundslt n t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0 \<or> n > 0"
+by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma tmboundslt0_split0: "tmboundslt 0 t \<Longrightarrow> Ipoly vs (fst(split0 t)) = 0"
+ by (induct t rule: split0.induct, auto simp add: Let_def split_def split0_stupid)
+
+lemma allpolys_split0: "allpolys isnpoly p \<Longrightarrow> allpolys isnpoly (snd (split0 p))"
+by (induct p rule: split0.induct, auto simp add: isnpoly_def Let_def split_def split0_stupid)
+
+lemma isnpoly_fst_split0: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
+ shows
+ "allpolys isnpoly p \<Longrightarrow> isnpoly (fst (split0 p))"
+ by (induct p rule: split0.induct,
+ auto simp add: polyadd_norm polysub_norm polyneg_norm polymul_norm
+ Let_def split_def split0_stupid)
+
+subsection{* Formulae *}
+
+datatype fm = T| F| Le tm | Lt tm | Eq tm | NEq tm|
+ NOT fm| And fm fm| Or fm fm| Imp fm fm| Iff fm fm| E fm| A fm
+
+
+ (* A size for fm *)
+consts fmsize :: "fm \<Rightarrow> nat"
+recdef fmsize "measure size"
+ "fmsize (NOT p) = 1 + fmsize p"
+ "fmsize (And p q) = 1 + fmsize p + fmsize q"
+ "fmsize (Or p q) = 1 + fmsize p + fmsize q"
+ "fmsize (Imp p q) = 3 + fmsize p + fmsize q"
+ "fmsize (Iff p q) = 3 + 2*(fmsize p + fmsize q)"
+ "fmsize (E p) = 1 + fmsize p"
+ "fmsize (A p) = 4+ fmsize p"
+ "fmsize p = 1"
+ (* several lemmas about fmsize *)
+lemma fmsize_pos: "fmsize p > 0"
+by (induct p rule: fmsize.induct) simp_all
+
+ (* Semantics of formulae (fm) *)
+consts Ifm ::"'a::{division_by_zero,ordered_field} list \<Rightarrow> 'a list \<Rightarrow> fm \<Rightarrow> bool"
+primrec
+ "Ifm vs bs T = True"
+ "Ifm vs bs F = False"
+ "Ifm vs bs (Lt a) = (Itm vs bs a < 0)"
+ "Ifm vs bs (Le a) = (Itm vs bs a \<le> 0)"
+ "Ifm vs bs (Eq a) = (Itm vs bs a = 0)"
+ "Ifm vs bs (NEq a) = (Itm vs bs a \<noteq> 0)"
+ "Ifm vs bs (NOT p) = (\<not> (Ifm vs bs p))"
+ "Ifm vs bs (And p q) = (Ifm vs bs p \<and> Ifm vs bs q)"
+ "Ifm vs bs (Or p q) = (Ifm vs bs p \<or> Ifm vs bs q)"
+ "Ifm vs bs (Imp p q) = ((Ifm vs bs p) \<longrightarrow> (Ifm vs bs q))"
+ "Ifm vs bs (Iff p q) = (Ifm vs bs p = Ifm vs bs q)"
+ "Ifm vs bs (E p) = (\<exists> x. Ifm vs (x#bs) p)"
+ "Ifm vs bs (A p) = (\<forall> x. Ifm vs (x#bs) p)"
+
+consts not:: "fm \<Rightarrow> fm"
+recdef not "measure size"
+ "not (NOT (NOT p)) = not p"
+ "not (NOT p) = p"
+ "not T = F"
+ "not F = T"
+ "not (Lt t) = Le (tmneg t)"
+ "not (Le t) = Lt (tmneg t)"
+ "not (Eq t) = NEq t"
+ "not (NEq t) = Eq t"
+ "not p = NOT p"
+lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)"
+by (induct p rule: not.induct) auto
+
+constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+ "conj p q \<equiv> (if (p = F \<or> 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 \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
+
+constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+ "disj p q \<equiv> (if (p = T \<or> 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 \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
+
+constdefs imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+ "imp p q \<equiv> (if (p = F \<or> q=T \<or> 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 \<or> q=T",simp_all add: imp_def)
+
+constdefs iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+ "iff p q \<equiv> (if (p = q) then T else if (p = NOT q \<or> NOT p = q) then F else
+ if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else
+ Iff p q)"
+lemma iff[simp]: "Ifm vs bs (iff p q) = Ifm vs bs (Iff p q)"
+ by (unfold iff_def,cases "p=q", simp,cases "p=NOT q", simp) (cases "NOT p= q", auto)
+ (* Quantifier freeness *)
+consts qfree:: "fm \<Rightarrow> bool"
+recdef qfree "measure size"
+ "qfree (E p) = False"
+ "qfree (A p) = False"
+ "qfree (NOT p) = qfree p"
+ "qfree (And p q) = (qfree p \<and> qfree q)"
+ "qfree (Or p q) = (qfree p \<and> qfree q)"
+ "qfree (Imp p q) = (qfree p \<and> qfree q)"
+ "qfree (Iff p q) = (qfree p \<and> qfree q)"
+ "qfree p = True"
+
+ (* Boundedness and substitution *)
+
+consts boundslt :: "nat \<Rightarrow> fm \<Rightarrow> bool"
+primrec
+ "boundslt n T = True"
+ "boundslt n F = True"
+ "boundslt n (Lt t) = (tmboundslt n t)"
+ "boundslt n (Le t) = (tmboundslt n t)"
+ "boundslt n (Eq t) = (tmboundslt n t)"
+ "boundslt n (NEq t) = (tmboundslt n t)"
+ "boundslt n (NOT p) = boundslt n p"
+ "boundslt n (And p q) = (boundslt n p \<and> boundslt n q)"
+ "boundslt n (Or p q) = (boundslt n p \<and> boundslt n q)"
+ "boundslt n (Imp p q) = ((boundslt n p) \<and> (boundslt n q))"
+ "boundslt n (Iff p q) = (boundslt n p \<and> boundslt n q)"
+ "boundslt n (E p) = boundslt (Suc n) p"
+ "boundslt n (A p) = boundslt (Suc n) p"
+
+consts
+ bound0:: "fm \<Rightarrow> bool" (* A Formula is independent of Bound 0 *)
+ bound:: "nat \<Rightarrow> fm \<Rightarrow> bool" (* A Formula is independent of Bound n *)
+ decr0 :: "fm \<Rightarrow> fm"
+ decr :: "nat \<Rightarrow> fm \<Rightarrow> fm"
+recdef bound0 "measure size"
+ "bound0 T = True"
+ "bound0 F = True"
+ "bound0 (Lt a) = tmbound0 a"
+ "bound0 (Le a) = tmbound0 a"
+ "bound0 (Eq a) = tmbound0 a"
+ "bound0 (NEq a) = tmbound0 a"
+ "bound0 (NOT p) = bound0 p"
+ "bound0 (And p q) = (bound0 p \<and> bound0 q)"
+ "bound0 (Or p q) = (bound0 p \<and> bound0 q)"
+ "bound0 (Imp p q) = ((bound0 p) \<and> (bound0 q))"
+ "bound0 (Iff p q) = (bound0 p \<and> bound0 q)"
+ "bound0 p = False"
+lemma bound0_I:
+ assumes bp: "bound0 p"
+ shows "Ifm vs (b#bs) p = Ifm vs (b'#bs) p"
+using bp tmbound0_I[where b="b" and bs="bs" and b'="b'"]
+by (induct p rule: bound0.induct,auto simp add: nth_pos2)
+
+primrec
+ "bound m T = True"
+ "bound m F = True"
+ "bound m (Lt t) = tmbound m t"
+ "bound m (Le t) = tmbound m t"
+ "bound m (Eq t) = tmbound m t"
+ "bound m (NEq t) = tmbound m t"
+ "bound m (NOT p) = bound m p"
+ "bound m (And p q) = (bound m p \<and> bound m q)"
+ "bound m (Or p q) = (bound m p \<and> bound m q)"
+ "bound m (Imp p q) = ((bound m p) \<and> (bound m q))"
+ "bound m (Iff p q) = (bound m p \<and> 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 \<le> length bs"
+ shows "Ifm vs (bs[n:=x]) p = Ifm vs bs p"
+ using bnd nb le tmbound_I[where bs=bs and vs = vs]
+proof(induct p arbitrary: bs n rule: bound.induct)
+ case (E p bs n)
+ {fix y
+ from prems have bnd: "boundslt (length (y#bs)) p"
+ and nb: "bound (Suc n) p" and le: "Suc n \<le> length (y#bs)" by simp+
+ from E.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" . }
+ thus ?case by simp
+next
+ case (A p bs n) {fix y
+ from prems have bnd: "boundslt (length (y#bs)) p"
+ and nb: "bound (Suc n) p" and le: "Suc n \<le> length (y#bs)" by simp+
+ from A.hyps[OF bnd nb le tmbound_I] have "Ifm vs ((y#bs)[Suc n:=x]) p = Ifm vs (y#bs) p" . }
+ thus ?case by simp
+qed auto
+
+recdef decr0 "measure size"
+ "decr0 (Lt a) = Lt (decrtm0 a)"
+ "decr0 (Le a) = Le (decrtm0 a)"
+ "decr0 (Eq a) = Eq (decrtm0 a)"
+ "decr0 (NEq a) = NEq (decrtm0 a)"
+ "decr0 (NOT p) = NOT (decr0 p)"
+ "decr0 (And p q) = conj (decr0 p) (decr0 q)"
+ "decr0 (Or p q) = disj (decr0 p) (decr0 q)"
+ "decr0 (Imp p q) = imp (decr0 p) (decr0 q)"
+ "decr0 (Iff p q) = iff (decr0 p) (decr0 q)"
+ "decr0 p = p"
+
+lemma decr0: assumes nb: "bound0 p"
+ shows "Ifm vs (x#bs) p = Ifm vs bs (decr0 p)"
+ using nb
+ by (induct p rule: decr0.induct, simp_all add: decrtm0)
+
+primrec
+ "decr m T = T"
+ "decr m F = F"
+ "decr m (Lt t) = (Lt (decrtm m t))"
+ "decr m (Le t) = (Le (decrtm m t))"
+ "decr m (Eq t) = (Eq (decrtm m t))"
+ "decr m (NEq t) = (NEq (decrtm m t))"
+ "decr m (NOT p) = NOT (decr m p)"
+ "decr m (And p q) = conj (decr m p) (decr m q)"
+ "decr m (Or p q) = disj (decr m p) (decr m q)"
+ "decr m (Imp p q) = imp (decr m p) (decr m q)"
+ "decr m (Iff p q) = iff (decr m p) (decr m q)"
+ "decr m (E p) = E (decr (Suc m) p)"
+ "decr m (A p) = A (decr (Suc m) p)"
+
+lemma decr: assumes bnd: "boundslt (length bs) p" and nb: "bound m p"
+ and nle: "m < length bs"
+ shows "Ifm vs (removen m bs) (decr m p) = Ifm vs bs p"
+ using bnd nb nle
+proof(induct p arbitrary: bs m rule: decr.induct)
+ case (E p bs m)
+ {fix x
+ from prems have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p"
+ and nle: "Suc m < length (x#bs)" by auto
+ from prems(4)[OF bnd nb nle] have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p".
+ } thus ?case by auto
+next
+ case (A p bs m)
+ {fix x
+ from prems have bnd: "boundslt (length (x#bs)) p" and nb: "bound (Suc m) p"
+ and nle: "Suc m < length (x#bs)" by auto
+ from prems(4)[OF bnd nb nle] have "Ifm vs (removen (Suc m) (x#bs)) (decr (Suc m) p) = Ifm vs (x#bs) p".
+ } thus ?case by auto
+qed (auto simp add: decrtm removen_nth)
+
+consts
+ subst0:: "tm \<Rightarrow> fm \<Rightarrow> fm"
+
+primrec
+ "subst0 t T = T"
+ "subst0 t F = F"
+ "subst0 t (Lt a) = Lt (tmsubst0 t a)"
+ "subst0 t (Le a) = Le (tmsubst0 t a)"
+ "subst0 t (Eq a) = Eq (tmsubst0 t a)"
+ "subst0 t (NEq a) = NEq (tmsubst0 t a)"
+ "subst0 t (NOT p) = NOT (subst0 t p)"
+ "subst0 t (And p q) = And (subst0 t p) (subst0 t q)"
+ "subst0 t (Or p q) = Or (subst0 t p) (subst0 t q)"
+ "subst0 t (Imp p q) = Imp (subst0 t p) (subst0 t q)"
+ "subst0 t (Iff p q) = Iff (subst0 t p) (subst0 t q)"
+ "subst0 t (E p) = E p"
+ "subst0 t (A p) = A p"
+
+lemma subst0: assumes qf: "qfree p"
+ shows "Ifm vs (x#bs) (subst0 t p) = Ifm vs ((Itm vs (x#bs) t)#bs) p"
+using qf tmsubst0[where x="x" and bs="bs" and t="t"]
+by (induct p rule: subst0.induct, auto)
+
+lemma subst0_nb:
+ assumes bp: "tmbound0 t" and qf: "qfree p"
+ shows "bound0 (subst0 t p)"
+using qf tmsubst0_nb[OF bp] bp
+by (induct p rule: subst0.induct, auto)
+
+consts subst:: "nat \<Rightarrow> tm \<Rightarrow> fm \<Rightarrow> fm"
+primrec
+ "subst n t T = T"
+ "subst n t F = F"
+ "subst n t (Lt a) = Lt (tmsubst n t a)"
+ "subst n t (Le a) = Le (tmsubst n t a)"
+ "subst n t (Eq a) = Eq (tmsubst n t a)"
+ "subst n t (NEq a) = NEq (tmsubst n t a)"
+ "subst n t (NOT p) = NOT (subst n t p)"
+ "subst n t (And p q) = And (subst n t p) (subst n t q)"
+ "subst n t (Or p q) = Or (subst n t p) (subst n t q)"
+ "subst n t (Imp p q) = Imp (subst n t p) (subst n t q)"
+ "subst n t (Iff p q) = Iff (subst n t p) (subst n t q)"
+ "subst n t (E p) = E (subst (Suc n) (incrtm0 t) p)"
+ "subst n t (A p) = A (subst (Suc n) (incrtm0 t) p)"
+
+lemma subst: assumes nb: "boundslt (length bs) p" and nlm: "n \<le> length bs"
+ shows "Ifm vs bs (subst n t p) = Ifm vs (bs[n:= Itm vs bs t]) p"
+ using nb nlm
+proof (induct p arbitrary: bs n t rule: subst0.induct)
+ case (E p bs n)
+ {fix x
+ from prems have bn: "boundslt (length (x#bs)) p" by simp
+ from prems have nlm: "Suc n \<le> length (x#bs)" by simp
+ from prems(3)[OF bn nlm] have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p" by simp
+ hence "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p"
+ by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) }
+thus ?case by simp
+next
+ case (A p bs n)
+ {fix x
+ from prems have bn: "boundslt (length (x#bs)) p" by simp
+ from prems have nlm: "Suc n \<le> length (x#bs)" by simp
+ from prems(3)[OF bn nlm] have "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs ((x#bs)[Suc n:= Itm vs (x#bs) (incrtm0 t)]) p" by simp
+ hence "Ifm vs (x#bs) (subst (Suc n) (incrtm0 t) p) = Ifm vs (x#bs[n:= Itm vs bs t]) p"
+ by (simp add: incrtm0[where x="x" and bs="bs" and t="t"]) }
+thus ?case by simp
+qed(auto simp add: tmsubst)
+
+lemma subst_nb: assumes tnb: "tmbound m t"
+shows "bound m (subst m t p)"
+using tnb tmsubst_nb incrtm0_tmbound
+by (induct p arbitrary: m t rule: subst.induct, auto)
+
+lemma not_qf[simp]: "qfree p \<Longrightarrow> qfree (not p)"
+by (induct p rule: not.induct, auto)
+lemma not_bn0[simp]: "bound0 p \<Longrightarrow> bound0 (not p)"
+by (induct p rule: not.induct, auto)
+lemma not_nb[simp]: "bound n p \<Longrightarrow> bound n (not p)"
+by (induct p rule: not.induct, auto)
+lemma not_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n (not p)"
+ by (induct p rule: not.induct, auto)
+
+lemma conj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (conj p q)"
+using conj_def by auto
+lemma conj_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
+using conj_def by auto
+lemma conj_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (conj p q)"
+using conj_def by auto
+lemma conj_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (conj p q)"
+using conj_def by auto
+
+lemma disj_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (disj p q)"
+using disj_def by auto
+lemma disj_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
+using disj_def by auto
+lemma disj_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (disj p q)"
+using disj_def by auto
+lemma disj_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (disj p q)"
+using disj_def by auto
+
+lemma imp_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (imp p q)"
+using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def)
+lemma imp_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
+using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def)
+lemma imp_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (imp p q)"
+using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def)
+lemma imp_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (imp p q)"
+using imp_def by auto
+
+lemma iff_qf[simp]: "\<lbrakk>qfree p ; qfree q\<rbrakk> \<Longrightarrow> qfree (iff p q)"
+ by (unfold iff_def,cases "p=q", auto)
+lemma iff_nb0[simp]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (iff p q)"
+using iff_def by (unfold iff_def,cases "p=q", auto)
+lemma iff_nb[simp]: "\<lbrakk>bound n p ; bound n q\<rbrakk> \<Longrightarrow> bound n (iff p q)"
+using iff_def by (unfold iff_def,cases "p=q", auto)
+lemma iff_blt[simp]: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (iff p q)"
+using iff_def by auto
+lemma decr0_qf: "bound0 p \<Longrightarrow> qfree (decr0 p)"
+by (induct p, simp_all)
+
+consts
+ isatom :: "fm \<Rightarrow> bool" (* test for atomicity *)
+recdef isatom "measure size"
+ "isatom T = True"
+ "isatom F = True"
+ "isatom (Lt a) = True"
+ "isatom (Le a) = True"
+ "isatom (Eq a) = True"
+ "isatom (NEq a) = True"
+ "isatom p = False"
+
+lemma bound0_qf: "bound0 p \<Longrightarrow> qfree p"
+by (induct p, simp_all)
+
+constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+ "djf f p q \<equiv> (if q=T then T else if q=F then f p else
+ (let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
+constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+ "evaldjf f ps \<equiv> foldr (djf f) ps F"
+
+lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)"
+by (cases "q=T", simp add: djf_def,cases "q=F",simp add: djf_def)
+(cases "f p", simp_all add: Let_def djf_def)
+
+lemma evaldjf_ex: "Ifm vs bs (evaldjf f ps) = (\<exists> p \<in> set ps. Ifm vs bs (f p))"
+ by(induct ps, simp_all add: evaldjf_def djf_Or)
+
+lemma evaldjf_bound0:
+ assumes nb: "\<forall> x\<in> set xs. bound0 (f x)"
+ shows "bound0 (evaldjf f xs)"
+ using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto)
+
+lemma evaldjf_qf:
+ assumes nb: "\<forall> x\<in> set xs. qfree (f x)"
+ shows "qfree (evaldjf f xs)"
+ using nb by (induct xs, auto simp add: evaldjf_def djf_def Let_def) (case_tac "f a", auto)
+
+consts disjuncts :: "fm \<Rightarrow> fm list"
+recdef disjuncts "measure size"
+ "disjuncts (Or p q) = (disjuncts p) @ (disjuncts q)"
+ "disjuncts F = []"
+ "disjuncts p = [p]"
+
+lemma disjuncts: "(\<exists> q\<in> set (disjuncts p). Ifm vs bs q) = Ifm vs bs p"
+by(induct p rule: disjuncts.induct, auto)
+
+lemma disjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). bound0 q"
+proof-
+ assume nb: "bound0 p"
+ hence "list_all bound0 (disjuncts p)" by (induct p rule:disjuncts.induct,auto)
+ thus ?thesis by (simp only: list_all_iff)
+qed
+
+lemma disjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (disjuncts p). qfree q"
+proof-
+ assume qf: "qfree p"
+ hence "list_all qfree (disjuncts p)"
+ by (induct p rule: disjuncts.induct, auto)
+ thus ?thesis by (simp only: list_all_iff)
+qed
+
+constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+ "DJ f p \<equiv> evaldjf f (disjuncts p)"
+
+lemma DJ: assumes fdj: "\<forall> 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) = (\<exists> q \<in> set (disjuncts p). Ifm vs bs (f q))"
+ by (simp add: DJ_def evaldjf_ex)
+ also have "\<dots> = Ifm vs bs (f p)" using fdj fF by (induct p rule: disjuncts.induct, auto)
+ finally show ?thesis .
+qed
+
+lemma DJ_qf: assumes
+ fqf: "\<forall> p. qfree p \<longrightarrow> qfree (f p)"
+ shows "\<forall>p. qfree p \<longrightarrow> 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 "\<forall> q\<in> set (disjuncts p). qfree q" .
+ with fqf have th':"\<forall> q\<in> 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: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
+ shows "\<forall> bs p. qfree p \<longrightarrow> qfree (DJ qe p) \<and> (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: "\<forall> p. qfree p \<longrightarrow> 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) = (\<exists> q\<in> set (disjuncts p). Ifm vs bs (qe q))"
+ by (simp add: DJ_def evaldjf_ex)
+ also have "\<dots> = (\<exists> q \<in> set(disjuncts p). Ifm vs bs (E q))" using qe disjuncts_qf[OF qf] by auto
+ also have "\<dots> = Ifm vs bs (E p)" by (induct p rule: disjuncts.induct, auto)
+ finally show "qfree (DJ qe p) \<and> Ifm vs bs (DJ qe p) = Ifm vs bs (E p)" using qfth by blast
+qed
+
+consts conjuncts :: "fm \<Rightarrow> fm list"
+
+recdef conjuncts "measure size"
+ "conjuncts (And p q) = (conjuncts p) @ (conjuncts q)"
+ "conjuncts T = []"
+ "conjuncts p = [p]"
+
+constdefs list_conj :: "fm list \<Rightarrow> fm"
+ "list_conj ps \<equiv> foldr conj ps T"
+
+constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+ "CJNB f p \<equiv> (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs
+ in conj (decr0 (list_conj yes)) (f (list_conj no)))"
+
+lemma conjuncts_qf: "qfree p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). qfree q"
+proof-
+ assume qf: "qfree p"
+ hence "list_all qfree (conjuncts p)"
+ by (induct p rule: conjuncts.induct, auto)
+ thus ?thesis by (simp only: list_all_iff)
+qed
+
+lemma conjuncts: "(\<forall> q\<in> set (conjuncts p). Ifm vs bs q) = Ifm vs bs p"
+by(induct p rule: conjuncts.induct, auto)
+
+lemma conjuncts_nb: "bound0 p \<Longrightarrow> \<forall> q\<in> set (conjuncts p). bound0 q"
+proof-
+ assume nb: "bound0 p"
+ hence "list_all bound0 (conjuncts p)" by (induct p rule:conjuncts.induct,auto)
+ thus ?thesis by (simp only: list_all_iff)
+qed
+
+fun islin :: "fm \<Rightarrow> bool" where
+ "islin (And p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
+| "islin (Or p q) = (islin p \<and> islin q \<and> p \<noteq> T \<and> p \<noteq> F \<and> q \<noteq> T \<and> q \<noteq> F)"
+| "islin (Eq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+| "islin (NEq (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+| "islin (Lt (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+| "islin (Le (CNP 0 c s)) = (isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s)"
+| "islin (NOT p) = False"
+| "islin (Imp p q) = False"
+| "islin (Iff p q) = False"
+| "islin p = bound0 p"
+
+lemma islin_stupid: assumes nb: "tmbound0 p"
+ shows "islin (Lt p)" and "islin (Le p)" and "islin (Eq p)" and "islin (NEq p)"
+ using nb by (cases p, auto, case_tac nat, auto)+
+
+definition "lt p = (case p of CP (C c) \<Rightarrow> if 0>\<^sub>N c then T else F| _ \<Rightarrow> Lt p)"
+definition "le p = (case p of CP (C c) \<Rightarrow> if 0\<ge>\<^sub>N c then T else F | _ \<Rightarrow> Le p)"
+definition "eq p = (case p of CP (C c) \<Rightarrow> if c = 0\<^sub>N then T else F | _ \<Rightarrow> Eq p)"
+definition "neq p = not (eq p)"
+
+lemma lt: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (lt p) = Ifm vs bs (Lt p)"
+ apply(simp add: lt_def)
+ apply(cases p, simp_all)
+ apply (case_tac poly, simp_all add: isnpoly_def)
+ done
+
+lemma le: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (le p) = Ifm vs bs (Le p)"
+ apply(simp add: le_def)
+ apply(cases p, simp_all)
+ apply (case_tac poly, simp_all add: isnpoly_def)
+ done
+
+lemma eq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (eq p) = Ifm vs bs (Eq p)"
+ apply(simp add: eq_def)
+ apply(cases p, simp_all)
+ apply (case_tac poly, simp_all add: isnpoly_def)
+ done
+
+lemma neq: "allpolys isnpoly p \<Longrightarrow> Ifm vs bs (neq p) = Ifm vs bs (NEq p)"
+ by(simp add: neq_def eq)
+
+lemma lt_lin: "tmbound0 p \<Longrightarrow> islin (lt p)"
+ apply (simp add: lt_def)
+ apply (cases p, simp_all)
+ apply (case_tac poly, simp_all)
+ apply (case_tac nat, simp_all)
+ done
+
+lemma le_lin: "tmbound0 p \<Longrightarrow> islin (le p)"
+ apply (simp add: le_def)
+ apply (cases p, simp_all)
+ apply (case_tac poly, simp_all)
+ apply (case_tac nat, simp_all)
+ done
+
+lemma eq_lin: "tmbound0 p \<Longrightarrow> islin (eq p)"
+ apply (simp add: eq_def)
+ apply (cases p, simp_all)
+ apply (case_tac poly, simp_all)
+ apply (case_tac nat, simp_all)
+ done
+
+lemma neq_lin: "tmbound0 p \<Longrightarrow> islin (neq p)"
+ apply (simp add: neq_def eq_def)
+ apply (cases p, simp_all)
+ apply (case_tac poly, simp_all)
+ apply (case_tac nat, simp_all)
+ done
+
+definition "simplt t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then lt s else Lt (CNP 0 c s))"
+definition "simple t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then le s else Le (CNP 0 c s))"
+definition "simpeq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then eq s else Eq (CNP 0 c s))"
+definition "simpneq t = (let (c,s) = split0 (simptm t) in if c= 0\<^sub>p then neq s else NEq (CNP 0 c s))"
+
+lemma simplt_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "islin (simplt t)"
+ unfolding simplt_def
+ using split0_nb0'
+by (auto simp add: lt_lin Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly])
+
+lemma simple_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "islin (simple t)"
+ unfolding simple_def
+ using split0_nb0'
+by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] le_lin)
+lemma simpeq_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "islin (simpeq t)"
+ unfolding simpeq_def
+ using split0_nb0'
+by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] eq_lin)
+
+lemma simpneq_islin[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "islin (simpneq t)"
+ unfolding simpneq_def
+ using split0_nb0'
+by (auto simp add: Let_def split_def isnpoly_fst_split0[OF simptm_allpolys_npoly] islin_stupid allpolys_split0[OF simptm_allpolys_npoly] neq_lin)
+
+lemma really_stupid: "\<not> (\<forall>c1 s'. (c1, s') \<noteq> split0 s)"
+ by (cases "split0 s", auto)
+lemma split0_npoly: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and n: "allpolys isnpoly t"
+ shows "isnpoly (fst (split0 t))" and "allpolys isnpoly (snd (split0 t))"
+ using n
+ by (induct t rule: split0.induct, auto simp add: Let_def split_def polyadd_norm polymul_norm polyneg_norm polysub_norm really_stupid)
+lemma simplt[simp]:
+ shows "Ifm vs bs (simplt t) = Ifm vs bs (Lt t)"
+proof-
+ have n: "allpolys isnpoly (simptm t)" by simp
+ let ?t = "simptm t"
+ {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis
+ using split0[of "simptm t" vs bs] lt[OF split0_npoly(2)[OF n], of vs bs]
+ by (simp add: simplt_def Let_def split_def lt)}
+ moreover
+ {assume "fst (split0 ?t) \<noteq> 0\<^sub>p"
+ hence ?thesis using split0[of "simptm t" vs bs] by (simp add: simplt_def Let_def split_def)
+ }
+ ultimately show ?thesis by blast
+qed
+
+lemma simple[simp]:
+ shows "Ifm vs bs (simple t) = Ifm vs bs (Le t)"
+proof-
+ have n: "allpolys isnpoly (simptm t)" by simp
+ let ?t = "simptm t"
+ {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis
+ using split0[of "simptm t" vs bs] le[OF split0_npoly(2)[OF n], of vs bs]
+ by (simp add: simple_def Let_def split_def le)}
+ moreover
+ {assume "fst (split0 ?t) \<noteq> 0\<^sub>p"
+ hence ?thesis using split0[of "simptm t" vs bs] by (simp add: simple_def Let_def split_def)
+ }
+ ultimately show ?thesis by blast
+qed
+
+lemma simpeq[simp]:
+ shows "Ifm vs bs (simpeq t) = Ifm vs bs (Eq t)"
+proof-
+ have n: "allpolys isnpoly (simptm t)" by simp
+ let ?t = "simptm t"
+ {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis
+ using split0[of "simptm t" vs bs] eq[OF split0_npoly(2)[OF n], of vs bs]
+ by (simp add: simpeq_def Let_def split_def)}
+ moreover
+ {assume "fst (split0 ?t) \<noteq> 0\<^sub>p"
+ hence ?thesis using split0[of "simptm t" vs bs] by (simp add: simpeq_def Let_def split_def)
+ }
+ ultimately show ?thesis by blast
+qed
+
+lemma simpneq[simp]:
+ shows "Ifm vs bs (simpneq t) = Ifm vs bs (NEq t)"
+proof-
+ have n: "allpolys isnpoly (simptm t)" by simp
+ let ?t = "simptm t"
+ {assume "fst (split0 ?t) = 0\<^sub>p" hence ?thesis
+ using split0[of "simptm t" vs bs] neq[OF split0_npoly(2)[OF n], of vs bs]
+ by (simp add: simpneq_def Let_def split_def )}
+ moreover
+ {assume "fst (split0 ?t) \<noteq> 0\<^sub>p"
+ hence ?thesis using split0[of "simptm t" vs bs] by (simp add: simpneq_def Let_def split_def)
+ }
+ ultimately show ?thesis by blast
+qed
+
+lemma lt_nb: "tmbound0 t \<Longrightarrow> bound0 (lt t)"
+ apply (simp add: lt_def)
+ apply (cases t, auto)
+ apply (case_tac poly, auto)
+ done
+
+lemma le_nb: "tmbound0 t \<Longrightarrow> bound0 (le t)"
+ apply (simp add: le_def)
+ apply (cases t, auto)
+ apply (case_tac poly, auto)
+ done
+
+lemma eq_nb: "tmbound0 t \<Longrightarrow> bound0 (eq t)"
+ apply (simp add: eq_def)
+ apply (cases t, auto)
+ apply (case_tac poly, auto)
+ done
+
+lemma neq_nb: "tmbound0 t \<Longrightarrow> bound0 (neq t)"
+ apply (simp add: neq_def eq_def)
+ apply (cases t, auto)
+ apply (case_tac poly, auto)
+ done
+
+lemma simplt_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "tmbound0 t \<Longrightarrow> bound0 (simplt t)"
+ using split0 [of "simptm t" vs bs]
+proof(simp add: simplt_def Let_def split_def)
+ assume nb: "tmbound0 t"
+ hence nb': "tmbound0 (simptm t)" by simp
+ let ?c = "fst (split0 (simptm t))"
+ from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+ have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto
+ from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
+ have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
+ from iffD1[OF isnpolyh_unique[OF ths] th]
+ have "fst (split0 (simptm t)) = 0\<^sub>p" .
+ thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (lt (snd (split0 (simptm t))))) \<and>
+ fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def lt_nb)
+qed
+
+lemma simple_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "tmbound0 t \<Longrightarrow> bound0 (simple t)"
+ using split0 [of "simptm t" vs bs]
+proof(simp add: simple_def Let_def split_def)
+ assume nb: "tmbound0 t"
+ hence nb': "tmbound0 (simptm t)" by simp
+ let ?c = "fst (split0 (simptm t))"
+ from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+ have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto
+ from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
+ have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
+ from iffD1[OF isnpolyh_unique[OF ths] th]
+ have "fst (split0 (simptm t)) = 0\<^sub>p" .
+ thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (le (snd (split0 (simptm t))))) \<and>
+ fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simplt_def Let_def split_def le_nb)
+qed
+
+lemma simpeq_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "tmbound0 t \<Longrightarrow> bound0 (simpeq t)"
+ using split0 [of "simptm t" vs bs]
+proof(simp add: simpeq_def Let_def split_def)
+ assume nb: "tmbound0 t"
+ hence nb': "tmbound0 (simptm t)" by simp
+ let ?c = "fst (split0 (simptm t))"
+ from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+ have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto
+ from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
+ have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
+ from iffD1[OF isnpolyh_unique[OF ths] th]
+ have "fst (split0 (simptm t)) = 0\<^sub>p" .
+ thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (eq (snd (split0 (simptm t))))) \<and>
+ fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpeq_def Let_def split_def eq_nb)
+qed
+
+lemma simpneq_nb[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "tmbound0 t \<Longrightarrow> bound0 (simpneq t)"
+ using split0 [of "simptm t" vs bs]
+proof(simp add: simpneq_def Let_def split_def)
+ assume nb: "tmbound0 t"
+ hence nb': "tmbound0 (simptm t)" by simp
+ let ?c = "fst (split0 (simptm t))"
+ from tmbound_split0[OF nb'[unfolded tmbound0_tmbound_iff[symmetric]]]
+ have th: "\<forall>bs. Ipoly bs ?c = Ipoly bs 0\<^sub>p" by auto
+ from isnpoly_fst_split0[OF simptm_allpolys_npoly[of t]]
+ have ths: "isnpolyh ?c 0" "isnpolyh 0\<^sub>p 0" by (simp_all add: isnpoly_def)
+ from iffD1[OF isnpolyh_unique[OF ths] th]
+ have "fst (split0 (simptm t)) = 0\<^sub>p" .
+ thus "(fst (split0 (simptm t)) = 0\<^sub>p \<longrightarrow> bound0 (neq (snd (split0 (simptm t))))) \<and>
+ fst (split0 (simptm t)) = 0\<^sub>p" by (simp add: simpneq_def Let_def split_def neq_nb)
+qed
+
+consts conjs :: "fm \<Rightarrow> fm list"
+recdef conjs "measure size"
+ "conjs (And p q) = (conjs p)@(conjs q)"
+ "conjs T = []"
+ "conjs p = [p]"
+lemma conjs_ci: "(\<forall> q \<in> set (conjs p). Ifm vs bs q) = Ifm vs bs p"
+by (induct p rule: conjs.induct, auto)
+constdefs list_disj :: "fm list \<Rightarrow> fm"
+ "list_disj ps \<equiv> foldr disj ps F"
+
+lemma list_conj: "Ifm vs bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm vs bs p)"
+ by (induct ps, auto simp add: list_conj_def)
+lemma list_conj_qf: " \<forall>p\<in> set ps. qfree p \<Longrightarrow> qfree (list_conj ps)"
+ by (induct ps, auto simp add: list_conj_def conj_qf)
+lemma list_disj: "Ifm vs bs (list_disj ps) = (\<exists>p\<in> set ps. Ifm vs bs p)"
+ by (induct ps, auto simp add: list_disj_def)
+
+lemma conj_boundslt: "boundslt n p \<Longrightarrow> boundslt n q \<Longrightarrow> boundslt n (conj p q)"
+ unfolding conj_def by auto
+
+lemma conjs_nb: "bound n p \<Longrightarrow> \<forall>q\<in> set (conjs p). bound n q"
+ apply (induct p rule: conjs.induct)
+ apply (unfold conjs.simps)
+ apply (unfold set_append)
+ apply (unfold ball_Un)
+ apply (unfold bound.simps)
+ apply auto
+ done
+
+lemma conjs_boundslt: "boundslt n p \<Longrightarrow> \<forall>q\<in> set (conjs p). boundslt n q"
+ apply (induct p rule: conjs.induct)
+ apply (unfold conjs.simps)
+ apply (unfold set_append)
+ apply (unfold ball_Un)
+ apply (unfold boundslt.simps)
+ apply blast
+by simp_all
+
+lemma list_conj_boundslt: " \<forall>p\<in> set ps. boundslt n p \<Longrightarrow> boundslt n (list_conj ps)"
+ unfolding list_conj_def
+ by (induct ps, auto simp add: conj_boundslt)
+
+lemma list_conj_nb: assumes bnd: "\<forall>p\<in> set ps. bound n p"
+ shows "bound n (list_conj ps)"
+ using bnd
+ unfolding list_conj_def
+ by (induct ps, auto simp add: conj_nb)
+
+lemma list_conj_nb': "\<forall>p\<in>set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
+unfolding list_conj_def by (induct ps , auto)
+
+lemma CJNB_qe:
+ assumes qe: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
+ shows "\<forall> bs p. qfree p \<longrightarrow> qfree (CJNB qe p) \<and> (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 "\<forall> q\<in> set ?yes. bound0 q" by blast
+ hence yes_nb: "bound0 ?cyes" by (simp add: list_conj_nb')
+ hence yes_qf: "qfree (decr0 ?cyes )" by (simp add: decr0_qf)
+ from conjuncts_qf[OF qfp] partition_set[OF part]
+ have " \<forall>q\<in> set ?no. qfree q" by auto
+ hence no_qf: "qfree ?cno"by (simp add: list_conj_qf)
+ with qe have cno_qf:"qfree (qe ?cno )"
+ and noE: "Ifm vs bs (qe ?cno) = Ifm vs bs (E ?cno)" by blast+
+ from cno_qf yes_qf have qf: "qfree (CJNB qe p)"
+ by (simp add: CJNB_def Let_def conj_qf split_def)
+ {fix bs
+ from conjuncts have "Ifm vs bs p = (\<forall>q\<in> set ?cjs. Ifm vs bs q)" by blast
+ also have "\<dots> = ((\<forall>q\<in> set ?yes. Ifm vs bs q) \<and> (\<forall>q\<in> set ?no. Ifm vs bs q))"
+ using partition_set[OF part] by auto
+ finally have "Ifm vs bs p = ((Ifm vs bs ?cyes) \<and> (Ifm vs bs ?cno))" using list_conj[of vs bs] by simp}
+ hence "Ifm vs bs (E p) = (\<exists>x. (Ifm vs (x#bs) ?cyes) \<and> (Ifm vs (x#bs) ?cno))" by simp
+ also have "\<dots> = (\<exists>x. (Ifm vs (y#bs) ?cyes) \<and> (Ifm vs (x#bs) ?cno))"
+ using bound0_I[OF yes_nb, where bs="bs" and b'="y"] by blast
+ also have "\<dots> = (Ifm vs bs (decr0 ?cyes) \<and> Ifm vs bs (E ?cno))"
+ by (auto simp add: decr0[OF yes_nb])
+ also have "\<dots> = (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) \<and> Ifm vs bs (CJNB qe p) = Ifm vs bs (E p)" by blast
+qed
+
+consts simpfm :: "fm \<Rightarrow> fm"
+recdef simpfm "measure fmsize"
+ "simpfm (Lt t) = simplt (simptm t)"
+ "simpfm (Le t) = simple (simptm t)"
+ "simpfm (Eq t) = simpeq(simptm t)"
+ "simpfm (NEq t) = simpneq(simptm t)"
+ "simpfm (And p q) = conj (simpfm p) (simpfm q)"
+ "simpfm (Or p q) = disj (simpfm p) (simpfm q)"
+ "simpfm (Imp p q) = disj (simpfm (NOT p)) (simpfm q)"
+ "simpfm (Iff p q) = disj (conj (simpfm p) (simpfm q)) (conj (simpfm (NOT p)) (simpfm (NOT q)))"
+ "simpfm (NOT (And p q)) = disj (simpfm (NOT p)) (simpfm (NOT q))"
+ "simpfm (NOT (Or p q)) = conj (simpfm (NOT p)) (simpfm (NOT q))"
+ "simpfm (NOT (Imp p q)) = conj (simpfm p) (simpfm (NOT q))"
+ "simpfm (NOT (Iff p q)) = disj (conj (simpfm p) (simpfm (NOT q))) (conj (simpfm (NOT p)) (simpfm q))"
+ "simpfm (NOT (Eq t)) = simpneq t"
+ "simpfm (NOT (NEq t)) = simpeq t"
+ "simpfm (NOT (Le t)) = simplt (Neg t)"
+ "simpfm (NOT (Lt t)) = simple (Neg t)"
+ "simpfm (NOT (NOT p)) = simpfm p"
+ "simpfm (NOT T) = F"
+ "simpfm (NOT F) = T"
+ "simpfm p = p"
+
+lemma simpfm[simp]: "Ifm vs bs (simpfm p) = Ifm vs bs p"
+by(induct p arbitrary: bs rule: simpfm.induct, auto)
+
+lemma simpfm_bound0: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "bound0 p \<Longrightarrow> bound0 (simpfm p)"
+by (induct p rule: simpfm.induct, auto)
+
+lemma lt_qf[simp]: "qfree (lt t)"
+ apply (cases t, auto simp add: lt_def)
+ by (case_tac poly, auto)
+
+lemma le_qf[simp]: "qfree (le t)"
+ apply (cases t, auto simp add: le_def)
+ by (case_tac poly, auto)
+
+lemma eq_qf[simp]: "qfree (eq t)"
+ apply (cases t, auto simp add: eq_def)
+ by (case_tac poly, auto)
+
+lemma neq_qf[simp]: "qfree (neq t)" by (simp add: neq_def)
+
+lemma simplt_qf[simp]: "qfree (simplt t)" by (simp add: simplt_def Let_def split_def)
+lemma simple_qf[simp]: "qfree (simple t)" by (simp add: simple_def Let_def split_def)
+lemma simpeq_qf[simp]: "qfree (simpeq t)" by (simp add: simpeq_def Let_def split_def)
+lemma simpneq_qf[simp]: "qfree (simpneq t)" by (simp add: simpneq_def Let_def split_def)
+
+lemma simpfm_qf[simp]: "qfree p \<Longrightarrow> qfree (simpfm p)"
+by (induct p rule: simpfm.induct, auto simp add: disj_qf imp_qf iff_qf conj_qf not_qf Let_def)
+
+lemma disj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (disj p q)" by (simp add: disj_def)
+lemma conj_lin: "islin p \<Longrightarrow> islin q \<Longrightarrow> islin (conj p q)" by (simp add: conj_def)
+
+lemma assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "qfree p \<Longrightarrow> islin (simpfm p)"
+ apply (induct p rule: simpfm.induct)
+ apply (simp_all add: conj_lin disj_lin)
+ done
+
+consts prep :: "fm \<Rightarrow> fm"
+recdef prep "measure fmsize"
+ "prep (E T) = T"
+ "prep (E F) = F"
+ "prep (E (Or p q)) = disj (prep (E p)) (prep (E q))"
+ "prep (E (Imp p q)) = disj (prep (E (NOT p))) (prep (E q))"
+ "prep (E (Iff p q)) = disj (prep (E (And p q))) (prep (E (And (NOT p) (NOT q))))"
+ "prep (E (NOT (And p q))) = disj (prep (E (NOT p))) (prep (E(NOT q)))"
+ "prep (E (NOT (Imp p q))) = prep (E (And p (NOT q)))"
+ "prep (E (NOT (Iff p q))) = disj (prep (E (And p (NOT q)))) (prep (E(And (NOT p) q)))"
+ "prep (E p) = E (prep p)"
+ "prep (A (And p q)) = conj (prep (A p)) (prep (A q))"
+ "prep (A p) = prep (NOT (E (NOT p)))"
+ "prep (NOT (NOT p)) = prep p"
+ "prep (NOT (And p q)) = disj (prep (NOT p)) (prep (NOT q))"
+ "prep (NOT (A p)) = prep (E (NOT p))"
+ "prep (NOT (Or p q)) = conj (prep (NOT p)) (prep (NOT q))"
+ "prep (NOT (Imp p q)) = conj (prep p) (prep (NOT q))"
+ "prep (NOT (Iff p q)) = disj (prep (And p (NOT q))) (prep (And (NOT p) q))"
+ "prep (NOT p) = not (prep p)"
+ "prep (Or p q) = disj (prep p) (prep q)"
+ "prep (And p q) = conj (prep p) (prep q)"
+ "prep (Imp p q) = prep (Or (NOT p) q)"
+ "prep (Iff p q) = disj (prep (And p q)) (prep (And (NOT p) (NOT q)))"
+ "prep p = p"
+(hints simp add: fmsize_pos)
+lemma prep: "Ifm vs bs (prep p) = Ifm vs bs p"
+by (induct p arbitrary: bs rule: prep.induct, auto)
+
+
+
+ (* Generic quantifier elimination *)
+consts qelim :: "fm \<Rightarrow> (fm \<Rightarrow> fm) \<Rightarrow> fm"
+recdef qelim "measure fmsize"
+ "qelim (E p) = (\<lambda> qe. DJ (CJNB qe) (qelim p qe))"
+ "qelim (A p) = (\<lambda> qe. not (qe ((qelim (NOT p) qe))))"
+ "qelim (NOT p) = (\<lambda> qe. not (qelim p qe))"
+ "qelim (And p q) = (\<lambda> qe. conj (qelim p qe) (qelim q qe))"
+ "qelim (Or p q) = (\<lambda> qe. disj (qelim p qe) (qelim q qe))"
+ "qelim (Imp p q) = (\<lambda> qe. imp (qelim p qe) (qelim q qe))"
+ "qelim (Iff p q) = (\<lambda> qe. iff (qelim p qe) (qelim q qe))"
+ "qelim p = (\<lambda> y. simpfm p)"
+
+
+lemma qelim:
+ assumes qe_inv: "\<forall> bs p. qfree p \<longrightarrow> qfree (qe p) \<and> (Ifm vs bs (qe p) = Ifm vs bs (E p))"
+ shows "\<And> bs. qfree (qelim p qe) \<and> (Ifm vs bs (qelim p qe) = Ifm vs bs p)"
+using qe_inv DJ_qe[OF CJNB_qe[OF qe_inv]]
+by (induct p rule: qelim.induct) auto
+
+subsection{* Core Procedure *}
+
+consts
+ plusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of +\<infinity>*)
+ minusinf:: "fm \<Rightarrow> fm" (* Virtual substitution of -\<infinity>*)
+recdef minusinf "measure size"
+ "minusinf (And p q) = conj (minusinf p) (minusinf q)"
+ "minusinf (Or p q) = disj (minusinf p) (minusinf q)"
+ "minusinf (Eq (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
+ "minusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
+ "minusinf (Lt (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP (~\<^sub>p c)))"
+ "minusinf (Le (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP (~\<^sub>p c)))"
+ "minusinf p = p"
+
+recdef plusinf "measure size"
+ "plusinf (And p q) = conj (plusinf p) (plusinf q)"
+ "plusinf (Or p q) = disj (plusinf p) (plusinf q)"
+ "plusinf (Eq (CNP 0 c e)) = conj (eq (CP c)) (eq e)"
+ "plusinf (NEq (CNP 0 c e)) = disj (not (eq e)) (not (eq (CP c)))"
+ "plusinf (Lt (CNP 0 c e)) = disj (conj (eq (CP c)) (lt e)) (lt (CP c))"
+ "plusinf (Le (CNP 0 c e)) = disj (conj (eq (CP c)) (le e)) (lt (CP c))"
+ "plusinf p = p"
+
+lemma minusinf_inf: assumes lp:"islin p"
+ shows "\<exists>z. \<forall>x < z. Ifm vs (x#bs) (minusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
+ using lp
+proof (induct p rule: minusinf.induct)
+ case 1 thus ?case by (auto,rule_tac x="min z za" in exI, auto)
+next
+ case 2 thus ?case by (auto,rule_tac x="min z za" in exI, auto)
+next
+ case (3 c e) hence nbe: "tmbound0 e" by simp
+ from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+ note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
+ let ?c = "Ipoly vs c"
+ let ?e = "Itm vs (y#bs) e"
+ have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+ moreover {assume "?c = 0" hence ?case
+ using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto}
+ moreover {assume cp: "?c > 0"
+ {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
+ using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e < 0" by simp
+ hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))"
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto}
+ moreover {assume cp: "?c < 0"
+ {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
+ using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e > 0" by simp
+ hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Eq (CNP 0 c e)))"
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto}
+ ultimately show ?case by blast
+next
+ case (4 c e) hence nbe: "tmbound0 e" by simp
+ from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+ note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
+ let ?c = "Ipoly vs c"
+ let ?e = "Itm vs (y#bs) e"
+ have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+ moreover {assume "?c = 0" hence ?case using eqs by auto}
+ moreover {assume cp: "?c > 0"
+ {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
+ using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e < 0" by simp
+ hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))"
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+ moreover {assume cp: "?c < 0"
+ {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
+ using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e > 0" by simp
+ hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (NEq (CNP 0 c e)))"
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+ ultimately show ?case by blast
+next
+ case (5 c e) hence nbe: "tmbound0 e" by simp
+ from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+ hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm)
+ note eqs = lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] lt[OF nc(2), where ?'a = 'a]
+ let ?c = "Ipoly vs c"
+ let ?e = "Itm vs (y#bs) e"
+ have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+ moreover {assume "?c = 0" hence ?case using eqs by auto}
+ moreover {assume cp: "?c > 0"
+ {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
+ using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e < 0" by simp
+ hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))"
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ moreover {assume cp: "?c < 0"
+ {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
+ using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e > 0" by simp
+ hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Lt (CNP 0 c e)))"
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto}
+ ultimately show ?case by blast
+next
+ case (6 c e) hence nbe: "tmbound0 e" by simp
+ from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+ hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm)
+ note eqs = lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] le[OF nc(2), where ?'a = 'a]
+ let ?c = "Ipoly vs c"
+ let ?e = "Itm vs (y#bs) e"
+ have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+ moreover {assume "?c = 0" hence ?case using eqs by auto}
+ moreover {assume cp: "?c > 0"
+ {fix x assume xz: "x < -?e / ?c" hence "?c * x < - ?e"
+ using pos_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e < 0" by simp
+ hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))"
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ moreover {assume cp: "?c < 0"
+ {fix x assume xz: "x < -?e / ?c" hence "?c * x > - ?e"
+ using neg_less_divide_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e > 0" by simp
+ hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (minusinf (Le (CNP 0 c e)))"
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ ultimately show ?case by blast
+qed (auto)
+
+lemma plusinf_inf: assumes lp:"islin p"
+ shows "\<exists>z. \<forall>x > z. Ifm vs (x#bs) (plusinf p) \<longleftrightarrow> Ifm vs (x#bs) p"
+ using lp
+proof (induct p rule: plusinf.induct)
+ case 1 thus ?case by (auto,rule_tac x="max z za" in exI, auto)
+next
+ case 2 thus ?case by (auto,rule_tac x="max z za" in exI, auto)
+next
+ case (3 c e) hence nbe: "tmbound0 e" by simp
+ from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+ note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
+ let ?c = "Ipoly vs c"
+ let ?e = "Itm vs (y#bs) e"
+ have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+ moreover {assume "?c = 0" hence ?case
+ using eq[OF nc(2), of vs] eq[OF nc(1), of vs] by auto}
+ moreover {assume cp: "?c > 0"
+ {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
+ using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e > 0" by simp
+ hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x" and vs=vs and bs=bs] by auto} hence ?case by auto}
+ moreover {assume cp: "?c < 0"
+ {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
+ using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e < 0" by simp
+ hence "Ifm vs (x#bs) (Eq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Eq (CNP 0 c e)))"
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] eqs by auto} hence ?case by auto}
+ ultimately show ?case by blast
+next
+ case (4 c e) hence nbe: "tmbound0 e" by simp
+ from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+ note eqs = eq[OF nc(1), where ?'a = 'a] eq[OF nc(2), where ?'a = 'a]
+ let ?c = "Ipoly vs c"
+ let ?e = "Itm vs (y#bs) e"
+ have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+ moreover {assume "?c = 0" hence ?case using eqs by auto}
+ moreover {assume cp: "?c > 0"
+ {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
+ using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e > 0" by simp
+ hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))"
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+ moreover {assume cp: "?c < 0"
+ {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
+ using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e < 0" by simp
+ hence "Ifm vs (x#bs) (NEq (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (NEq (CNP 0 c e)))"
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] by auto} hence ?case by auto}
+ ultimately show ?case by blast
+next
+ case (5 c e) hence nbe: "tmbound0 e" by simp
+ from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+ hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm)
+ note eqs = lt[OF nc(1), where ?'a = 'a] lt[OF nc', where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] lt[OF nc(2), where ?'a = 'a]
+ let ?c = "Ipoly vs c"
+ let ?e = "Itm vs (y#bs) e"
+ have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+ moreover {assume "?c = 0" hence ?case using eqs by auto}
+ moreover {assume cp: "?c > 0"
+ {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
+ using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e > 0" by simp
+ hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))"
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ moreover {assume cp: "?c < 0"
+ {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
+ using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e < 0" by simp
+ hence "Ifm vs (x#bs) (Lt (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Lt (CNP 0 c e)))"
+ using eqs tmbound0_I[OF nbe, where b="y" and b'="x"] cp by auto} hence ?case by auto}
+ ultimately show ?case by blast
+next
+ case (6 c e) hence nbe: "tmbound0 e" by simp
+ from prems have nc: "allpolys isnpoly (CP c)" "allpolys isnpoly e" by simp_all
+ hence nc': "allpolys isnpoly (CP (~\<^sub>p c))" by (simp add: polyneg_norm)
+ note eqs = lt[OF nc(1), where ?'a = 'a] eq [OF nc(1), where ?'a = 'a] le[OF nc(2), where ?'a = 'a]
+ let ?c = "Ipoly vs c"
+ let ?e = "Itm vs (y#bs) e"
+ have "?c=0 \<or> ?c > 0 \<or> ?c < 0" by arith
+ moreover {assume "?c = 0" hence ?case using eqs by auto}
+ moreover {assume cp: "?c > 0"
+ {fix x assume xz: "x > -?e / ?c" hence "?c * x > - ?e"
+ using pos_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e > 0" by simp
+ hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))"
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ moreover {assume cp: "?c < 0"
+ {fix x assume xz: "x > -?e / ?c" hence "?c * x < - ?e"
+ using neg_divide_less_eq[OF cp, where a="x" and b="-?e"] by (simp add: mult_commute)
+ hence "?c * x + ?e < 0" by simp
+ hence "Ifm vs (x#bs) (Le (CNP 0 c e)) = Ifm vs (x#bs) (plusinf (Le (CNP 0 c e)))"
+ using tmbound0_I[OF nbe, where b="y" and b'="x"] cp eqs by auto} hence ?case by auto}
+ ultimately show ?case by blast
+qed (auto)
+
+lemma minusinf_nb: "islin p \<Longrightarrow> bound0 (minusinf p)"
+ by (induct p rule: minusinf.induct, auto simp add: eq_nb lt_nb le_nb)
+lemma plusinf_nb: "islin p \<Longrightarrow> bound0 (plusinf p)"
+ by (induct p rule: minusinf.induct, auto simp add: eq_nb lt_nb le_nb)
+
+lemma minusinf_ex: assumes lp: "islin p" and ex: "Ifm vs (x#bs) (minusinf p)"
+ shows "\<exists>x. Ifm vs (x#bs) p"
+proof-
+ from bound0_I [OF minusinf_nb[OF lp], where b="a" and bs ="bs"] ex
+ have th: "\<forall> x. Ifm vs (x#bs) (minusinf p)" by auto
+ from minusinf_inf[OF lp, where bs="bs"]
+ obtain z where z_def: "\<forall>x<z. Ifm vs (x # bs) (minusinf p) = Ifm vs (x # bs) p" by blast
+ from th have "Ifm vs ((z - 1)#bs) (minusinf p)" by simp
+ moreover have "z - 1 < z" by simp
+ ultimately show ?thesis using z_def by auto
+qed
+
+lemma plusinf_ex: assumes lp: "islin p" and ex: "Ifm vs (x#bs) (plusinf p)"
+ shows "\<exists>x. Ifm vs (x#bs) p"
+proof-
+ from bound0_I [OF plusinf_nb[OF lp], where b="a" and bs ="bs"] ex
+ have th: "\<forall> x. Ifm vs (x#bs) (plusinf p)" by auto
+ from plusinf_inf[OF lp, where bs="bs"]
+ obtain z where z_def: "\<forall>x>z. Ifm vs (x # bs) (plusinf p) = Ifm vs (x # bs) p" by blast
+ from th have "Ifm vs ((z + 1)#bs) (plusinf p)" by simp
+ moreover have "z + 1 > z" by simp
+ ultimately show ?thesis using z_def by auto
+qed
+
+fun uset :: "fm \<Rightarrow> (poly \<times> 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 "\<forall> (c,s) \<in> set (uset p). isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
+using lp by(induct p rule: uset.induct,auto)
+
+lemma minusinf_uset0:
+ assumes lp: "islin p"
+ and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))"
+ and ex: "Ifm vs (x#bs) p" (is "?I x p")
+ shows "\<exists> (c,s) \<in> set (uset p). x \<ge> - Itm vs (x#bs) s / Ipoly vs c"
+proof-
+ have "\<exists> (c,s) \<in> set (uset p). (Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or> (Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s)"
+ using lp nmi ex
+ apply (induct p rule: minusinf.induct, auto simp add: eq le lt nth_pos2 polyneg_norm)
+ apply (auto simp add: linorder_not_less order_le_less)
+ done
+ then obtain c s where csU: "(c,s) \<in> set (uset p)" and x: "(Ipoly vs c < 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s) \<or> (Ipoly vs c > 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s)" by blast
+ hence "x \<ge> (- Itm vs (x#bs) s) / Ipoly vs c"
+ using divide_le_eq[of "- Itm vs (x#bs) s" "Ipoly vs c" x]
+ by (auto simp add: mult_commute del: divide_minus_left)
+ thus ?thesis using csU by auto
+qed
+
+lemma minusinf_uset:
+ assumes lp: "islin p"
+ and nmi: "\<not> (Ifm vs (a#bs) (minusinf p))"
+ and ex: "Ifm vs (x#bs) p" (is "?I x p")
+ shows "\<exists> (c,s) \<in> set (uset p). x \<ge> - Itm vs (a#bs) s / Ipoly vs c"
+proof-
+ from nmi have nmi': "\<not> (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) \<in> set (uset p)" and th: "x \<ge> - 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: "\<not> (Ifm vs (x#bs) (plusinf p))"
+ and ex: "Ifm vs (x#bs) p" (is "?I x p")
+ shows "\<exists> (c,s) \<in> set (uset p). x \<le> - Itm vs (x#bs) s / Ipoly vs c"
+proof-
+ have "\<exists> (c,s) \<in> set (uset p). (Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s) \<or> (Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s)"
+ using lp nmi ex
+ apply (induct p rule: minusinf.induct, auto simp add: eq le lt nth_pos2 polyneg_norm)
+ apply (auto simp add: linorder_not_less order_le_less)
+ done
+ then obtain c s where csU: "(c,s) \<in> set (uset p)" and x: "(Ipoly vs c < 0 \<and> Ipoly vs c * x \<ge> - Itm vs (x#bs) s) \<or> (Ipoly vs c > 0 \<and> Ipoly vs c * x \<le> - Itm vs (x#bs) s)" by blast
+ hence "x \<le> (- Itm vs (x#bs) s) / Ipoly vs c"
+ using le_divide_eq[of x "- Itm vs (x#bs) s" "Ipoly vs c"]
+ by (auto simp add: mult_commute del: divide_minus_left)
+ thus ?thesis using csU by auto
+qed
+
+lemma plusinf_uset:
+ assumes lp: "islin p"
+ and nmi: "\<not> (Ifm vs (a#bs) (plusinf p))"
+ and ex: "Ifm vs (x#bs) p" (is "?I x p")
+ shows "\<exists> (c,s) \<in> set (uset p). x \<le> - Itm vs (a#bs) s / Ipoly vs c"
+proof-
+ from nmi have nmi': "\<not> (Ifm vs (x#bs) (plusinf p))"
+ by (simp add: bound0_I[OF plusinf_nb[OF lp], where b=x and b'=a])
+ from plusinf_uset0[OF lp nmi' ex]
+ obtain c s where csU: "(c,s) \<in> set (uset p)" and th: "x \<le> - Itm vs (x#bs) s / Ipoly vs c" by blast
+ from uset_l[OF lp, rule_format, OF csU] have nb: "tmbound0 s" by simp
+ from th tmbound0_I[OF nb, of vs x bs a] csU show ?thesis by auto
+qed
+
+lemma lin_dense:
+ assumes lp: "islin p"
+ and noS: "\<forall> t. l < t \<and> t< u \<longrightarrow> t \<notin> (\<lambda> (c,t). - Itm vs (x#bs) t / Ipoly vs c) ` set (uset p)"
+ (is "\<forall> t. _ \<and> _ \<longrightarrow> t \<notin> (\<lambda> (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 \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
+ and px: "Ifm vs (x # bs) (Lt (CNP 0 c s))"
+ and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
+ from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp
+ hence ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c" by auto
+ have ccs: "?N c = 0 \<or> ?N c < 0 \<or> ?N c > 0" by dlo
+ moreover
+ {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
+ moreover
+ {assume c: "?N c > 0"
+ from px pos_less_divide_eq[OF c, where a="x" and b="-?Nt x s"]
+ have px': "x < - ?Nt x s / ?N c"
+ by (auto simp add: not_less ring_simps)
+ {assume y: "y < - ?Nt x s / ?N c"
+ hence "y * ?N c < - ?Nt x s"
+ by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+ hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+ hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
+ moreover
+ {assume y: "y > -?Nt x s / ?N c"
+ with yu have eu: "u > - ?Nt x s / ?N c" by auto
+ with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
+ with lx px' have "False" by simp hence ?case by simp }
+ ultimately have ?case using ycs by blast
+ }
+ moreover
+ {assume c: "?N c < 0"
+ from px neg_divide_less_eq[OF c, where a="x" and b="-?Nt x s"]
+ have px': "x > - ?Nt x s / ?N c"
+ by (auto simp add: not_less ring_simps)
+ {assume y: "y > - ?Nt x s / ?N c"
+ hence "y * ?N c < - ?Nt x s"
+ by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+ hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+ hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
+ moreover
+ {assume y: "y < -?Nt x s / ?N c"
+ with ly have eu: "l < - ?Nt x s / ?N c" by auto
+ with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
+ with xu px' have "False" by simp hence ?case by simp }
+ ultimately have ?case using ycs by blast
+ }
+ ultimately show ?case by blast
+next
+ case (6 c s)
+ from "6.prems"
+ have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
+ and px: "Ifm vs (x # bs) (Le (CNP 0 c s))"
+ and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
+ from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp
+ hence ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c" by auto
+ have ccs: "?N c = 0 \<or> ?N c < 0 \<or> ?N c > 0" by dlo
+ moreover
+ {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
+ moreover
+ {assume c: "?N c > 0"
+ from px pos_le_divide_eq[OF c, where a="x" and b="-?Nt x s"]
+ have px': "x <= - ?Nt x s / ?N c" by (simp add: not_less ring_simps)
+ {assume y: "y < - ?Nt x s / ?N c"
+ hence "y * ?N c < - ?Nt x s"
+ by (simp add: pos_less_divide_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+ hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+ hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
+ moreover
+ {assume y: "y > -?Nt x s / ?N c"
+ with yu have eu: "u > - ?Nt x s / ?N c" by auto
+ with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
+ with lx px' have "False" by simp hence ?case by simp }
+ ultimately have ?case using ycs by blast
+ }
+ moreover
+ {assume c: "?N c < 0"
+ from px neg_divide_le_eq[OF c, where a="x" and b="-?Nt x s"]
+ have px': "x >= - ?Nt x s / ?N c" by (simp add: ring_simps)
+ {assume y: "y > - ?Nt x s / ?N c"
+ hence "y * ?N c < - ?Nt x s"
+ by (simp add: neg_divide_less_eq[OF c, where a="y" and b="-?Nt x s", symmetric])
+ hence "?N c * y + ?Nt x s < 0" by (simp add: ring_simps)
+ hence ?case using tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"] by simp}
+ moreover
+ {assume y: "y < -?Nt x s / ?N c"
+ with ly have eu: "l < - ?Nt x s / ?N c" by auto
+ with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
+ with xu px' have "False" by simp hence ?case by simp }
+ ultimately have ?case using ycs by blast
+ }
+ ultimately show ?case by blast
+next
+ case (3 c s)
+ from "3.prems"
+ have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
+ and px: "Ifm vs (x # bs) (Eq (CNP 0 c s))"
+ and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
+ from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp
+ hence ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c" by auto
+ have ccs: "?N c = 0 \<or> ?N c < 0 \<or> ?N c > 0" by dlo
+ moreover
+ {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
+ moreover
+ {assume c: "?N c > 0" hence cnz: "?N c \<noteq> 0" by simp
+ from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz
+ have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps)
+ {assume y: "y < -?Nt x s / ?N c"
+ with ly have eu: "l < - ?Nt x s / ?N c" by auto
+ with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
+ with xu px' have "False" by simp hence ?case by simp }
+ moreover
+ {assume y: "y > -?Nt x s / ?N c"
+ with yu have eu: "u > - ?Nt x s / ?N c" by auto
+ with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
+ with lx px' have "False" by simp hence ?case by simp }
+ ultimately have ?case using ycs by blast
+ }
+ moreover
+ {assume c: "?N c < 0" hence cnz: "?N c \<noteq> 0" by simp
+ from px eq_divide_eq[of "x" "-?Nt x s" "?N c"] cnz
+ have px': "x = - ?Nt x s / ?N c" by (simp add: ring_simps)
+ {assume y: "y < -?Nt x s / ?N c"
+ with ly have eu: "l < - ?Nt x s / ?N c" by auto
+ with noS ly yu have th: "- ?Nt x s / ?N c \<ge> u" by (cases "- ?Nt x s / ?N c < u", auto)
+ with xu px' have "False" by simp hence ?case by simp }
+ moreover
+ {assume y: "y > -?Nt x s / ?N c"
+ with yu have eu: "u > - ?Nt x s / ?N c" by auto
+ with noS ly yu have th: "- ?Nt x s / ?N c \<le> l" by (cases "- ?Nt x s / ?N c > l", auto)
+ with lx px' have "False" by simp hence ?case by simp }
+ ultimately have ?case using ycs by blast
+ }
+ ultimately show ?case by blast
+next
+ case (4 c s)
+ from "4.prems"
+ have lin: "isnpoly c" "c \<noteq> 0\<^sub>p" "tmbound0 s" "allpolys isnpoly s"
+ and px: "Ifm vs (x # bs) (NEq (CNP 0 c s))"
+ and noS: "\<forall>t. l < t \<and> t < u \<longrightarrow> t \<noteq> - Itm vs (x # bs) s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp_all
+ from ly yu noS have yne: "y \<noteq> - ?Nt x s / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" by simp
+ hence ycs: "y < - ?Nt x s / ?N c \<or> y > -?Nt x s / ?N c" by auto
+ have ccs: "?N c = 0 \<or> ?N c \<noteq> 0" by dlo
+ moreover
+ {assume "?N c = 0" hence ?case using px by (simp add: tmbound0_I[OF lin(3), where bs="bs" and b="x" and b'="y"])}
+ moreover
+ {assume c: "?N c \<noteq> 0"
+ from yne c eq_divide_eq[of "y" "- ?Nt x s" "?N c"] have ?case
+ by (simp add: ring_simps tmbound0_I[OF lin(3), of vs x bs y] sum_eq[symmetric]) }
+ ultimately show ?case by blast
+qed (auto simp add: nth_pos2 tmbound0_I[where vs=vs and bs="bs" and b="y" and b'="x"] bound0_I[where vs=vs and bs="bs" and b="y" and b'="x"])
+
+lemma one_plus_one_pos[simp]: "(1::'a::{ordered_field}) + 1 > 0"
+proof-
+ have op: "(1::'a) > 0" by simp
+ from add_pos_pos[OF op op] show ?thesis .
+qed
+
+lemma one_plus_one_nonzero[simp]: "(1::'a::{ordered_field}) + 1 \<noteq> 0"
+ using one_plus_one_pos[where ?'a = 'a] by (simp add: less_le)
+
+lemma half_sum_eq: "(u + u) / (1+1) = (u::'a::{ordered_field})"
+proof-
+ have "(u + u) = (1 + 1) * u" by (simp add: ring_simps)
+ hence "(u + u) / (1+1) = (1 + 1)*u / (1 + 1)" by simp
+ with nonzero_mult_divide_cancel_left[OF one_plus_one_nonzero, of u] show ?thesis by simp
+qed
+
+lemma inf_uset:
+ assumes lp: "islin p"
+ and nmi: "\<not> (Ifm vs (x#bs) (minusinf p))" (is "\<not> (Ifm vs (x#bs) (?M p))")
+ and npi: "\<not> (Ifm vs (x#bs) (plusinf p))" (is "\<not> (Ifm vs (x#bs) (?P p))")
+ and ex: "\<exists> x. Ifm vs (x#bs) p" (is "\<exists> x. ?I x p")
+ shows "\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). ?I ((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) / (1 + 1)) p"
+proof-
+ let ?Nt = "\<lambda> 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': "\<not> (?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': "\<not> (?I a (?P p))" by simp
+ have "\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). ?I ((- ?Nt a t/?N c + - ?Nt a s /?N d) / (1 + 1)) p"
+ proof-
+ let ?M = "(\<lambda> (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 "\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). a \<le> - ?Nt x t / ?N c \<and> a \<ge> - ?Nt x s / ?N d" by blast
+ then obtain "c" "t" "d" "s" where
+ ctU: "(c,t) \<in> ?U" and dsU: "(d,s) \<in> ?U"
+ and xs1: "a \<le> - ?Nt x s / ?N d" and tx1: "a \<ge> - ?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 \<le> - ?Nt a s / ?N d" and tx: "a \<ge> - ?Nt a t / ?N c" by auto
+ from ctU have Mne: "?M \<noteq> {}" by auto
+ hence Une: "?U \<noteq> {}" by simp
+ let ?l = "Min ?M"
+ let ?u = "Max ?M"
+ have linM: "?l \<in> ?M" using fM Mne by simp
+ have uinM: "?u \<in> ?M" using fM Mne by simp
+ have ctM: "- ?Nt a t / ?N c \<in> ?M" using ctU by auto
+ have dsM: "- ?Nt a s / ?N d \<in> ?M" using dsU by auto
+ have lM: "\<forall> t\<in> ?M. ?l \<le> t" using Mne fM by auto
+ have Mu: "\<forall> t\<in> ?M. t \<le> ?u" using Mne fM by auto
+ have "?l \<le> - ?Nt a t / ?N c" using ctM Mne by simp hence lx: "?l \<le> a" using tx by simp
+ have "- ?Nt a s / ?N d \<le> ?u" using dsM Mne by simp hence xu: "a \<le> ?u" using xs by simp
+ from finite_set_intervals2[where P="\<lambda> x. ?I x p",OF pa lx xu linM uinM fM lM Mu]
+ have "(\<exists> s\<in> ?M. ?I s p) \<or>
+ (\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p)" .
+ moreover {fix u assume um: "u\<in> ?M" and pu: "?I u p"
+ hence "\<exists> (nu,tu) \<in> ?U. u = - ?Nt a tu / ?N nu" by auto
+ then obtain "tu" "nu" where tuU: "(nu,tu) \<in> ?U" and tuu:"u= - ?Nt a tu / ?N nu" by blast
+ from half_sum_eq[of u] pu tuu
+ have "?I (((- ?Nt a tu / ?N nu) + (- ?Nt a tu / ?N nu)) / (1 + 1)) p" by simp
+ with tuU have ?thesis by blast}
+ moreover{
+ assume "\<exists> t1\<in> ?M. \<exists> t2 \<in> ?M. (\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M) \<and> t1 < a \<and> a < t2 \<and> ?I a p"
+ then obtain t1 and t2 where t1M: "t1 \<in> ?M" and t2M: "t2\<in> ?M"
+ and noM: "\<forall> y. t1 < y \<and> y < t2 \<longrightarrow> y \<notin> ?M" and t1x: "t1 < a" and xt2: "a < t2" and px: "?I a p"
+ by blast
+ from t1M have "\<exists> (t1n,t1u) \<in> ?U. t1 = - ?Nt a t1u / ?N t1n" by auto
+ then obtain "t1u" "t1n" where t1uU: "(t1n,t1u) \<in> ?U" and t1u: "t1 = - ?Nt a t1u / ?N t1n" by blast
+ from t2M have "\<exists> (t2n,t2u) \<in> ?U. t2 = - ?Nt a t2u / ?N t2n" by auto
+ then obtain "t2u" "t2n" where t2uU: "(t2n,t2u) \<in> ?U" and t2u: "t2 = - ?Nt a t2u / ?N t2n" by blast
+ from t1x xt2 have t1t2: "t1 < t2" by simp
+ let ?u = "(t1 + t2) / (1 + 1)"
+ from less_half_sum[OF t1t2] gt_half_sum[OF t1t2] have t1lu: "t1 < ?u" and ut2: "?u < t2" by auto
+ from lin_dense[OF lp noM t1x xt2 px t1lu ut2] have "?I ?u p" .
+ with t1uU t2uU t1u t2u have ?thesis by blast}
+ ultimately show ?thesis by blast
+ qed
+ then obtain "l" "n" "s" "m" where lnU: "(n,l) \<in> ?U" and smU:"(m,s) \<in> ?U"
+ and pu: "?I ((- ?Nt a l / ?N n + - ?Nt a s / ?N m) / (1 + 1)) p" by blast
+ from lnU smU uset_l[OF lp] have nbl: "tmbound0 l" and nbs: "tmbound0 s" by auto
+ from tmbound0_I[OF nbl, where bs="bs" and b="a" and b'="x"]
+ tmbound0_I[OF nbs, where bs="bs" and b="a" and b'="x"] pu
+ have "?I ((- ?Nt x l / ?N n + - ?Nt x s / ?N m) / (1 + 1)) p" by simp
+ with lnU smU
+ show ?thesis by auto
+qed
+
+ (* The Ferrante - Rackoff Theorem *)
+
+theorem fr_eq:
+ assumes lp: "islin p"
+ shows "(\<exists> x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> (\<exists> (n,t) \<in> set (uset p). \<exists> (m,s) \<in> set (uset p). Ifm vs (((- Itm vs (x#bs) t / Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /(1 + 1))#bs) p))"
+ (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
+proof
+ assume px: "\<exists> x. ?I x p"
+ have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
+ moreover {assume "?M \<or> ?P" hence "?D" by blast}
+ moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
+ from inf_uset[OF lp nmi npi] have "?F" using px by blast hence "?D" by blast}
+ ultimately show "?D" by blast
+next
+ assume "?D"
+ moreover {assume m:"?M" from minusinf_ex[OF lp m] have "?E" .}
+ moreover {assume p: "?P" from plusinf_ex[OF lp p] have "?E" . }
+ moreover {assume f:"?F" hence "?E" by blast}
+ ultimately show "?E" by blast
+qed
+
+section{* First implementation : Naive by encoding all case splits locally *}
+definition "msubsteq c t d s a r =
+ evaldjf (split conj)
+ [(let cd = c *\<^sub>p d in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (conj (Eq (CP c)) (NEq (CP d)) , Eq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (NEq (CP c)) (Eq (CP d)) , Eq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (Eq (CP c)) (Eq (CP d)) , Eq r)]"
+
+lemma msubsteq_nb: assumes lp: "islin (Eq (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
+ shows "bound0 (msubsteq c t d s a r)"
+proof-
+ have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (NEq (CP cd), Eq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (conj (Eq (CP c)) (NEq (CP d)) , Eq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (NEq (CP c)) (Eq (CP d)) , Eq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (Eq (CP c)) (Eq (CP d)) , Eq r)]. bound0 (split conj x)"
+ using lp by (simp add: Let_def t s )
+ from evaldjf_bound0[OF th] show ?thesis by (simp add: msubsteq_def)
+qed
+
+lemma msubsteq: assumes lp: "islin (Eq (CNP 0 a r))"
+ shows "Ifm vs (x#bs) (msubsteq c t d s a r) = Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Eq (CNP 0 a r))" (is "?lhs = ?rhs")
+proof-
+ let ?Nt = "\<lambda>(x::'a) t. Itm vs (x#bs) t"
+ let ?N = "\<lambda>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 \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all
+ note r= tmbound0_I[OF lin(3), of vs _ bs x]
+ have cd_cs: "?c * ?d \<noteq> 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d \<noteq> 0) \<or> (?c \<noteq> 0 \<and> ?d = 0)" by auto
+ moreover
+ {assume c: "?c = 0" and d: "?d=0"
+ hence ?thesis by (simp add: r[of 0] msubsteq_def Let_def evaldjf_ex)}
+ moreover
+ {assume c: "?c = 0" and d: "?d\<noteq>0"
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?s / ((1 + 1)*?d)" by simp
+ have "?rhs = Ifm vs (-?s / ((1 + 1)*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (-?s / ((1 + 1)*?d)) + ?r = 0" by (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) = 0"
+ using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp
+ also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r= 0"
+ by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
+
+ also have "\<dots> \<longleftrightarrow> - (?a * ?s) + (1 + 1)*?d*?r = 0" using d by simp
+ finally have ?thesis using c d
+ apply (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ apply simp
+ done}
+ moreover
+ {assume c: "?c \<noteq> 0" and d: "?d=0"
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?t / ((1 + 1)*?c)" by simp
+ have "?rhs = Ifm vs (-?t / ((1 + 1)*?c) # bs) (Eq (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (-?t / ((1 + 1)*?c)) + ?r = 0" by (simp add: r[of "- (?t/ ((1 + 1)* ?c))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) = 0"
+ using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp
+ also have "\<dots> \<longleftrightarrow> (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r= 0"
+ by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
+ also have "\<dots> \<longleftrightarrow> - (?a * ?t) + (1 + 1)*?c*?r = 0" using c by simp
+ finally have ?thesis using c d
+ apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubsteq_def Let_def evaldjf_ex del: one_add_one_is_two)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ apply simp
+ done }
+ moreover
+ {assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *(1 + 1) \<noteq> 0" by simp
+ from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+ have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
+ by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Eq (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r = 0"
+ by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+ also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) =0 "
+ using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r =0"
+ using nonzero_mult_divide_cancel_left[OF dc] c d
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using c d
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubsteq_def Let_def evaldjf_ex ring_simps)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ apply (simp add: ring_simps)
+ done }
+ ultimately show ?thesis by blast
+qed
+
+
+definition "msubstneq c t d s a r =
+ evaldjf (split conj)
+ [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (conj (Eq (CP c)) (NEq (CP d)) , NEq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (NEq (CP c)) (Eq (CP d)) , NEq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (Eq (CP c)) (Eq (CP d)) , NEq r)]"
+
+lemma msubstneq_nb: assumes lp: "islin (NEq (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
+ shows "bound0 (msubstneq c t d s a r)"
+proof-
+ have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (NEq (CP cd), NEq (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (conj (Eq (CP c)) (NEq (CP d)) , NEq (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (NEq (CP c)) (Eq (CP d)) , NEq (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (Eq (CP c)) (Eq (CP d)) , NEq r)]. bound0 (split conj x)"
+ using lp by (simp add: Let_def t s )
+ from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstneq_def)
+qed
+
+lemma msubstneq: assumes lp: "islin (Eq (CNP 0 a r))"
+ shows "Ifm vs (x#bs) (msubstneq c t d s a r) = Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (NEq (CNP 0 a r))" (is "?lhs = ?rhs")
+proof-
+ let ?Nt = "\<lambda>(x::'a) t. Itm vs (x#bs) t"
+ let ?N = "\<lambda>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 \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all
+ note r= tmbound0_I[OF lin(3), of vs _ bs x]
+ have cd_cs: "?c * ?d \<noteq> 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d \<noteq> 0) \<or> (?c \<noteq> 0 \<and> ?d = 0)" by auto
+ moreover
+ {assume c: "?c = 0" and d: "?d=0"
+ hence ?thesis by (simp add: r[of 0] msubstneq_def Let_def evaldjf_ex)}
+ moreover
+ {assume c: "?c = 0" and d: "?d\<noteq>0"
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?s / ((1 + 1)*?d)" by simp
+ have "?rhs = Ifm vs (-?s / ((1 + 1)*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (-?s / ((1 + 1)*?d)) + ?r \<noteq> 0" by (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a * (-?s / ((1 + 1)*?d)) + ?r) \<noteq> 0"
+ using d mult_cancel_left[of "(1 + 1)*?d" "(?a * (-?s / ((1 + 1)*?d)) + ?r)" 0] by simp
+ also have "\<dots> \<longleftrightarrow> (- ?a * ?s) * ((1 + 1)*?d / ((1 + 1)*?d)) + (1 + 1)*?d*?r\<noteq> 0"
+ by (simp add: ring_simps right_distrib[of "(1 + 1)*?d"] del: right_distrib)
+
+ also have "\<dots> \<longleftrightarrow> - (?a * ?s) + (1 + 1)*?d*?r \<noteq> 0" using d by simp
+ finally have ?thesis using c d
+ apply (simp add: r[of "- (Itm vs (x # bs) s / ((1 + 1) * \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ apply simp
+ done}
+ moreover
+ {assume c: "?c \<noteq> 0" and d: "?d=0"
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = -?t / ((1 + 1)*?c)" by simp
+ have "?rhs = Ifm vs (-?t / ((1 + 1)*?c) # bs) (NEq (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (-?t / ((1 + 1)*?c)) + ?r \<noteq> 0" by (simp add: r[of "- (?t/ ((1 + 1)* ?c))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a * (-?t / ((1 + 1)*?c)) + ?r) \<noteq> 0"
+ using c mult_cancel_left[of "(1 + 1)*?c" "(?a * (-?t / ((1 + 1)*?c)) + ?r)" 0] by simp
+ also have "\<dots> \<longleftrightarrow> (?a * -?t)* ((1 + 1)*?c) / ((1 + 1)*?c) + (1 + 1)*?c*?r \<noteq> 0"
+ by (simp add: ring_simps right_distrib[of "(1 + 1)*?c"] del: right_distrib)
+ also have "\<dots> \<longleftrightarrow> - (?a * ?t) + (1 + 1)*?c*?r \<noteq> 0" using c by simp
+ finally have ?thesis using c d
+ apply (simp add: r[of "- (?t/ ((1 + 1)*?c))"] msubstneq_def Let_def evaldjf_ex del: one_add_one_is_two)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ apply simp
+ done }
+ moreover
+ {assume c: "?c \<noteq> 0" and d: "?d\<noteq>0" hence dc: "?c * ?d *(1 + 1) \<noteq> 0" by simp
+ from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+ have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
+ by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (NEq (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r \<noteq> 0"
+ by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+ also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) \<noteq> 0 "
+ using c d mult_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r \<noteq> 0"
+ using nonzero_mult_divide_cancel_left[OF dc] c d
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using c d
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstneq_def Let_def evaldjf_ex ring_simps)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ apply (simp add: ring_simps)
+ done }
+ ultimately show ?thesis by blast
+qed
+
+definition "msubstlt c t d s a r =
+ evaldjf (split conj)
+ [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (let cd = c *\<^sub>p d in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Lt (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (lt (CP c)) (Eq (CP d)) , Lt (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Lt (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (lt (CP d)) (Eq (CP c)) , Lt (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (Eq (CP c)) (Eq (CP d)) , Lt r)]"
+
+lemma msubstlt_nb: assumes lp: "islin (Lt (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
+ shows "bound0 (msubstlt c t d s a r)"
+proof-
+ have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Lt (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (let cd = c *\<^sub>p d in (lt (CP cd), Lt (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Lt (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (lt (CP c)) (Eq (CP d)) , Lt (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Lt (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (lt (CP d)) (Eq (CP c)) , Lt (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (Eq (CP c)) (Eq (CP d)) , Lt r)]. bound0 (split conj x)"
+ using lp by (simp add: Let_def t s lt_nb )
+ from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstlt_def)
+qed
+
+
+lemma msubstlt: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Lt (CNP 0 a r))"
+ shows "Ifm vs (x#bs) (msubstlt c t d s a r) \<longleftrightarrow>
+ Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Lt (CNP 0 a r))" (is "?lhs = ?rhs")
+proof-
+ let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
+ let ?N = "\<lambda>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 \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all
+ note r= tmbound0_I[OF lin(3), of vs _ bs x]
+ have cd_cs: "?c * ?d < 0 \<or> ?c * ?d > 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d < 0) \<or> (?c = 0 \<and> ?d > 0) \<or> (?c < 0 \<and> ?d = 0) \<or> (?c > 0 \<and> ?d = 0)" by auto
+ moreover
+ {assume c: "?c=0" and d: "?d=0"
+ hence ?thesis using nc nd by (simp add: polyneg_norm lt r[of 0] msubstlt_def Let_def evaldjf_ex)}
+ moreover
+ {assume dc: "?c*?d > 0"
+ from mult_pos_pos[OF one_plus_one_pos dc] have dc': "(1 + 1)*?c *?d > 0" by simp
+ hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
+ from dc' have dc'': "\<not> (1 + 1)*?c *?d < 0" by simp
+ from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+ have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
+ by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0"
+ by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+ also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) < 0"
+
+ using dc' dc'' mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r < 0"
+ using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using dc c d nc nd dc'
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ by (simp add: ring_simps order_less_not_sym[OF dc])}
+ moreover
+ {assume dc: "?c*?d < 0"
+
+ from dc one_plus_one_pos[where ?'a='a] have dc': "(1 + 1)*?c *?d < 0"
+ apply (simp add: mult_less_0_iff field_simps)
+ apply (rule add_neg_neg)
+ apply (simp_all add: mult_less_0_iff)
+ done
+ hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
+ from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+ have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
+ by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r < 0"
+ by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+
+ also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) > 0"
+
+ using dc' order_less_not_sym[OF dc'] mult_less_cancel_left_disj[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp
+ also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r < 0"
+ using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using dc c d nc nd
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ by (simp add: ring_simps order_less_not_sym[OF dc]) }
+ moreover
+ {assume c: "?c > 0" and d: "?d=0"
+ from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff)
+ from c have c': "(1 + 1)*?c \<noteq> 0" by simp
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) < 0"
+ using c mult_less_cancel_left_disj[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
+ also have "\<dots> \<longleftrightarrow> - ?a*?t+ (1 + 1)*?c *?r < 0"
+ using nonzero_mult_divide_cancel_left[OF c'] c
+ by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+ finally have ?thesis using c d nc nd
+ apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ using c order_less_not_sym[OF c] less_imp_neq[OF c]
+ by (simp add: ring_simps ) }
+ moreover
+ {assume c: "?c < 0" and d: "?d=0" hence c': "(1 + 1)*?c \<noteq> 0" by simp
+ from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff)
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r < 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) > 0"
+ using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_less_cancel_left_disj[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
+ also have "\<dots> \<longleftrightarrow> ?a*?t - (1 + 1)*?c *?r < 0"
+ using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using c d nc nd
+ apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ using c order_less_not_sym[OF c] less_imp_neq[OF c]
+ by (simp add: ring_simps ) }
+ moreover
+ moreover
+ {assume c: "?c = 0" and d: "?d>0"
+ from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff)
+ from d have d': "(1 + 1)*?d \<noteq> 0" by simp
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) < 0"
+ using d mult_less_cancel_left_disj[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
+ also have "\<dots> \<longleftrightarrow> - ?a*?s+ (1 + 1)*?d *?r < 0"
+ using nonzero_mult_divide_cancel_left[OF d'] d
+ by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+ finally have ?thesis using c d nc nd
+ apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ using d order_less_not_sym[OF d] less_imp_neq[OF d]
+ by (simp add: ring_simps ) }
+ moreover
+ {assume c: "?c = 0" and d: "?d<0" hence d': "(1 + 1)*?d \<noteq> 0" by simp
+ from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff)
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Lt (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r < 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) > 0"
+ using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_less_cancel_left_disj[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
+ also have "\<dots> \<longleftrightarrow> ?a*?s - (1 + 1)*?d *?r < 0"
+ using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using c d nc nd
+ apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstlt_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ using d order_less_not_sym[OF d] less_imp_neq[OF d]
+ by (simp add: ring_simps ) }
+ultimately show ?thesis by blast
+qed
+
+definition "msubstle c t d s a r =
+ evaldjf (split conj)
+ [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (let cd = c *\<^sub>p d in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Le (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (lt (CP c)) (Eq (CP d)) , Le (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Le (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (lt (CP d)) (Eq (CP c)) , Le (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (Eq (CP c)) (Eq (CP d)) , Le r)]"
+
+lemma msubstle_nb: assumes lp: "islin (Le (CNP 0 a r))" and t: "tmbound0 t" and s: "tmbound0 s"
+ shows "bound0 (msubstle c t d s a r)"
+proof-
+ have th: "\<forall>x\<in> set [(let cd = c *\<^sub>p d in (lt (CP (~\<^sub>p cd)), Le (Add (Mul (~\<^sub>p a) (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (let cd = c *\<^sub>p d in (lt (CP cd), Le (Sub (Mul a (Add (Mul d t) (Mul c s))) (Mul (2\<^sub>p *\<^sub>p cd) r)))),
+ (conj (lt (CP (~\<^sub>p c))) (Eq (CP d)) , Le (Add (Mul (~\<^sub>p a) t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (lt (CP c)) (Eq (CP d)) , Le (Sub (Mul a t) (Mul (2\<^sub>p *\<^sub>p c) r))),
+ (conj (lt (CP (~\<^sub>p d))) (Eq (CP c)) , Le (Add (Mul (~\<^sub>p a) s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (lt (CP d)) (Eq (CP c)) , Le (Sub (Mul a s) (Mul (2\<^sub>p *\<^sub>p d) r))),
+ (conj (Eq (CP c)) (Eq (CP d)) , Le r)]. bound0 (split conj x)"
+ using lp by (simp add: Let_def t s lt_nb )
+ from evaldjf_bound0[OF th] show ?thesis by (simp add: msubstle_def)
+qed
+
+lemma msubstle: assumes nc: "isnpoly c" and nd: "isnpoly d" and lp: "islin (Le (CNP 0 a r))"
+ shows "Ifm vs (x#bs) (msubstle c t d s a r) \<longleftrightarrow>
+ Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) (Le (CNP 0 a r))" (is "?lhs = ?rhs")
+proof-
+ let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
+ let ?N = "\<lambda>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 \<noteq> 0\<^sub>p" "tmbound0 r" "allpolys isnpoly r" by simp_all
+ note r= tmbound0_I[OF lin(3), of vs _ bs x]
+ have cd_cs: "?c * ?d < 0 \<or> ?c * ?d > 0 \<or> (?c = 0 \<and> ?d = 0) \<or> (?c = 0 \<and> ?d < 0) \<or> (?c = 0 \<and> ?d > 0) \<or> (?c < 0 \<and> ?d = 0) \<or> (?c > 0 \<and> ?d = 0)" by auto
+ moreover
+ {assume c: "?c=0" and d: "?d=0"
+ hence ?thesis using nc nd by (simp add: polyneg_norm polymul_norm lt r[of 0] msubstle_def Let_def evaldjf_ex)}
+ moreover
+ {assume dc: "?c*?d > 0"
+ from mult_pos_pos[OF one_plus_one_pos dc] have dc': "(1 + 1)*?c *?d > 0" by simp
+ hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
+ from dc' have dc'': "\<not> (1 + 1)*?c *?d < 0" by simp
+ from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+ have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
+ by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0"
+ by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+ also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) <= 0"
+
+ using dc' dc'' mult_le_cancel_left[of "(1 + 1) * ?c * ?d" "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r" 0] by simp
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )) + (1 + 1)*?c*?d*?r <= 0"
+ using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using dc c d nc nd dc'
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ by (simp add: ring_simps order_less_not_sym[OF dc])}
+ moreover
+ {assume dc: "?c*?d < 0"
+
+ from dc one_plus_one_pos[where ?'a='a] have dc': "(1 + 1)*?c *?d < 0"
+ by (simp add: mult_less_0_iff field_simps add_neg_neg add_pos_pos)
+ hence c:"?c \<noteq> 0" and d: "?d\<noteq> 0" by auto
+ from add_frac_eq[OF c d, of "- ?t" "- ?s"]
+ have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)"
+ by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r <= 0"
+ by (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"])
+
+ also have "\<dots> \<longleftrightarrow> ((1 + 1) * ?c * ?d) * (?a * (- (?d * ?t + ?c* ?s )/ ((1 + 1)*?c*?d)) + ?r) >= 0"
+
+ using dc' order_less_not_sym[OF dc'] mult_le_cancel_left[of "(1 + 1) * ?c * ?d" 0 "?a * (- (?d * ?t + ?c* ?s)/ ((1 + 1)*?c*?d)) + ?r"] by simp
+ also have "\<dots> \<longleftrightarrow> ?a * ((?d * ?t + ?c* ?s )) - (1 + 1)*?c*?d*?r <= 0"
+ using nonzero_mult_divide_cancel_left[of "(1 + 1)*?c*?d"] c d
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using dc c d nc nd
+ apply (simp add: r[of "(- (?d * ?t) + - (?c *?s)) / ((1 + 1) * ?c * ?d)"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ by (simp add: ring_simps order_less_not_sym[OF dc]) }
+ moreover
+ {assume c: "?c > 0" and d: "?d=0"
+ from c have c'': "(1 + 1)*?c > 0" by (simp add: zero_less_mult_iff)
+ from c have c': "(1 + 1)*?c \<noteq> 0" by simp
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) <= 0"
+ using c mult_le_cancel_left[of "(1 + 1) * ?c" "?a* (- ?t / ((1 + 1)*?c))+ ?r" 0] c' c'' order_less_not_sym[OF c''] by simp
+ also have "\<dots> \<longleftrightarrow> - ?a*?t+ (1 + 1)*?c *?r <= 0"
+ using nonzero_mult_divide_cancel_left[OF c'] c
+ by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+ finally have ?thesis using c d nc nd
+ apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ using c order_less_not_sym[OF c] less_imp_neq[OF c]
+ by (simp add: ring_simps ) }
+ moreover
+ {assume c: "?c < 0" and d: "?d=0" hence c': "(1 + 1)*?c \<noteq> 0" by simp
+ from c have c'': "(1 + 1)*?c < 0" by (simp add: mult_less_0_iff)
+ from d have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?t / ((1 + 1)*?c)" by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- ?t / ((1 + 1)*?c) # bs) (Le (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a* (- ?t / ((1 + 1)*?c))+ ?r <= 0" by (simp add: r[of "- (?t / ((1 + 1)*?c))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?c * (?a* (- ?t / ((1 + 1)*?c))+ ?r) >= 0"
+ using c order_less_not_sym[OF c''] less_imp_neq[OF c''] c'' mult_le_cancel_left[of "(1 + 1) * ?c" 0 "?a* (- ?t / ((1 + 1)*?c))+ ?r"] by simp
+ also have "\<dots> \<longleftrightarrow> ?a*?t - (1 + 1)*?c *?r <= 0"
+ using nonzero_mult_divide_cancel_left[OF c'] c order_less_not_sym[OF c''] less_imp_neq[OF c''] c''
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using c d nc nd
+ apply(simp add: r[of "- (?t / ((1 + 1)*?c))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ using c order_less_not_sym[OF c] less_imp_neq[OF c]
+ by (simp add: ring_simps ) }
+ moreover
+ moreover
+ {assume c: "?c = 0" and d: "?d>0"
+ from d have d'': "(1 + 1)*?d > 0" by (simp add: zero_less_mult_iff)
+ from d have d': "(1 + 1)*?d \<noteq> 0" by simp
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) <= 0"
+ using d mult_le_cancel_left[of "(1 + 1) * ?d" "?a* (- ?s / ((1 + 1)*?d))+ ?r" 0] d' d'' order_less_not_sym[OF d''] by simp
+ also have "\<dots> \<longleftrightarrow> - ?a*?s+ (1 + 1)*?d *?r <= 0"
+ using nonzero_mult_divide_cancel_left[OF d'] d
+ by (simp add: ring_simps diff_divide_distrib less_le del: left_distrib)
+ finally have ?thesis using c d nc nd
+ apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ using d order_less_not_sym[OF d] less_imp_neq[OF d]
+ by (simp add: ring_simps ) }
+ moreover
+ {assume c: "?c = 0" and d: "?d<0" hence d': "(1 + 1)*?d \<noteq> 0" by simp
+ from d have d'': "(1 + 1)*?d < 0" by (simp add: mult_less_0_iff)
+ from c have th: "(- ?t / ?c + - ?s / ?d)/(1 + 1) = - ?s / ((1 + 1)*?d)" by (simp add: ring_simps)
+ have "?rhs \<longleftrightarrow> Ifm vs (- ?s / ((1 + 1)*?d) # bs) (Le (CNP 0 a r))" by (simp only: th)
+ also have "\<dots> \<longleftrightarrow> ?a* (- ?s / ((1 + 1)*?d))+ ?r <= 0" by (simp add: r[of "- (?s / ((1 + 1)*?d))"])
+ also have "\<dots> \<longleftrightarrow> (1 + 1)*?d * (?a* (- ?s / ((1 + 1)*?d))+ ?r) >= 0"
+ using d order_less_not_sym[OF d''] less_imp_neq[OF d''] d'' mult_le_cancel_left[of "(1 + 1) * ?d" 0 "?a* (- ?s / ((1 + 1)*?d))+ ?r"] by simp
+ also have "\<dots> \<longleftrightarrow> ?a*?s - (1 + 1)*?d *?r <= 0"
+ using nonzero_mult_divide_cancel_left[OF d'] d order_less_not_sym[OF d''] less_imp_neq[OF d''] d''
+ by (simp add: ring_simps diff_divide_distrib del: left_distrib)
+ finally have ?thesis using c d nc nd
+ apply(simp add: r[of "- (?s / ((1 + 1)*?d))"] msubstle_def Let_def evaldjf_ex ring_simps lt polyneg_norm polymul_norm)
+ apply (simp only: one_add_one_is_two[symmetric] of_int_add)
+ using d order_less_not_sym[OF d] less_imp_neq[OF d]
+ by (simp add: ring_simps ) }
+ultimately show ?thesis by blast
+qed
+
+
+fun msubst :: "fm \<Rightarrow> (poly \<times> tm) \<times> (poly \<times> tm) \<Rightarrow> fm" where
+ "msubst (And p q) ((c,t), (d,s)) = conj (msubst p ((c,t),(d,s))) (msubst q ((c,t),(d,s)))"
+| "msubst (Or p q) ((c,t), (d,s)) = disj (msubst p ((c,t),(d,s))) (msubst q ((c,t), (d,s)))"
+| "msubst (Eq (CNP 0 a r)) ((c,t),(d,s)) = msubsteq c t d s a r"
+| "msubst (NEq (CNP 0 a r)) ((c,t),(d,s)) = msubstneq c t d s a r"
+| "msubst (Lt (CNP 0 a r)) ((c,t),(d,s)) = msubstlt c t d s a r"
+| "msubst (Le (CNP 0 a r)) ((c,t),(d,s)) = msubstle c t d s a r"
+| "msubst p ((c,t),(d,s)) = p"
+
+lemma msubst_I: assumes lp: "islin p" and nc: "isnpoly c" and nd: "isnpoly d"
+ shows "Ifm vs (x#bs) (msubst p ((c,t),(d,s))) = Ifm vs (((- Itm vs (x#bs) t / Ipoly vs c + - Itm vs (x#bs) s / Ipoly vs d) /(1 + 1))#bs) p"
+ using lp
+by (induct p rule: islin.induct, auto simp add: tmbound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) /(1 + 1)" and b'=x and bs = bs and vs=vs] bound0_I[where b="(- (Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>) + - (Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>)) /(1 + 1)" and b'=x and bs = bs and vs=vs] msubsteq msubstneq msubstlt[OF nc nd] msubstle[OF nc nd])
+
+lemma msubst_nb: assumes lp: "islin p" and t: "tmbound0 t" and s: "tmbound0 s"
+ shows "bound0 (msubst p ((c,t),(d,s)))"
+ using lp t s
+ by (induct p rule: islin.induct, auto simp add: msubsteq_nb msubstneq_nb msubstlt_nb msubstle_nb)
+
+lemma fr_eq_msubst:
+ assumes lp: "islin p"
+ shows "(\<exists> x. Ifm vs (x#bs) p) = ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> (\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). Ifm vs (x#bs) (msubst p ((c,t),(d,s)))))"
+ (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?F)" is "?E = ?D")
+proof-
+from uset_l[OF lp] have th: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s" by blast
+{fix c t d s assume ctU: "(c,t) \<in>set (uset p)" and dsU: "(d,s) \<in>set (uset p)"
+ and pts: "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p"
+ from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d" by simp_all
+ from msubst_I[OF lp norm, of vs x bs t s] pts
+ have "Ifm vs (x # bs) (msubst p ((c, t), d, s))" ..}
+moreover
+{fix c t d s assume ctU: "(c,t) \<in>set (uset p)" and dsU: "(d,s) \<in>set (uset p)"
+ and pts: "Ifm vs (x # bs) (msubst p ((c, t), d, s))"
+ from th[rule_format, OF ctU] th[rule_format, OF dsU] have norm:"isnpoly c" "isnpoly d" by simp_all
+ from msubst_I[OF lp norm, of vs x bs t s] pts
+ have "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p" ..}
+ultimately have th': "(\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> set (uset p). Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1+1) # bs) p) \<longleftrightarrow> ?F" by blast
+from fr_eq[OF lp, of vs bs x, simplified th'] show ?thesis .
+qed
+
+text {* Rest of the implementation *}
+
+consts alluopairs:: "'a list \<Rightarrow> ('a \<times> 'a) list"
+primrec
+ "alluopairs [] = []"
+ "alluopairs (x#xs) = (map (Pair x) (x#xs))@(alluopairs xs)"
+
+lemma alluopairs_set1: "set (alluopairs xs) \<le> {(x,y). x\<in> set xs \<and> y\<in> set xs}"
+by (induct xs, auto)
+
+lemma alluopairs_set:
+ "\<lbrakk>x\<in> set xs ; y \<in> set xs\<rbrakk> \<Longrightarrow> (x,y) \<in> set (alluopairs xs) \<or> (y,x) \<in> set (alluopairs xs) "
+by (induct xs, auto)
+
+lemma alluopairs_ex:
+ assumes Pc: "\<forall> x \<in> set xs. \<forall>y\<in> set xs. P x y = P y x"
+ shows "(\<exists> x \<in> set xs. \<exists> y \<in> set xs. P x y) = (\<exists> (x,y) \<in> set (alluopairs xs). P x y)"
+proof
+ assume "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y"
+ then obtain x y where x: "x \<in> set xs" and y:"y \<in> set xs" and P: "P x y" by blast
+ from alluopairs_set[OF x y] P Pc x y show"\<exists>(x, y)\<in>set (alluopairs xs). P x y"
+ by auto
+next
+ assume "\<exists>(x, y)\<in>set (alluopairs xs). P x y"
+ then obtain "x" and "y" where xy:"(x,y) \<in> set (alluopairs xs)" and P: "P x y" by blast+
+ from xy have "x \<in> set xs \<and> y\<in> set xs" using alluopairs_set1 by blast
+ with P show "\<exists>x\<in>set xs. \<exists>y\<in>set xs. P x y" by blast
+qed
+
+lemma nth_pos2: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
+using Nat.gr0_conv_Suc
+by clarsimp
+
+lemma filter_length: "length (List.filter P xs) < Suc (length xs)"
+ apply (induct xs, auto) done
+
+consts remdps:: "'a list \<Rightarrow> 'a list"
+
+recdef remdps "measure size"
+ "remdps [] = []"
+ "remdps (x#xs) = (x#(remdps (List.filter (\<lambda> y. y \<noteq> x) xs)))"
+(hints simp add: filter_length[rule_format])
+
+lemma remdps_set[simp]: "set (remdps xs) = set xs"
+ by (induct xs rule: remdps.induct, auto)
+
+lemma simpfm_lin: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "qfree p \<Longrightarrow> islin (simpfm p)"
+ by (induct p rule: simpfm.induct, auto simp add: conj_lin disj_lin)
+
+definition
+ "ferrack p \<equiv> let q = simpfm p ; mp = minusinf q ; pp = plusinf q
+ in if (mp = T \<or> pp = T) then T
+ else (let U = alluopairs (remdps (uset q))
+ in decr0 (disj mp (disj pp (evaldjf (simpfm o (msubst q)) U ))))"
+
+lemma ferrack:
+ assumes qf: "qfree p"
+ shows "qfree (ferrack p) \<and> ((Ifm vs bs (ferrack p)) = (Ifm vs bs (E p)))"
+ (is "_ \<and> (?rhs = ?lhs)")
+proof-
+ let ?I = "\<lambda> x p. Ifm vs (x#bs) p"
+ let ?N = "\<lambda> t. Ipoly vs t"
+ let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
+ let ?q = "simpfm p"
+ let ?U = "remdps(uset ?q)"
+ let ?Up = "alluopairs ?U"
+ let ?mp = "minusinf ?q"
+ let ?pp = "plusinf ?q"
+ let ?I = "\<lambda>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: "\<forall>(c, s)\<in>set ?U. isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
+ by simp
+ {fix c t d s assume ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U"
+ from U_l ctU dsU have norm: "isnpoly c" "isnpoly d" by auto
+ from msubst_I[OF lq norm, of vs x bs t s] msubst_I[OF lq norm(2,1), of vs x bs s t]
+ have "?I (msubst ?q ((c,t),(d,s))) = ?I (msubst ?q ((d,s),(c,t)))" by (simp add: ring_simps)}
+ hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (msubst ?q (x, y)) \<longleftrightarrow> ?I (msubst ?q (y, x))" by clarsimp
+ {fix x assume xUp: "x \<in> set ?Up"
+ then obtain c t d s where ctU: "(c,t) \<in> set ?U" and dsU: "(d,s) \<in> set ?U"
+ and x: "x = ((c,t),(d,s))" using alluopairs_set1[of ?U] by auto
+ from U_l[rule_format, OF ctU] U_l[rule_format, OF dsU]
+ have nbs: "tmbound0 t" "tmbound0 s" by simp_all
+ from simpfm_bound0[OF msubst_nb[OF lq nbs, of c d]]
+ have "bound0 ((simpfm o (msubst (simpfm p))) x)" using x by simp}
+ with evaldjf_bound0[of ?Up "(simpfm o (msubst (simpfm p)))"]
+ have "bound0 (evaldjf (simpfm o (msubst (simpfm p))) ?Up)" by blast
+ with mp_nb pp_nb
+ have th1: "bound0 (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up )))" by (simp add: disj_nb)
+ from decr0_qf[OF th1] have thqf: "qfree (ferrack p)" by (simp add: ferrack_def Let_def)
+ have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
+ also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists>(c, t)\<in>set ?U. \<exists>(d, s)\<in>set ?U. ?I (msubst (simpfm p) ((c, t), d, s)))" using fr_eq_msubst[OF lq, of vs bs x] by simp
+ also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> (\<exists> (x,y) \<in> set ?Up. ?I ((simpfm o (msubst ?q)) (x,y)))" using alluopairs_ex[OF th0] by simp
+ also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (evaldjf (simpfm o (msubst ?q)) ?Up)"
+ by (simp add: evaldjf_ex)
+ also have "\<dots> \<longleftrightarrow> ?I (disj ?mp (disj ?pp (evaldjf (simpfm o (msubst ?q)) ?Up)))" by simp
+ also have "\<dots> \<longleftrightarrow> ?rhs" using decr0[OF th1, of vs x bs]
+ apply (simp add: ferrack_def Let_def)
+ by (cases "?mp = T \<or> ?pp = T", auto)
+ finally show ?thesis using thqf by blast
+qed
+
+definition "frpar p = simpfm (qelim p ferrack)"
+lemma frpar: "qfree (frpar p) \<and> (Ifm vs bs (frpar p) \<longleftrightarrow> Ifm vs bs p)"
+proof-
+ from ferrack have th: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack p) \<and> Ifm vs bs (ferrack p) = Ifm vs bs (E p)" by blast
+ from qelim[OF th, of p bs] show ?thesis unfolding frpar_def by auto
+qed
+
+declare polyadd.simps[code]
+lemma [simp,code]: "polyadd (CN c n p, CN c' n' p') =
+ (if n < n' then CN (polyadd(c,CN c' n' p')) n p
+ else if n'<n then CN (polyadd(CN c n p, c')) n' p'
+ else (let cc' = polyadd (c,c') ;
+ pp' = polyadd (p,p')
+ in (if pp' = 0\<^sub>p then cc' else CN cc' n pp')))"
+ by (simp add: Let_def stupid)
+
+
+
+(*
+lemmas [code func] = polysub_def
+lemmas [code func del] = Zero_nat_def
+code_gen "frpar" in SML to FRParTest
+*)
+
+section{* Second implemenation: Case splits not local *}
+
+lemma fr_eq2: assumes lp: "islin p"
+ shows "(\<exists> x. Ifm vs (x#bs) p) \<longleftrightarrow>
+ ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or>
+ (Ifm vs (0#bs) p) \<or>
+ (\<exists> (n,t) \<in> set (uset p). Ipoly vs n \<noteq> 0 \<and> Ifm vs ((- Itm vs (x#bs) t / (Ipoly vs n * (1 + 1)))#bs) p) \<or>
+ (\<exists> (n,t) \<in> set (uset p). \<exists> (m,s) \<in> set (uset p). Ipoly vs n \<noteq> 0 \<and> Ipoly vs m \<noteq> 0 \<and> Ifm vs (((- Itm vs (x#bs) t / Ipoly vs n + - Itm vs (x#bs) s / Ipoly vs m) /(1 + 1))#bs) p))"
+ (is "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?Z \<or> ?U \<or> ?F)" is "?E = ?D")
+proof
+ assume px: "\<exists> x. ?I x p"
+ have "?M \<or> ?P \<or> (\<not> ?M \<and> \<not> ?P)" by blast
+ moreover {assume "?M \<or> ?P" hence "?D" by blast}
+ moreover {assume nmi: "\<not> ?M" and npi: "\<not> ?P"
+ from inf_uset[OF lp nmi npi, OF px]
+ obtain c t d s where ct: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)" "?I ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / ((1\<Colon>'a) + (1\<Colon>'a))) p"
+ by auto
+ let ?c = "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
+ let ?d = "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>"
+ let ?s = "Itm vs (x # bs) s"
+ let ?t = "Itm vs (x # bs) t"
+ have eq2: "\<And>(x::'a). x + x = (1 + 1) * x"
+ by (simp add: ring_simps)
+ {assume "?c = 0 \<and> ?d = 0"
+ with ct have ?D by simp}
+ moreover
+ {assume z: "?c = 0" "?d \<noteq> 0"
+ from z have ?D using ct by auto}
+ moreover
+ {assume z: "?c \<noteq> 0" "?d = 0"
+ with ct have ?D by auto }
+ moreover
+ {assume z: "?c \<noteq> 0" "?d \<noteq> 0"
+ from z have ?F using ct
+ apply - apply (rule bexI[where x = "(c,t)"], simp_all)
+ by (rule bexI[where x = "(d,s)"], simp_all)
+ hence ?D by blast}
+ ultimately have ?D by auto}
+ ultimately show "?D" by blast
+next
+ assume "?D"
+ moreover {assume m:"?M" from minusinf_ex[OF lp m] have "?E" .}
+ moreover {assume p: "?P" from plusinf_ex[OF lp p] have "?E" . }
+ moreover {assume f:"?F" hence "?E" by blast}
+ ultimately show "?E" by blast
+qed
+
+definition "msubsteq2 c t a b = Eq (Add (Mul a t) (Mul c b))"
+definition "msubstltpos c t a b = Lt (Add (Mul a t) (Mul c b))"
+definition "msubstlepos c t a b = Le (Add (Mul a t) (Mul c b))"
+definition "msubstltneg c t a b = Lt (Neg (Add (Mul a t) (Mul c b)))"
+definition "msubstleneg c t a b = Le (Neg (Add (Mul a t) (Mul c b)))"
+
+lemma msubsteq2:
+ assumes nz: "Ipoly vs c \<noteq> 0" and l: "islin (Eq (CNP 0 a b))"
+ shows "Ifm vs (x#bs) (msubsteq2 c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Eq (CNP 0 a b))" (is "?lhs = ?rhs")
+ using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+ by (simp add: msubsteq2_def field_simps)
+
+lemma msubstltpos:
+ assumes nz: "Ipoly vs c > 0" and l: "islin (Lt (CNP 0 a b))"
+ shows "Ifm vs (x#bs) (msubstltpos c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Lt (CNP 0 a b))" (is "?lhs = ?rhs")
+ using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+ by (simp add: msubstltpos_def field_simps)
+
+lemma msubstlepos:
+ assumes nz: "Ipoly vs c > 0" and l: "islin (Le (CNP 0 a b))"
+ shows "Ifm vs (x#bs) (msubstlepos c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Le (CNP 0 a b))" (is "?lhs = ?rhs")
+ using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+ by (simp add: msubstlepos_def field_simps)
+
+lemma msubstltneg:
+ assumes nz: "Ipoly vs c < 0" and l: "islin (Lt (CNP 0 a b))"
+ shows "Ifm vs (x#bs) (msubstltneg c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Lt (CNP 0 a b))" (is "?lhs = ?rhs")
+ using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+ by (simp add: msubstltneg_def field_simps del: minus_add_distrib)
+
+lemma msubstleneg:
+ assumes nz: "Ipoly vs c < 0" and l: "islin (Le (CNP 0 a b))"
+ shows "Ifm vs (x#bs) (msubstleneg c t a b) = Ifm vs (((Itm vs (x#bs) t / Ipoly vs c ))#bs) (Le (CNP 0 a b))" (is "?lhs = ?rhs")
+ using nz l tmbound0_I[of b vs x bs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" , symmetric]
+ by (simp add: msubstleneg_def field_simps del: minus_add_distrib)
+
+fun msubstpos :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm" where
+ "msubstpos (And p q) c t = And (msubstpos p c t) (msubstpos q c t)"
+| "msubstpos (Or p q) c t = Or (msubstpos p c t) (msubstpos q c t)"
+| "msubstpos (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
+| "msubstpos (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
+| "msubstpos (Lt (CNP 0 a r)) c t = msubstltpos c t a r"
+| "msubstpos (Le (CNP 0 a r)) c t = msubstlepos c t a r"
+| "msubstpos p c t = p"
+
+lemma msubstpos_I:
+ assumes lp: "islin p" and pos: "Ipoly vs c > 0"
+ shows "Ifm vs (x#bs) (msubstpos p c t) = Ifm vs (Itm vs (x#bs) t / Ipoly vs c #bs) p"
+ using lp pos
+ by (induct p rule: islin.induct, auto simp add: msubsteq2 msubstltpos[OF pos] msubstlepos[OF pos] tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] field_simps)
+
+fun msubstneg :: "fm \<Rightarrow> poly \<Rightarrow> tm \<Rightarrow> fm" where
+ "msubstneg (And p q) c t = And (msubstneg p c t) (msubstneg q c t)"
+| "msubstneg (Or p q) c t = Or (msubstneg p c t) (msubstneg q c t)"
+| "msubstneg (Eq (CNP 0 a r)) c t = msubsteq2 c t a r"
+| "msubstneg (NEq (CNP 0 a r)) c t = NOT (msubsteq2 c t a r)"
+| "msubstneg (Lt (CNP 0 a r)) c t = msubstltneg c t a r"
+| "msubstneg (Le (CNP 0 a r)) c t = msubstleneg c t a r"
+| "msubstneg p c t = p"
+
+lemma msubstneg_I:
+ assumes lp: "islin p" and pos: "Ipoly vs c < 0"
+ shows "Ifm vs (x#bs) (msubstneg p c t) = Ifm vs (Itm vs (x#bs) t / Ipoly vs c #bs) p"
+ using lp pos
+ by (induct p rule: islin.induct, auto simp add: msubsteq2 msubstltneg[OF pos] msubstleneg[OF pos] tmbound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup>" bs x] bound0_I[of _ vs "Itm vs (x # bs) t / \<lparr>c\<rparr>\<^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 \<noteq> 0"
+ shows "Ifm vs (x#bs) (msubst2 p c t) = Ifm vs (Itm vs (x#bs) t / Ipoly vs c #bs) p"
+proof-
+ let ?c = "Ipoly vs c"
+ from nc have anc: "allpolys isnpoly (CP c)" "allpolys isnpoly (CP (~\<^sub>p c))"
+ by (simp_all add: polyneg_norm)
+ from nz have "?c > 0 \<or> ?c < 0" by arith
+ moreover
+ {assume c: "?c < 0"
+ from c msubstneg_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
+ have ?thesis by (auto simp add: msubst2_def)}
+ moreover
+ {assume c: "?c > 0"
+ from c msubstpos_I[OF lp c, of x bs t] lt[OF anc(1), of vs "x#bs"] lt[OF anc(2), of vs "x#bs"]
+ have ?thesis by (auto simp add: msubst2_def)}
+ ultimately show ?thesis by blast
+qed
+
+term msubsteq2
+lemma msubsteq2_nb: "tmbound0 t \<Longrightarrow> islin (Eq (CNP 0 a r)) \<Longrightarrow> bound0 (msubsteq2 c t a r)"
+ by (simp add: msubsteq2_def)
+
+lemma msubstltpos_nb: "tmbound0 t \<Longrightarrow> islin (Lt (CNP 0 a r)) \<Longrightarrow> bound0 (msubstltpos c t a r)"
+ by (simp add: msubstltpos_def)
+lemma msubstltneg_nb: "tmbound0 t \<Longrightarrow> islin (Lt (CNP 0 a r)) \<Longrightarrow> bound0 (msubstltneg c t a r)"
+ by (simp add: msubstltneg_def)
+
+lemma msubstlepos_nb: "tmbound0 t \<Longrightarrow> islin (Le (CNP 0 a r)) \<Longrightarrow> bound0 (msubstlepos c t a r)"
+ by (simp add: msubstlepos_def)
+lemma msubstleneg_nb: "tmbound0 t \<Longrightarrow> islin (Le (CNP 0 a r)) \<Longrightarrow> bound0 (msubstleneg c t a r)"
+ by (simp add: msubstleneg_def)
+
+lemma msubstpos_nb: assumes lp: "islin p" and tnb: "tmbound0 t"
+ shows "bound0 (msubstpos p c t)"
+using lp tnb
+by (induct p c t rule: msubstpos.induct, auto simp add: msubsteq2_nb msubstltpos_nb msubstlepos_nb)
+
+lemma msubstneg_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
+ shows "bound0 (msubstneg p c t)"
+using lp tnb
+by (induct p c t rule: msubstneg.induct, auto simp add: msubsteq2_nb msubstltneg_nb msubstleneg_nb)
+
+lemma msubst2_nb: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})" and lp: "islin p" and tnb: "tmbound0 t"
+ shows "bound0 (msubst2 p c t)"
+using lp tnb
+by (simp add: msubst2_def msubstneg_nb msubstpos_nb conj_nb disj_nb lt_nb simpfm_bound0)
+
+lemma of_int2: "of_int 2 = 1 + 1"
+proof-
+ have "(2::int) = 1 + 1" by simp
+ hence "of_int 2 = of_int (1 + 1)" by simp
+ thus ?thesis unfolding of_int_add by simp
+qed
+
+lemma of_int_minus2: "of_int (-2) = - (1 + 1)"
+proof-
+ have th: "(-2::int) = - 2" by simp
+ show ?thesis unfolding th by (simp only: of_int_minus of_int2)
+qed
+
+
+lemma islin_qf: "islin p \<Longrightarrow> qfree p"
+ by (induct p rule: islin.induct, auto simp add: bound0_qf)
+lemma fr_eq_msubst2:
+ assumes lp: "islin p"
+ shows "(\<exists> x. Ifm vs (x#bs) p) \<longleftrightarrow> ((Ifm vs (x#bs) (minusinf p)) \<or> (Ifm vs (x#bs) (plusinf p)) \<or> Ifm vs (x#bs) (subst0 (CP 0\<^sub>p) p) \<or> (\<exists>(n, t)\<in>set (uset p). Ifm vs (x# bs) (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<or> (\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> 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 "(\<exists> x. ?I x p) = (?M \<or> ?P \<or> ?Pz \<or> ?PU \<or> ?F)" is "?E = ?D")
+proof-
+ from uset_l[OF lp] have th: "\<forall>(c, s)\<in>set (uset p). isnpoly c \<and> tmbound0 s" by blast
+ let ?I = "\<lambda>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: "(\<exists>(n, t)\<in>set (uset p). ?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)) \<longleftrightarrow> (\<exists>(n, t)\<in>set (uset p). \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p)"
+ proof-
+ {fix n t assume H: "(n, t)\<in>set (uset p)" "?I(msubst2 p (n *\<^sub>p C (-2, 1)) t)"
+ from H(1) th have "isnpoly n" by blast
+ hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" by (simp_all add: polymul_norm n2)
+ have nn': "allpolys isnpoly (CP (~\<^sub>p (n *\<^sub>p C (-2, 1))))"
+ by (simp add: polyneg_norm nn)
+ hence nn2: "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>n \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" using H(2) nn' nn
+ by (auto simp add: msubst2_def lt zero_less_mult_iff mult_less_0_iff)
+ from msubst2[OF lp nn nn2(1), of x bs t]
+ have "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p"
+ using H(2) nn2 by (simp add: of_int_minus2 del: minus_add_distrib)}
+ moreover
+ {fix n t assume H: "(n, t)\<in>set (uset p)" "\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs (- Itm vs (x # bs) t / (\<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> * (1 + 1)) # bs) p"
+ from H(1) th have "isnpoly n" by blast
+ hence nn: "isnpoly (n *\<^sub>p (C (-2,1)))" "\<lparr>n *\<^sub>p(C (-2,1)) \<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
+ using H(2) by (simp_all add: polymul_norm n2)
+ from msubst2[OF lp nn, of x bs t] have "?I (msubst2 p (n *\<^sub>p (C (-2,1))) t)" using H(2,3) by (simp add: of_int_minus2 del: minus_add_distrib)}
+ ultimately show ?thesis by blast
+ qed
+ have eq2: "(\<exists> (c,t) \<in> set (uset p). \<exists> (d,s) \<in> 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)))) \<longleftrightarrow> (\<exists>(n, t)\<in>set (uset p).
+ \<exists>(m, s)\<in>set (uset p). \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>n\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>m\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p)"
+ proof-
+ {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)"
+ "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))"
+ from H(1,2) th have "isnpoly c" "isnpoly d" by blast+
+ hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)"
+ by (simp_all add: polymul_norm n2)
+ have stupid: "allpolys isnpoly (CP (~\<^sub>p (C (-2, 1) *\<^sub>p c *\<^sub>p d)))" "allpolys isnpoly (CP ((C (-2, 1) *\<^sub>p c *\<^sub>p d)))"
+ by (simp_all add: polyneg_norm nn)
+ have nn': "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
+ using H(3) by (auto simp add: msubst2_def lt[OF stupid(1)] lt[OF stupid(2)] zero_less_mult_iff mult_less_0_iff)
+ from msubst2[OF lp nn nn'(1), of x bs ] H(3) nn'
+ have "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0 \<and> Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p"
+ apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib)
+ by (simp add: mult_commute)}
+ moreover
+ {fix c t d s assume H: "(c,t) \<in> set (uset p)" "(d,s) \<in> set (uset p)"
+ "\<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "Ifm vs ((- Itm vs (x # bs) t / \<lparr>c\<rparr>\<^sub>p\<^bsup>vs\<^esup> + - Itm vs (x # bs) s / \<lparr>d\<rparr>\<^sub>p\<^bsup>vs\<^esup>) / (1 + 1) # bs) p"
+ from H(1,2) th have "isnpoly c" "isnpoly d" by blast+
+ hence nn: "isnpoly (C (-2, 1) *\<^sub>p c*\<^sub>p d)" "\<lparr>(C (-2, 1) *\<^sub>p c*\<^sub>p d)\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
+ using H(3,4) by (simp_all add: polymul_norm n2)
+ from msubst2[OF lp nn, of x bs ] H(3,4,5)
+ have "Ifm vs (x#bs) (msubst2 p (C (-2, 1) *\<^sub>p c*\<^sub>p d) (Add (Mul d t) (Mul c s)))" apply (simp add: add_divide_distrib of_int_minus2 del: minus_add_distrib)by (simp add: mult_commute)}
+ ultimately show ?thesis by blast
+ qed
+ from fr_eq2[OF lp, of vs bs x] show ?thesis
+ unfolding eq0 eq1 eq2 by blast
+qed
+
+definition
+"ferrack2 p \<equiv> let q = simpfm p ; mp = minusinf q ; pp = plusinf q
+ in if (mp = T \<or> pp = T) then T
+ else (let U = remdps (uset q)
+ in decr0 (list_disj [mp, pp, simpfm (subst0 (CP 0\<^sub>p) q), evaldjf (\<lambda>(c,t). msubst2 q (c *\<^sub>p C (-2, 1)) t) U,
+ evaldjf (\<lambda>((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) \<and> ((Ifm vs bs (ferrack2 p)) = (Ifm vs bs (E p)))"
+ (is "_ \<and> (?rhs = ?lhs)")
+proof-
+ let ?J = "\<lambda> x p. Ifm vs (x#bs) p"
+ let ?N = "\<lambda> t. Ipoly vs t"
+ let ?Nt = "\<lambda>x t. Itm vs (x#bs) t"
+ let ?q = "simpfm p"
+ let ?qz = "subst0 (CP 0\<^sub>p) ?q"
+ let ?U = "remdps(uset ?q)"
+ let ?Up = "alluopairs ?U"
+ let ?mp = "minusinf ?q"
+ let ?pp = "plusinf ?q"
+ let ?I = "\<lambda>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: "\<forall>(c, s)\<in>set ?U. isnpoly c \<and> c \<noteq> 0\<^sub>p \<and> tmbound0 s \<and> allpolys isnpoly s"
+ by simp
+ have bnd0: "\<forall>x \<in> set ?U. bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) x)"
+ proof-
+ {fix c t assume ct: "(c,t) \<in> set ?U"
+ hence tnb: "tmbound0 t" using U_l by blast
+ from msubst2_nb[OF lq tnb]
+ have "bound0 ((\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) (c,t))" by simp}
+ thus ?thesis by auto
+ qed
+ have bnd1: "\<forall>x \<in> set ?Up. bound0 ((\<lambda>((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) x)"
+ proof-
+ {fix b a d c assume badc: "((b,a),(d,c)) \<in> set ?Up"
+ from badc U_l alluopairs_set1[of ?U]
+ have nb: "tmbound0 (Add (Mul d a) (Mul b c))" by auto
+ from msubst2_nb[OF lq nb] have "bound0 ((\<lambda>((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) ((b,a),(d,c)))" by simp}
+ thus ?thesis by auto
+ qed
+ have stupid: "bound0 F" by simp
+ let ?R = "list_disj [?mp, ?pp, simpfm (subst0 (CP 0\<^sub>p) ?q), evaldjf (\<lambda>(c,t). msubst2 ?q (c *\<^sub>p C (-2, 1)) t) ?U,
+ evaldjf (\<lambda>((b,a),(d,c)). msubst2 ?q (C (-2, 1) *\<^sub>p b*\<^sub>p d) (Add (Mul d a) (Mul b c))) (alluopairs ?U)]"
+ from subst0_nb[of "CP 0\<^sub>p" ?q] q_qf evaldjf_bound0[OF bnd1] evaldjf_bound0[OF bnd0] mp_nb pp_nb stupid
+ have nb: "bound0 ?R "
+ by (simp add: list_disj_def disj_nb0 simpfm_bound0)
+ let ?s = "\<lambda>((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) \<in> set ?U" and dcU: "(d,c) \<in> 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 \<longleftrightarrow> ?rhs")
+ proof
+ assume H: ?lhs
+ hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
+ by (auto simp add: msubst2_def lt[OF stupid(3)] lt[OF stupid(1)] mult_less_0_iff zero_less_mult_iff)
+ from msubst2[OF lq norm2(1) z(1), of x bs]
+ msubst2[OF lq norm2(2) z(2), of x bs] H
+ show ?rhs by (simp add: ring_simps)
+ next
+ assume H: ?rhs
+ hence z: "\<lparr>C (-2, 1) *\<^sub>p b *\<^sub>p d\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0" "\<lparr>C (-2, 1) *\<^sub>p d *\<^sub>p b\<rparr>\<^sub>p\<^bsup>vs\<^esup> \<noteq> 0"
+ by (auto simp add: msubst2_def lt[OF stupid(4)] lt[OF stupid(2)] mult_less_0_iff zero_less_mult_iff)
+ from msubst2[OF lq norm2(1) z(1), of x bs]
+ msubst2[OF lq norm2(2) z(2), of x bs] H
+ show ?lhs by (simp add: ring_simps)
+ qed}
+ hence th0: "\<forall>x \<in> set ?U. \<forall>y \<in> set ?U. ?I (?s (x, y)) \<longleftrightarrow> ?I (?s (y, x))"
+ by clarsimp
+
+ have "?lhs \<longleftrightarrow> (\<exists>x. Ifm vs (x#bs) ?q)" by simp
+ also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists>(b, a)\<in>set ?U. \<exists>(d, c)\<in>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 "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists> x\<in>set ?U. \<exists> y \<in>set ?U. ?I (?s (x,y)))"
+ by (simp add: split_def)
+ also have "\<dots> \<longleftrightarrow> ?I ?mp \<or> ?I ?pp \<or> ?I (subst0 (CP 0\<^sub>p) ?q) \<or> (\<exists>(n,t) \<in> set ?U. ?I (msubst2 ?q (n *\<^sub>p C (-2, 1)) t)) \<or> (\<exists> (x,y) \<in> set ?Up. ?I (?s (x,y)))"
+ using alluopairs_ex[OF th0] by simp
+ also have "\<dots> \<longleftrightarrow> ?I ?R"
+ by (simp add: list_disj_def evaldjf_ex split_def)
+ also have "\<dots> \<longleftrightarrow> ?rhs"
+ unfolding ferrack2_def
+ apply (cases "?mp = T")
+ apply (simp add: list_disj_def)
+ apply (cases "?pp = T")
+ apply (simp add: list_disj_def)
+ by (simp_all add: Let_def decr0[OF nb])
+ finally show ?thesis using decr0_qf[OF nb]
+ by (simp add: ferrack2_def Let_def)
+qed
+
+lemma frpar2: "qfree (frpar2 p) \<and> (Ifm vs bs (frpar2 p) \<longleftrightarrow> Ifm vs bs p)"
+proof-
+ from ferrack2 have th: "\<forall>bs p. qfree p \<longrightarrow> qfree (ferrack2 p) \<and> Ifm vs bs (ferrack2 p) = Ifm vs bs (E p)" by blast
+ from qelim[OF th, of "prep p" bs]
+show ?thesis unfolding frpar2_def by (auto simp add: prep)
+qed
+
+code_module FRPar
+ contains
+ frpar = "frpar"
+ frpar2 = "frpar2"
+ test = "%x . frpar (E(Lt (Mul 1\<^sub>p (Bound 0))))"
+
+ML{*
+
+structure ReflectedFRPar =
+struct
+
+val bT = HOLogic.boolT;
+fun num rT x = HOLogic.mk_number rT x;
+fun rrelT rT = [rT,rT] ---> rT;
+fun rrT rT = [rT, rT] ---> bT;
+fun divt rT = Const(@{const_name "HOL.divide"},rrelT rT);
+fun timest rT = Const(@{const_name "HOL.times"},rrelT rT);
+fun plust rT = Const(@{const_name "HOL.plus"},rrelT rT);
+fun minust rT = Const(@{const_name "HOL.minus"},rrelT rT);
+fun uminust rT = Const(@{const_name "HOL.uminus"}, rT --> rT);
+fun powt rT = Const(@{const_name "power"}, [rT,@{typ "nat"}] ---> rT);
+val brT = [bT, bT] ---> bT;
+val nott = @{term "Not"};
+val conjt = @{term "op &"};
+val disjt = @{term "op |"};
+val impt = @{term "op -->"};
+val ifft = @{term "op = :: bool => _"}
+fun llt rT = Const(@{const_name "HOL.less"},rrT rT);
+fun lle rT = Const(@{const_name "HOL.less"},rrT rT);
+fun eqt rT = Const("op =",rrT rT);
+fun rz rT = Const(@{const_name "HOL.zero"},rT);
+
+fun dest_nat t = case t of
+ Const ("Suc",_)$t' => 1 + dest_nat t'
+| _ => (snd o HOLogic.dest_number) t;
+
+fun num_of_term m t =
+ case t of
+ Const(@{const_name "uminus"},_)$t => FRPar.Neg (num_of_term m t)
+ | Const(@{const_name "HOL.plus"},_)$a$b => FRPar.Add (num_of_term m a, num_of_term m b)
+ | Const(@{const_name "HOL.minus"},_)$a$b => FRPar.Sub (num_of_term m a, num_of_term m b)
+ | Const(@{const_name "HOL.times"},_)$a$b => FRPar.Mul (num_of_term m a, num_of_term m b)
+ | Const(@{const_name "power"},_)$a$n => FRPar.Pw (num_of_term m a, dest_nat n)
+ | Const(@{const_name "HOL.divide"},_)$a$b => FRPar.C (HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd)
+ | _ => (FRPar.C (HOLogic.dest_number t |> snd,1)
+ handle TERM _ => FRPar.Bound (AList.lookup (op aconv) m t |> valOf));
+
+fun tm_of_term m m' t =
+ case t of
+ Const(@{const_name "uminus"},_)$t => FRPar.tm_Neg (tm_of_term m m' t)
+ | Const(@{const_name "HOL.plus"},_)$a$b => FRPar.tm_Add (tm_of_term m m' a, tm_of_term m m' b)
+ | Const(@{const_name "HOL.minus"},_)$a$b => FRPar.tm_Sub (tm_of_term m m' a, tm_of_term m m' b)
+ | Const(@{const_name "HOL.times"},_)$a$b => FRPar.tm_Mul (num_of_term m' a, tm_of_term m m' b)
+ | _ => (FRPar.CP (num_of_term m' t)
+ handle TERM _ => FRPar.tm_Bound (AList.lookup (op aconv) m t |> valOf)
+ | Option => FRPar.tm_Bound (AList.lookup (op aconv) m t |> valOf));
+
+fun term_of_num T m t =
+ case t of
+ FRPar.C (a,b) => (if b = 1 then num T a else if b=0 then (rz T)
+ else (divt T) $ num T a $ num T b)
+| FRPar.Bound i => AList.lookup (op = : int*int -> bool) m i |> valOf
+| FRPar.Add(a,b) => (plust T)$(term_of_num T m a)$(term_of_num T m b)
+| FRPar.Mul(a,b) => (timest T)$(term_of_num T m a)$(term_of_num T m b)
+| FRPar.Sub(a,b) => (minust T)$(term_of_num T m a)$(term_of_num T m b)
+| FRPar.Neg a => (uminust T)$(term_of_num T m a)
+| FRPar.Pw(a,n) => (powt T)$(term_of_num T m t)$(HOLogic.mk_number HOLogic.natT n)
+| FRPar.CN(c,n,p) => term_of_num T m (FRPar.Add(c,FRPar.Mul(FRPar.Bound n, p)))
+| _ => error "term_of_num: Unknown term";
+
+fun term_of_tm T m m' t =
+ case t of
+ FRPar.CP p => term_of_num T m' p
+| FRPar.tm_Bound i => AList.lookup (op = : int*int -> bool) m i |> valOf
+| FRPar.tm_Add(a,b) => (plust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b)
+| FRPar.tm_Mul(a,b) => (timest T)$(term_of_num T m' a)$(term_of_tm T m m' b)
+| FRPar.tm_Sub(a,b) => (minust T)$(term_of_tm T m m' a)$(term_of_tm T m m' b)
+| FRPar.tm_Neg a => (uminust T)$(term_of_tm T m m' a)
+| FRPar.CNP(n,c,p) => term_of_tm T m m' (FRPar.tm_Add(FRPar.tm_Mul(c, FRPar.tm_Bound n), p))
+| _ => error "term_of_tm: Unknown term";
+
+fun fm_of_term m m' fm =
+ case fm of
+ Const("True",_) => FRPar.T
+ | Const("False",_) => FRPar.F
+ | Const("Not",_)$p => FRPar.NOT (fm_of_term m m' p)
+ | Const("op &",_)$p$q => FRPar.And(fm_of_term m m' p, fm_of_term m m' q)
+ | Const("op |",_)$p$q => FRPar.Or(fm_of_term m m' p, fm_of_term m m' q)
+ | Const("op -->",_)$p$q => FRPar.Imp(fm_of_term m m' p, fm_of_term m m' q)
+ | Const("op =",ty)$p$q =>
+ if domain_type ty = bT then FRPar.Iff(fm_of_term m m' p, fm_of_term m m' q)
+ else FRPar.Eq (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+ | Const(@{const_name "HOL.less"},_)$p$q =>
+ FRPar.Lt (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+ | Const(@{const_name "HOL.less_eq"},_)$p$q =>
+ FRPar.Le (FRPar.tm_Sub(tm_of_term m m' p, tm_of_term m m' q))
+ | Const("Ex",_)$Abs(xn,xT,p) =>
+ let val (xn', p') = variant_abs (xn,xT,p)
+ val x = Free(xn',xT)
+ fun incr i = i + 1
+ val m0 = (x,0):: (map (apsnd incr) m)
+ in FRPar.E (fm_of_term m0 m' p') end
+ | Const("All",_)$Abs(xn,xT,p) =>
+ let val (xn', p') = variant_abs (xn,xT,p)
+ val x = Free(xn',xT)
+ fun incr i = i + 1
+ val m0 = (x,0):: (map (apsnd incr) m)
+ in FRPar.A (fm_of_term m0 m' p') end
+ | _ => error "fm_of_term";
+
+
+fun term_of_fm T m m' t =
+ case t of
+ FRPar.T => Const("True",bT)
+ | FRPar.F => Const("False",bT)
+ | FRPar.NOT p => nott $ (term_of_fm T m m' p)
+ | FRPar.And (p,q) => conjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+ | FRPar.Or (p,q) => disjt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+ | FRPar.Imp (p,q) => impt $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+ | FRPar.Iff (p,q) => ifft $ (term_of_fm T m m' p) $ (term_of_fm T m m' q)
+ | FRPar.Lt p => (llt T) $ (term_of_tm T m m' p) $ (rz T)
+ | FRPar.Le p => (lle T) $ (term_of_tm T m m' p) $ (rz T)
+ | FRPar.Eq p => (eqt T) $ (term_of_tm T m m' p) $ (rz T)
+ | FRPar.NEq p => nott $ ((eqt T) $ (term_of_tm T m m' p) $ (rz T))
+ | _ => error "term_of_fm: quantifiers!!!!???";
+
+fun frpar_oracle (T,m, m', fm) =
+ let
+ val t = HOLogic.dest_Trueprop fm
+ val im = 0 upto (length m - 1)
+ val im' = 0 upto (length m' - 1)
+ in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm T (im ~~ m) (im' ~~ m')
+ (FRPar.frpar (fm_of_term (m ~~ im) (m' ~~ im') t))))
+ end;
+
+fun frpar_oracle2 (T,m, m', fm) =
+ let
+ val t = HOLogic.dest_Trueprop fm
+ val im = 0 upto (length m - 1)
+ val im' = 0 upto (length m' - 1)
+ in HOLogic.mk_Trueprop (HOLogic.mk_eq(t, term_of_fm T (im ~~ m) (im' ~~ m')
+ (FRPar.frpar2 (fm_of_term (m ~~ im) (m' ~~ im') t))))
+ end;
+
+end;
+
+
+*}
+
+oracle frpar_oracle = {* fn (ty, ts, ts', ct) =>
+ let
+ val thy = Thm.theory_of_cterm ct
+ in cterm_of thy (ReflectedFRPar.frpar_oracle (ty,ts, ts', term_of ct))
+ end *}
+
+oracle frpar_oracle2 = {* fn (ty, ts, ts', ct) =>
+ let
+ val thy = Thm.theory_of_cterm ct
+ in cterm_of thy (ReflectedFRPar.frpar_oracle2 (ty,ts, ts', term_of ct))
+ end *}
+
+ML{*
+structure FRParTac =
+struct
+
+fun frpar_tac T ps ctxt i =
+ (ObjectLogic.full_atomize_tac i)
+ THEN (fn st =>
+ let
+ val g = List.nth (cprems_of st, i - 1)
+ val thy = ProofContext.theory_of ctxt
+ val fs = subtract (op aconv) (map Free (Term.add_frees (term_of g) [])) ps
+ val th = frpar_oracle (T, fs,ps, (* Pattern.eta_long [] *)g)
+ in rtac (th RS iffD2) i st end);
+
+fun frpar2_tac T ps ctxt i =
+ (ObjectLogic.full_atomize_tac i)
+ THEN (fn st =>
+ let
+ val g = List.nth (cprems_of st, i - 1)
+ val thy = ProofContext.theory_of ctxt
+ val fs = subtract (op aconv) (map Free (Term.add_frees (term_of g) [])) ps
+ val th = frpar_oracle2 (T, fs,ps, (* Pattern.eta_long [] *)g)
+ in rtac (th RS iffD2) i st end);
+
+end;
+
+*}
+
+method_setup frpar = {*
+let
+ fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
+ fun simple_keyword k = Scan.lift (Args.$$$ k) >> K ()
+ val parsN = "pars"
+ val typN = "type"
+ val any_keyword = keyword parsN || keyword typN
+ val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat
+ val cterms = thms >> map Drule.dest_term;
+ val terms = Scan.repeat (Scan.unless any_keyword Args.term)
+ val typ = Scan.unless any_keyword Args.typ
+in
+ (keyword typN |-- typ) -- (keyword parsN |-- terms) >>
+ (fn (T,ps) => fn ctxt => SIMPLE_METHOD' (FRParTac.frpar_tac T ps ctxt))
+end
+*} "Parametric QE for linear Arithmetic over fields, Version 1"
+
+method_setup frpar2 = {*
+let
+ fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ()
+ fun simple_keyword k = Scan.lift (Args.$$$ k) >> K ()
+ val parsN = "pars"
+ val typN = "type"
+ val any_keyword = keyword parsN || keyword typN
+ val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat
+ val cterms = thms >> map Drule.dest_term;
+ val terms = Scan.repeat (Scan.unless any_keyword Args.term)
+ val typ = Scan.unless any_keyword Args.typ
+in
+ (keyword typN |-- typ) -- (keyword parsN |-- terms) >>
+ (fn (T,ps) => fn ctxt => SIMPLE_METHOD' (FRParTac.frpar2_tac T ps ctxt))
+end
+*} "Parametric QE for linear Arithmetic over fields, Version 2"
+
+
+lemma "\<exists>(x::'a::{division_by_zero,ordered_field,number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
+ apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "y::'a::{division_by_zero,ordered_field,number_ring}")
+ apply (simp add: ring_simps)
+ apply (rule spec[where x=y])
+ apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "z::'a::{division_by_zero,ordered_field,number_ring}")
+ by simp
+
+text{* Collins/Jones Problem *}
+(*
+lemma "\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
+proof-
+ have "(\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
+by (simp add: ring_simps)
+have "?rhs"
+
+ apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "a::'a::{division_by_zero,ordered_field,number_ring}" "b::'a::{division_by_zero,ordered_field,number_ring}")
+ apply (simp add: ring_simps)
+oops
+*)
+(*
+lemma "ALL (x::'a::{division_by_zero,ordered_field,number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
+apply (frpar type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "t::'a::{division_by_zero,ordered_field,number_ring}")
+oops
+*)
+
+lemma "\<exists>(x::'a::{division_by_zero,ordered_field,number_ring}). y \<noteq> -1 \<longrightarrow> (y + 1)*x < 0"
+ apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "y::'a::{division_by_zero,ordered_field,number_ring}")
+ apply (simp add: ring_simps)
+ apply (rule spec[where x=y])
+ apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "z::'a::{division_by_zero,ordered_field,number_ring}")
+ by simp
+
+text{* Collins/Jones Problem *}
+
+(*
+lemma "\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0"
+proof-
+ have "(\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < (2 - 3*r) *(a^2 + b^2) + (2*a)*r \<and> (2 - 3*r) *(a^2 + b^2) + 4*a*r - 2*a - r < 0) \<longleftrightarrow> (\<exists>(r::'a::{division_by_zero,ordered_field,number_ring}). 0 < r \<and> r < 1 \<and> 0 < 2 *(a^2 + b^2) - (3*(a^2 + b^2)) * r + (2*a)*r \<and> 2*(a^2 + b^2) - (3*(a^2 + b^2) - 4*a + 1)*r - 2*a < 0)" (is "?lhs \<longleftrightarrow> ?rhs")
+by (simp add: ring_simps)
+have "?rhs"
+ apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "a::'a::{division_by_zero,ordered_field,number_ring}" "b::'a::{division_by_zero,ordered_field,number_ring}")
+ apply simp
+oops
+*)
+
+(*
+lemma "ALL (x::'a::{division_by_zero,ordered_field,number_ring}) y. (1 - t)*x \<le> (1+t)*y \<and> (1 - t)*y \<le> (1+t)*x --> 0 \<le> y"
+apply (frpar2 type: "'a::{division_by_zero,ordered_field,number_ring}" pars: "t::'a::{division_by_zero,ordered_field,number_ring}")
+apply (simp add: field_simps linorder_neq_iff[symmetric])
+apply ferrack
+oops
+*)
+end
\ No newline at end of file
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Sun Oct 25 08:57:55 2009 +0100
@@ -0,0 +1,1743 @@
+(* Title: HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy
+ Author: Amine Chaieb
+*)
+
+header {* Implementation and verification of mutivariate polynomials Library *}
+
+
+theory Reflected_Multivariate_Polynomial
+imports Parity Abstract_Rat Efficient_Nat List Polynomial_List
+begin
+
+ (* Impelementation *)
+
+subsection{* Datatype of polynomial expressions *}
+
+datatype poly = C Num| Bound nat| Add poly poly|Sub poly poly
+ | Mul poly poly| Neg poly| Pw poly nat| CN poly nat poly
+
+ML{* @{term "Add"}*}
+syntax "_poly0" :: "poly" ("0\<^sub>p")
+translations "0\<^sub>p" \<rightleftharpoons> "C (0\<^sub>N)"
+syntax "_poly" :: "int \<Rightarrow> poly" ("_\<^sub>p")
+translations "i\<^sub>p" \<rightleftharpoons> "C (i\<^sub>N)"
+
+subsection{* Boundedness, substitution and all that *}
+consts polysize:: "poly \<Rightarrow> nat"
+primrec
+ "polysize (C c) = 1"
+ "polysize (Bound n) = 1"
+ "polysize (Neg p) = 1 + polysize p"
+ "polysize (Add p q) = 1 + polysize p + polysize q"
+ "polysize (Sub p q) = 1 + polysize p + polysize q"
+ "polysize (Mul p q) = 1 + polysize p + polysize q"
+ "polysize (Pw p n) = 1 + polysize p"
+ "polysize (CN c n p) = 4 + polysize c + polysize p"
+
+consts
+ polybound0:: "poly \<Rightarrow> bool" (* a poly is INDEPENDENT of Bound 0 *)
+ polysubst0:: "poly \<Rightarrow> poly \<Rightarrow> poly" (* substitute a poly into a poly for Bound 0 *)
+primrec
+ "polybound0 (C c) = True"
+ "polybound0 (Bound n) = (n>0)"
+ "polybound0 (Neg a) = polybound0 a"
+ "polybound0 (Add a b) = (polybound0 a \<and> polybound0 b)"
+ "polybound0 (Sub a b) = (polybound0 a \<and> polybound0 b)"
+ "polybound0 (Mul a b) = (polybound0 a \<and> polybound0 b)"
+ "polybound0 (Pw p n) = (polybound0 p)"
+ "polybound0 (CN c n p) = (n \<noteq> 0 \<and> polybound0 c \<and> polybound0 p)"
+primrec
+ "polysubst0 t (C c) = (C c)"
+ "polysubst0 t (Bound n) = (if n=0 then t else Bound n)"
+ "polysubst0 t (Neg a) = Neg (polysubst0 t a)"
+ "polysubst0 t (Add a b) = Add (polysubst0 t a) (polysubst0 t b)"
+ "polysubst0 t (Sub a b) = Sub (polysubst0 t a) (polysubst0 t b)"
+ "polysubst0 t (Mul a b) = Mul (polysubst0 t a) (polysubst0 t b)"
+ "polysubst0 t (Pw p n) = Pw (polysubst0 t p) n"
+ "polysubst0 t (CN c n p) = (if n=0 then Add (polysubst0 t c) (Mul t (polysubst0 t p))
+ else CN (polysubst0 t c) n (polysubst0 t p))"
+
+consts
+ decrpoly:: "poly \<Rightarrow> poly"
+recdef decrpoly "measure polysize"
+ "decrpoly (Bound n) = Bound (n - 1)"
+ "decrpoly (Neg a) = Neg (decrpoly a)"
+ "decrpoly (Add a b) = Add (decrpoly a) (decrpoly b)"
+ "decrpoly (Sub a b) = Sub (decrpoly a) (decrpoly b)"
+ "decrpoly (Mul a b) = Mul (decrpoly a) (decrpoly b)"
+ "decrpoly (Pw p n) = Pw (decrpoly p) n"
+ "decrpoly (CN c n p) = CN (decrpoly c) (n - 1) (decrpoly p)"
+ "decrpoly a = a"
+
+subsection{* Degrees and heads and coefficients *}
+
+consts degree:: "poly \<Rightarrow> nat"
+recdef degree "measure size"
+ "degree (CN c 0 p) = 1 + degree p"
+ "degree p = 0"
+consts head:: "poly \<Rightarrow> poly"
+
+recdef head "measure size"
+ "head (CN c 0 p) = head p"
+ "head p = p"
+ (* More general notions of degree and head *)
+consts degreen:: "poly \<Rightarrow> nat \<Rightarrow> nat"
+recdef degreen "measure size"
+ "degreen (CN c n p) = (\<lambda>m. if n=m then 1 + degreen p n else 0)"
+ "degreen p = (\<lambda>m. 0)"
+
+consts headn:: "poly \<Rightarrow> nat \<Rightarrow> poly"
+recdef headn "measure size"
+ "headn (CN c n p) = (\<lambda>m. if n \<le> m then headn p m else CN c n p)"
+ "headn p = (\<lambda>m. p)"
+
+consts coefficients:: "poly \<Rightarrow> poly list"
+recdef coefficients "measure size"
+ "coefficients (CN c 0 p) = c#(coefficients p)"
+ "coefficients p = [p]"
+
+consts isconstant:: "poly \<Rightarrow> bool"
+recdef isconstant "measure size"
+ "isconstant (CN c 0 p) = False"
+ "isconstant p = True"
+
+consts behead:: "poly \<Rightarrow> poly"
+recdef behead "measure size"
+ "behead (CN c 0 p) = (let p' = behead p in if p' = 0\<^sub>p then c else CN c 0 p')"
+ "behead p = 0\<^sub>p"
+
+consts headconst:: "poly \<Rightarrow> Num"
+recdef headconst "measure size"
+ "headconst (CN c n p) = headconst p"
+ "headconst (C n) = n"
+
+subsection{* Operations for normalization *}
+consts
+ polyadd :: "poly\<times>poly \<Rightarrow> poly"
+ polyneg :: "poly \<Rightarrow> poly" ("~\<^sub>p")
+ polysub :: "poly\<times>poly \<Rightarrow> poly"
+ polymul :: "poly\<times>poly \<Rightarrow> poly"
+ polypow :: "nat \<Rightarrow> poly \<Rightarrow> poly"
+syntax "_polyadd" :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "+\<^sub>p" 60)
+translations "a +\<^sub>p b" \<rightleftharpoons> "polyadd (a,b)"
+syntax "_polymul" :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "*\<^sub>p" 60)
+translations "a *\<^sub>p b" \<rightleftharpoons> "polymul (a,b)"
+syntax "_polysub" :: "poly \<Rightarrow> poly \<Rightarrow> poly" (infixl "-\<^sub>p" 60)
+translations "a -\<^sub>p b" \<rightleftharpoons> "polysub (a,b)"
+syntax "_polypow" :: "nat \<Rightarrow> poly \<Rightarrow> poly" (infixl "^\<^sub>p" 60)
+translations "a ^\<^sub>p k" \<rightleftharpoons> "polypow k a"
+
+recdef polyadd "measure (\<lambda> (a,b). polysize a + polysize b)"
+ "polyadd (C c, C c') = C (c+\<^sub>Nc')"
+ "polyadd (C c, CN c' n' p') = CN (polyadd (C c, c')) n' p'"
+ "polyadd (CN c n p, C c') = CN (polyadd (c, C c')) n p"
+stupid: "polyadd (CN c n p, CN c' n' p') =
+ (if n < n' then CN (polyadd(c,CN c' n' p')) n p
+ else if n'<n then CN (polyadd(CN c n p, c')) n' p'
+ else (let cc' = polyadd (c,c') ;
+ pp' = polyadd (p,p')
+ in (if pp' = 0\<^sub>p then cc' else CN cc' n pp')))"
+ "polyadd (a, b) = Add a b"
+(hints recdef_simp add: Let_def measure_def split_def inv_image_def)
+
+(*
+declare stupid [simp del, code del]
+
+lemma [simp,code]: "polyadd (CN c n p, CN c' n' p') =
+ (if n < n' then CN (polyadd(c,CN c' n' p')) n p
+ else if n'<n then CN (polyadd(CN c n p, c')) n' p'
+ else (let cc' = polyadd (c,c') ;
+ pp' = polyadd (p,p')
+ in (if pp' = 0\<^sub>p then cc' else CN cc' n pp')))"
+ by (simp add: Let_def stupid)
+*)
+
+recdef polyneg "measure size"
+ "polyneg (C c) = C (~\<^sub>N c)"
+ "polyneg (CN c n p) = CN (polyneg c) n (polyneg p)"
+ "polyneg a = Neg a"
+
+defs polysub_def[code]: "polysub \<equiv> \<lambda> (p,q). polyadd (p,polyneg q)"
+
+recdef polymul "measure (\<lambda>(a,b). size a + size b)"
+ "polymul(C c, C c') = C (c*\<^sub>Nc')"
+ "polymul(C c, CN c' n' p') =
+ (if c = 0\<^sub>N then 0\<^sub>p else CN (polymul(C c,c')) n' (polymul(C c, p')))"
+ "polymul(CN c n p, C c') =
+ (if c' = 0\<^sub>N then 0\<^sub>p else CN (polymul(c,C c')) n (polymul(p, C c')))"
+ "polymul(CN c n p, CN c' n' p') =
+ (if n<n' then CN (polymul(c,CN c' n' p')) n (polymul(p,CN c' n' p'))
+ else if n' < n
+ then CN (polymul(CN c n p,c')) n' (polymul(CN c n p,p'))
+ else polyadd(polymul(CN c n p, c'),CN 0\<^sub>p n' (polymul(CN c n p, p'))))"
+ "polymul (a,b) = Mul a b"
+recdef polypow "measure id"
+ "polypow 0 = (\<lambda>p. 1\<^sub>p)"
+ "polypow n = (\<lambda>p. let q = polypow (n div 2) p ; d = polymul(q,q) in
+ if even n then d else polymul(p,d))"
+
+consts polynate :: "poly \<Rightarrow> poly"
+recdef polynate "measure polysize"
+ "polynate (Bound n) = CN 0\<^sub>p n 1\<^sub>p"
+ "polynate (Add p q) = (polynate p +\<^sub>p polynate q)"
+ "polynate (Sub p q) = (polynate p -\<^sub>p polynate q)"
+ "polynate (Mul p q) = (polynate p *\<^sub>p polynate q)"
+ "polynate (Neg p) = (~\<^sub>p (polynate p))"
+ "polynate (Pw p n) = ((polynate p) ^\<^sub>p n)"
+ "polynate (CN c n p) = polynate (Add c (Mul (Bound n) p))"
+ "polynate (C c) = C (normNum c)"
+
+fun poly_cmul :: "Num \<Rightarrow> poly \<Rightarrow> poly" where
+ "poly_cmul y (C x) = C (y *\<^sub>N x)"
+| "poly_cmul y (CN c n p) = CN (poly_cmul y c) n (poly_cmul y p)"
+| "poly_cmul y p = C y *\<^sub>p p"
+
+constdefs monic:: "poly \<Rightarrow> (poly \<times> bool)"
+ "monic p \<equiv> (let h = headconst p in if h = 0\<^sub>N then (p,False) else ((C (Ninv h)) *\<^sub>p p, 0>\<^sub>N h))"
+
+subsection{* Pseudo-division *}
+
+constdefs shift1:: "poly \<Rightarrow> poly"
+ "shift1 p \<equiv> CN 0\<^sub>p 0 p"
+consts funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
+
+primrec
+ "funpow 0 f x = x"
+ "funpow (Suc n) f x = funpow n f (f x)"
+function (tailrec) polydivide_aux :: "(poly \<times> nat \<times> poly \<times> nat \<times> poly) \<Rightarrow> (nat \<times> poly)"
+ where
+ "polydivide_aux (a,n,p,k,s) =
+ (if s = 0\<^sub>p then (k,s)
+ else (let b = head s; m = degree s in
+ (if m < n then (k,s) else
+ (let p'= funpow (m - n) shift1 p in
+ (if a = b then polydivide_aux (a,n,p,k,s -\<^sub>p p')
+ else polydivide_aux (a,n,p,Suc k, (a *\<^sub>p s) -\<^sub>p (b *\<^sub>p p')))))))"
+ by pat_completeness auto
+
+
+constdefs polydivide:: "poly \<Rightarrow> poly \<Rightarrow> (nat \<times> poly)"
+ "polydivide s p \<equiv> polydivide_aux (head p,degree p,p,0, s)"
+
+fun poly_deriv_aux :: "nat \<Rightarrow> poly \<Rightarrow> poly" where
+ "poly_deriv_aux n (CN c 0 p) = CN (poly_cmul ((int n)\<^sub>N) c) 0 (poly_deriv_aux (n + 1) p)"
+| "poly_deriv_aux n p = poly_cmul ((int n)\<^sub>N) p"
+
+fun poly_deriv :: "poly \<Rightarrow> poly" where
+ "poly_deriv (CN c 0 p) = poly_deriv_aux 1 p"
+| "poly_deriv p = 0\<^sub>p"
+
+ (* Verification *)
+lemma nth_pos2[simp]: "0 < n \<Longrightarrow> (x#xs) ! n = xs ! (n - 1)"
+using Nat.gr0_conv_Suc
+by clarsimp
+
+subsection{* Semantics of the polynomial representation *}
+
+consts Ipoly :: "'a list \<Rightarrow> poly \<Rightarrow> 'a::{ring_char_0,power,division_by_zero,field}"
+primrec
+ "Ipoly bs (C c) = INum c"
+ "Ipoly bs (Bound n) = bs!n"
+ "Ipoly bs (Neg a) = - Ipoly bs a"
+ "Ipoly bs (Add a b) = Ipoly bs a + Ipoly bs b"
+ "Ipoly bs (Sub a b) = Ipoly bs a - Ipoly bs b"
+ "Ipoly bs (Mul a b) = Ipoly bs a * Ipoly bs b"
+ "Ipoly bs (Pw t n) = (Ipoly bs t) ^ n"
+ "Ipoly bs (CN c n p) = (Ipoly bs c) + (bs!n)*(Ipoly bs p)"
+syntax "_Ipoly" :: "poly \<Rightarrow> 'a list \<Rightarrow>'a::{ring_char_0,power,division_by_zero,field}" ("\<lparr>_\<rparr>\<^sub>p\<^bsup>_\<^esup>")
+translations "\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup>" \<rightleftharpoons> "Ipoly bs p"
+
+lemma Ipoly_CInt: "Ipoly bs (C (i,1)) = of_int i"
+ by (simp add: INum_def)
+lemma Ipoly_CRat: "Ipoly bs (C (i, j)) = of_int i / of_int j"
+ by (simp add: INum_def)
+
+lemmas RIpoly_eqs = Ipoly.simps(2-7) Ipoly_CInt Ipoly_CRat
+
+subsection {* Normal form and normalization *}
+
+consts isnpolyh:: "poly \<Rightarrow> nat \<Rightarrow> bool"
+recdef isnpolyh "measure size"
+ "isnpolyh (C c) = (\<lambda>k. isnormNum c)"
+ "isnpolyh (CN c n p) = (\<lambda>k. n\<ge> k \<and> (isnpolyh c (Suc n)) \<and> (isnpolyh p n) \<and> (p \<noteq> 0\<^sub>p))"
+ "isnpolyh p = (\<lambda>k. False)"
+
+lemma isnpolyh_mono: "\<lbrakk>n' \<le> n ; isnpolyh p n\<rbrakk> \<Longrightarrow> isnpolyh p n'"
+by (induct p rule: isnpolyh.induct, auto)
+
+constdefs isnpoly:: "poly \<Rightarrow> bool"
+ "isnpoly p \<equiv> isnpolyh p 0"
+
+text{* polyadd preserves normal forms *}
+
+lemma polyadd_normh: "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1\<rbrakk>
+ \<Longrightarrow> isnpolyh (polyadd(p,q)) (min n0 n1)"
+proof(induct p q arbitrary: n0 n1 rule: polyadd.induct)
+ case (2 a b c' n' p' n0 n1)
+ from prems have th1: "isnpolyh (C (a,b)) (Suc n')" by simp
+ from prems(3) have th2: "isnpolyh c' (Suc n')" and nplen1: "n' \<ge> n1" by simp_all
+ with isnpolyh_mono have cp: "isnpolyh c' (Suc n')" by simp
+ with prems(1)[OF th1 th2] have th3:"isnpolyh (C (a,b) +\<^sub>p c') (Suc n')" by simp
+ from nplen1 have n01len1: "min n0 n1 \<le> n'" by simp
+ thus ?case using prems th3 by simp
+next
+ case (3 c' n' p' a b n1 n0)
+ from prems have th1: "isnpolyh (C (a,b)) (Suc n')" by simp
+ from prems(2) have th2: "isnpolyh c' (Suc n')" and nplen1: "n' \<ge> n1" by simp_all
+ with isnpolyh_mono have cp: "isnpolyh c' (Suc n')" by simp
+ with prems(1)[OF th2 th1] have th3:"isnpolyh (c' +\<^sub>p C (a,b)) (Suc n')" by simp
+ from nplen1 have n01len1: "min n0 n1 \<le> n'" by simp
+ thus ?case using prems th3 by simp
+next
+ case (4 c n p c' n' p' n0 n1)
+ hence nc: "isnpolyh c (Suc n)" and np: "isnpolyh p n" by simp_all
+ from prems have nc': "isnpolyh c' (Suc n')" and np': "isnpolyh p' n'" by simp_all
+ from prems have ngen0: "n \<ge> n0" by simp
+ from prems have n'gen1: "n' \<ge> n1" by simp
+ have "n < n' \<or> n' < n \<or> n = n'" by auto
+ moreover {assume eq: "n = n'" hence eq': "\<not> n' < n \<and> \<not> n < n'" by simp
+ with prems(2)[rule_format, OF eq' nc nc']
+ have ncc':"isnpolyh (c +\<^sub>p c') (Suc n)" by auto
+ hence ncc'n01: "isnpolyh (c +\<^sub>p c') (min n0 n1)"
+ using isnpolyh_mono[where n'="min n0 n1" and n="Suc n"] ngen0 n'gen1 by auto
+ from eq prems(1)[rule_format, OF eq' np np'] have npp': "isnpolyh (p +\<^sub>p p') n" by simp
+ have minle: "min n0 n1 \<le> n'" using ngen0 n'gen1 eq by simp
+ from minle npp' ncc'n01 prems ngen0 n'gen1 ncc' have ?case by (simp add: Let_def)}
+ moreover {assume lt: "n < n'"
+ have "min n0 n1 \<le> n0" by simp
+ with prems have th1:"min n0 n1 \<le> n" by auto
+ from prems have th21: "isnpolyh c (Suc n)" by simp
+ from prems have th22: "isnpolyh (CN c' n' p') n'" by simp
+ from lt have th23: "min (Suc n) n' = Suc n" by arith
+ from prems(4)[rule_format, OF lt th21 th22]
+ have "isnpolyh (polyadd (c, CN c' n' p')) (Suc n)" using th23 by simp
+ with prems th1 have ?case by simp }
+ moreover {assume gt: "n' < n" hence gt': "n' < n \<and> \<not> n < n'" by simp
+ have "min n0 n1 \<le> n1" by simp
+ with prems have th1:"min n0 n1 \<le> n'" by auto
+ from prems have th21: "isnpolyh c' (Suc n')" by simp_all
+ from prems have th22: "isnpolyh (CN c n p) n" by simp
+ from gt have th23: "min n (Suc n') = Suc n'" by arith
+ from prems(3)[rule_format, OF gt' th22 th21]
+ have "isnpolyh (polyadd (CN c n p,c')) (Suc n')" using th23 by simp
+ with prems th1 have ?case by simp}
+ ultimately show ?case by blast
+qed auto
+
+lemma polyadd[simp]: "Ipoly bs (polyadd (p,q)) = (Ipoly bs p) + (Ipoly bs q)"
+by (induct p q rule: polyadd.induct, auto simp add: Let_def ring_simps right_distrib[symmetric] simp del: right_distrib)
+
+lemma polyadd_norm: "\<lbrakk> isnpoly p ; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polyadd(p,q))"
+ using polyadd_normh[of "p" "0" "q" "0"] isnpoly_def by simp
+
+text{* The degree of addition and other general lemmas needed for the normal form of polymul*}
+
+lemma polyadd_different_degreen:
+ "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1; degreen p m \<noteq> degreen q m ; m \<le> min n0 n1\<rbrakk> \<Longrightarrow>
+ degreen (polyadd(p,q)) m = max (degreen p m) (degreen q m)"
+proof (induct p q arbitrary: m n0 n1 rule: polyadd.induct)
+ case (4 c n p c' n' p' m n0 n1)
+ thus ?case
+ apply (cases "n' < n", simp_all add: Let_def)
+ apply (cases "n = n'", simp_all)
+ apply (cases "n' = m", simp_all add: Let_def)
+ by (erule allE[where x="m"], erule allE[where x="Suc m"],
+ erule allE[where x="m"], erule allE[where x="Suc m"],
+ clarsimp,erule allE[where x="m"],erule allE[where x="Suc m"], simp)
+qed simp_all
+
+lemma headnz[simp]: "\<lbrakk>isnpolyh p n ; p \<noteq> 0\<^sub>p\<rbrakk> \<Longrightarrow> headn p m \<noteq> 0\<^sub>p"
+ by (induct p arbitrary: n rule: headn.induct, auto)
+lemma degree_isnpolyh_Suc[simp]: "isnpolyh p (Suc n) \<Longrightarrow> degree p = 0"
+ by (induct p arbitrary: n rule: degree.induct, auto)
+lemma degreen_0[simp]: "isnpolyh p n \<Longrightarrow> m < n \<Longrightarrow> degreen p m = 0"
+ by (induct p arbitrary: n rule: degreen.induct, auto)
+
+lemma degree_isnpolyh_Suc': "n > 0 \<Longrightarrow> isnpolyh p n \<Longrightarrow> degree p = 0"
+ by (induct p arbitrary: n rule: degree.induct, auto)
+
+lemma degree_npolyhCN[simp]: "isnpolyh (CN c n p) n0 \<Longrightarrow> degree c = 0"
+ using degree_isnpolyh_Suc by auto
+lemma degreen_npolyhCN[simp]: "isnpolyh (CN c n p) n0 \<Longrightarrow> degreen c n = 0"
+ using degreen_0 by auto
+
+
+lemma degreen_polyadd:
+ assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \<le> max n0 n1"
+ shows "degreen (p +\<^sub>p q) m \<le> max (degreen p m) (degreen q m)"
+ using np nq m
+proof (induct p q arbitrary: n0 n1 m rule: polyadd.induct)
+ case (2 c c' n' p' n0 n1) thus ?case by (cases n', simp_all)
+next
+ case (3 c n p c' n0 n1) thus ?case by (cases n, auto)
+next
+ case (4 c n p c' n' p' n0 n1 m)
+ thus ?case
+ apply (cases "n < n'", simp_all add: Let_def)
+ apply (cases "n' < n", simp_all)
+ apply (erule allE[where x="n"],erule allE[where x="Suc n"],clarify)
+ apply (erule allE[where x="n'"],erule allE[where x="Suc n'"],clarify)
+ by (erule allE[where x="m"],erule allE[where x="m"], auto)
+qed auto
+
+
+lemma polyadd_eq_const_degreen: "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; polyadd (p,q) = C c\<rbrakk>
+ \<Longrightarrow> degreen p m = degreen q m"
+proof (induct p q arbitrary: m n0 n1 c rule: polyadd.induct)
+ case (4 c n p c' n' p' m n0 n1 x)
+ hence z: "CN c n p +\<^sub>p CN c' n' p' = C x" by simp
+ {assume nn': "n' < n" hence ?case using prems by simp}
+ moreover
+ {assume nn':"\<not> n' < n" hence "n < n' \<or> n = n'" by arith
+ moreover {assume "n < n'" with prems have ?case by simp }
+ moreover {assume eq: "n = n'" hence ?case using prems
+ by (cases "p +\<^sub>p p' = 0\<^sub>p", auto simp add: Let_def) }
+ ultimately have ?case by blast}
+ ultimately show ?case by blast
+qed simp_all
+
+lemma polymul_properties:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \<le> min n0 n1"
+ shows "isnpolyh (p *\<^sub>p q) (min n0 n1)"
+ and "(p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p)"
+ and "degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \<or> q = 0\<^sub>p) then 0
+ else degreen p m + degreen q m)"
+ using np nq m
+proof(induct p q arbitrary: n0 n1 m rule: polymul.induct)
+ case (2 a b c' n' p')
+ let ?c = "(a,b)"
+ { case (1 n0 n1)
+ hence n: "isnpolyh (C ?c) n'" "isnpolyh c' (Suc n')" "isnpolyh p' n'" "isnormNum ?c"
+ "isnpolyh (CN c' n' p') n1"
+ by simp_all
+ {assume "?c = 0\<^sub>N" hence ?case by auto}
+ moreover {assume cnz: "?c \<noteq> 0\<^sub>N"
+ from "2.hyps"(1)[rule_format,where xb="n'", OF cnz n(1) n(3)]
+ "2.hyps"(2)[rule_format, where x="Suc n'"
+ and xa="Suc n'" and xb = "n'", OF cnz ] cnz n have ?case
+ by (auto simp add: min_def)}
+ ultimately show ?case by blast
+ next
+ case (2 n0 n1) thus ?case by auto
+ next
+ case (3 n0 n1) thus ?case using "2.hyps" by auto }
+next
+ case (3 c n p a b){
+ let ?c' = "(a,b)"
+ case (1 n0 n1)
+ hence n: "isnpolyh (C ?c') n" "isnpolyh c (Suc n)" "isnpolyh p n" "isnormNum ?c'"
+ "isnpolyh (CN c n p) n0"
+ by simp_all
+ {assume "?c' = 0\<^sub>N" hence ?case by auto}
+ moreover {assume cnz: "?c' \<noteq> 0\<^sub>N"
+ from "3.hyps"(1)[rule_format,where xb="n", OF cnz n(3) n(1)]
+ "3.hyps"(2)[rule_format, where x="Suc n"
+ and xa="Suc n" and xb = "n", OF cnz ] cnz n have ?case
+ by (auto simp add: min_def)}
+ ultimately show ?case by blast
+ next
+ case (2 n0 n1) thus ?case apply auto done
+ next
+ case (3 n0 n1) thus ?case using "3.hyps" by auto }
+next
+ case (4 c n p c' n' p')
+ let ?cnp = "CN c n p" let ?cnp' = "CN c' n' p'"
+ {fix n0 n1
+ assume "isnpolyh ?cnp n0" and "isnpolyh ?cnp' n1"
+ hence cnp: "isnpolyh ?cnp n" and cnp': "isnpolyh ?cnp' n'"
+ and np: "isnpolyh p n" and nc: "isnpolyh c (Suc n)"
+ and np': "isnpolyh p' n'" and nc': "isnpolyh c' (Suc n')"
+ and nn0: "n \<ge> n0" and nn1:"n' \<ge> n1"
+ by simp_all
+ have "n < n' \<or> n' < n \<or> n' = n" by auto
+ moreover
+ {assume nn': "n < n'"
+ with "4.hyps"(5)[rule_format, OF nn' np cnp', where xb ="n"]
+ "4.hyps"(6)[rule_format, OF nn' nc cnp', where xb="n"] nn' nn0 nn1 cnp
+ have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+ by (simp add: min_def) }
+ moreover
+
+ {assume nn': "n > n'" hence stupid: "n' < n \<and> \<not> n < n'" by arith
+ with "4.hyps"(3)[rule_format, OF stupid cnp np', where xb="n'"]
+ "4.hyps"(4)[rule_format, OF stupid cnp nc', where xb="Suc n'"]
+ nn' nn0 nn1 cnp'
+ have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+ by (cases "Suc n' = n", simp_all add: min_def)}
+ moreover
+ {assume nn': "n' = n" hence stupid: "\<not> n' < n \<and> \<not> n < n'" by arith
+ from "4.hyps"(1)[rule_format, OF stupid cnp np', where xb="n"]
+ "4.hyps"(2)[rule_format, OF stupid cnp nc', where xb="n"] nn' cnp cnp' nn1
+
+ have "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)"
+ by simp (rule polyadd_normh,simp_all add: min_def isnpolyh_mono[OF nn0]) }
+ ultimately show "isnpolyh (?cnp *\<^sub>p ?cnp') (min n0 n1)" by blast }
+ note th = this
+ {fix n0 n1 m
+ assume np: "isnpolyh ?cnp n0" and np':"isnpolyh ?cnp' n1"
+ and m: "m \<le> min n0 n1"
+ let ?d = "degreen (?cnp *\<^sub>p ?cnp') m"
+ let ?d1 = "degreen ?cnp m"
+ let ?d2 = "degreen ?cnp' m"
+ let ?eq = "?d = (if ?cnp = 0\<^sub>p \<or> ?cnp' = 0\<^sub>p then 0 else ?d1 + ?d2)"
+ have "n'<n \<or> n < n' \<or> n' = n" by auto
+ moreover
+ {assume "n' < n \<or> n < n'"
+ with "4.hyps" np np' m
+ have ?eq apply (cases "n' < n", simp_all)
+ apply (erule allE[where x="n"],erule allE[where x="n"],auto)
+ done }
+ moreover
+ {assume nn': "n' = n" hence nn:"\<not> n' < n \<and> \<not> n < n'" by arith
+ from "4.hyps"(1)[rule_format, OF nn, where x="n" and xa ="n'" and xb="n"]
+ "4.hyps"(2)[rule_format, OF nn, where x="n" and xa ="Suc n'" and xb="n"]
+ np np' nn'
+ have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n"
+ "isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n"
+ "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)"
+ "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p" by (auto simp add: min_def)
+ {assume mn: "m = n"
+ from "4.hyps"(1)[rule_format, OF nn norm(1,4), where xb="n"]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2), where xb="n"] norm nn' mn
+ have degs: "degreen (?cnp *\<^sub>p c') n =
+ (if c'=0\<^sub>p then 0 else ?d1 + degreen c' n)"
+ "degreen (?cnp *\<^sub>p p') n = ?d1 + degreen p' n" by (simp_all add: min_def)
+ from degs norm
+ have th1: "degreen(?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n" by simp
+ hence neq: "degreen (?cnp *\<^sub>p c') n \<noteq> degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n"
+ by simp
+ have nmin: "n \<le> min n n" by (simp add: min_def)
+ from polyadd_different_degreen[OF norm(3,6) neq nmin] th1
+ have deg: "degreen (CN c n p *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n = degreen (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
+ from "4.hyps"(1)[rule_format, OF nn norm(1,4), where xb="n"]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2), where xb="n"]
+ mn norm m nn' deg
+ have ?eq by simp}
+ moreover
+ {assume mn: "m \<noteq> n" hence mn': "m < n" using m np by auto
+ from nn' m np have max1: "m \<le> max n n" by simp
+ hence min1: "m \<le> min n n" by simp
+ hence min2: "m \<le> min n (Suc n)" by simp
+ {assume "c' = 0\<^sub>p"
+ from `c' = 0\<^sub>p` have ?eq
+ using "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2) min2] mn nn'
+ apply simp
+ done}
+ moreover
+ {assume cnz: "c' \<noteq> 0\<^sub>p"
+ from "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2) min2]
+ degreen_polyadd[OF norm(3,6) max1]
+
+ have "degreen (?cnp *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (?cnp *\<^sub>p p')) m
+ \<le> max (degreen (?cnp *\<^sub>p c') m) (degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) m)"
+ using mn nn' cnz np np' by simp
+ with "4.hyps"(1)[rule_format, OF nn norm(1,4) min1]
+ "4.hyps"(2)[rule_format, OF nn norm(1,2) min2]
+ degreen_0[OF norm(3) mn'] have ?eq using nn' mn cnz np np' by clarsimp}
+ ultimately have ?eq by blast }
+ ultimately have ?eq by blast}
+ ultimately show ?eq by blast}
+ note degth = this
+ { case (2 n0 n1)
+ hence np: "isnpolyh ?cnp n0" and np': "isnpolyh ?cnp' n1"
+ and m: "m \<le> min n0 n1" by simp_all
+ hence mn: "m \<le> n" by simp
+ let ?c0p = "CN 0\<^sub>p n (?cnp *\<^sub>p p')"
+ {assume C: "?cnp *\<^sub>p c' +\<^sub>p ?c0p = 0\<^sub>p" "n' = n"
+ hence nn: "\<not>n' < n \<and> \<not> n<n'" by simp
+ from "4.hyps"(1) [rule_format, OF nn, where x="n" and xa = "n" and xb="n"]
+ "4.hyps"(2) [rule_format, OF nn, where x="n" and xa = "Suc n" and xb="n"]
+ np np' C(2) mn
+ have norm: "isnpolyh ?cnp n" "isnpolyh c' (Suc n)" "isnpolyh (?cnp *\<^sub>p c') n"
+ "isnpolyh p' n" "isnpolyh (?cnp *\<^sub>p p') n" "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n"
+ "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)"
+ "?cnp *\<^sub>p p' \<noteq> 0\<^sub>p"
+ "degreen (?cnp *\<^sub>p c') n = (if c'=0\<^sub>p then 0 else degreen ?cnp n + degreen c' n)"
+ "degreen (?cnp *\<^sub>p p') n = degreen ?cnp n + degreen p' n"
+ by (simp_all add: min_def)
+
+ from norm have cn: "isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
+ have degneq: "degreen (?cnp *\<^sub>p c') n < degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) n"
+ using norm by simp
+ from polyadd_eq_const_degreen[OF norm(3) cn C(1), where m="n"] degneq
+ have "False" by simp }
+ thus ?case using "4.hyps" by clarsimp}
+qed auto
+
+lemma polymul[simp]: "Ipoly bs (p *\<^sub>p q) = (Ipoly bs p) * (Ipoly bs q)"
+by(induct p q rule: polymul.induct, auto simp add: ring_simps)
+
+lemma polymul_normh:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> isnpolyh (p *\<^sub>p q) (min n0 n1)"
+ using polymul_properties(1) by blast
+lemma polymul_eq0_iff:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \<or> q = 0\<^sub>p) "
+ using polymul_properties(2) by blast
+lemma polymul_degreen:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; m \<le> min n0 n1\<rbrakk> \<Longrightarrow> degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \<or> q = 0\<^sub>p) then 0 else degreen p m + degreen q m)"
+ using polymul_properties(3) by blast
+lemma polymul_norm:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polymul (p,q))"
+ using polymul_normh[of "p" "0" "q" "0"] isnpoly_def by simp
+
+lemma headconst_zero: "isnpolyh p n0 \<Longrightarrow> headconst p = 0\<^sub>N \<longleftrightarrow> p = 0\<^sub>p"
+ by (induct p arbitrary: n0 rule: headconst.induct, auto)
+
+lemma headconst_isnormNum: "isnpolyh p n0 \<Longrightarrow> isnormNum (headconst p)"
+ by (induct p arbitrary: n0, auto)
+
+lemma monic_eqI: assumes np: "isnpolyh p n0"
+ shows "INum (headconst p) * Ipoly bs (fst (monic p)) = (Ipoly bs p ::'a::{ring_char_0,power,division_by_zero,field})"
+ unfolding monic_def Let_def
+proof(cases "headconst p = 0\<^sub>N", simp_all add: headconst_zero[OF np])
+ let ?h = "headconst p"
+ assume pz: "p \<noteq> 0\<^sub>p"
+ {assume hz: "INum ?h = (0::'a)"
+ from headconst_isnormNum[OF np] have norm: "isnormNum ?h" "isnormNum 0\<^sub>N" by simp_all
+ from isnormNum_unique[where ?'a = 'a, OF norm] hz have "?h = 0\<^sub>N" by simp
+ with headconst_zero[OF np] have "p =0\<^sub>p" by blast with pz have "False" by blast}
+ thus "INum (headconst p) = (0::'a) \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = 0" by blast
+qed
+
+
+
+
+text{* polyneg is a negation and preserves normal form *}
+lemma polyneg[simp]: "Ipoly bs (polyneg p) = - Ipoly bs p"
+by (induct p rule: polyneg.induct, auto)
+
+lemma polyneg0: "isnpolyh p n \<Longrightarrow> ((~\<^sub>p p) = 0\<^sub>p) = (p = 0\<^sub>p)"
+ by (induct p arbitrary: n rule: polyneg.induct, auto simp add: Nneg_def)
+lemma polyneg_polyneg: "isnpolyh p n0 \<Longrightarrow> ~\<^sub>p (~\<^sub>p p) = p"
+ by (induct p arbitrary: n0 rule: polyneg.induct, auto)
+lemma polyneg_normh: "\<And>n. isnpolyh p n \<Longrightarrow> isnpolyh (polyneg p) n "
+by (induct p rule: polyneg.induct, auto simp add: polyneg0)
+
+lemma polyneg_norm: "isnpoly p \<Longrightarrow> isnpoly (polyneg p)"
+ using isnpoly_def polyneg_normh by simp
+
+
+text{* polysub is a substraction and preserves normalform *}
+lemma polysub[simp]: "Ipoly bs (polysub (p,q)) = (Ipoly bs p) - (Ipoly bs q)"
+by (simp add: polysub_def polyneg polyadd)
+lemma polysub_normh: "\<And> n0 n1. \<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> isnpolyh (polysub(p,q)) (min n0 n1)"
+by (simp add: polysub_def polyneg_normh polyadd_normh)
+
+lemma polysub_norm: "\<lbrakk> isnpoly p; isnpoly q\<rbrakk> \<Longrightarrow> isnpoly (polysub(p,q))"
+ using polyadd_norm polyneg_norm by (simp add: polysub_def)
+lemma polysub_same_0[simp]: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "isnpolyh p n0 \<Longrightarrow> polysub (p, p) = 0\<^sub>p"
+unfolding polysub_def split_def fst_conv snd_conv
+by (induct p arbitrary: n0,auto simp add: Let_def Nsub0[simplified Nsub_def])
+
+lemma polysub_0:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1\<rbrakk> \<Longrightarrow> (p -\<^sub>p q = 0\<^sub>p) = (p = q)"
+ unfolding polysub_def split_def fst_conv snd_conv
+ apply (induct p q arbitrary: n0 n1 rule:polyadd.induct, simp_all add: Nsub0[simplified Nsub_def])
+ apply (clarsimp simp add: Let_def)
+ apply (case_tac "n < n'", simp_all)
+ apply (case_tac "n' < n", simp_all)
+ apply (erule impE)+
+ apply (rule_tac x="Suc n" in exI, simp)
+ apply (rule_tac x="n" in exI, simp)
+ apply (erule impE)+
+ apply (rule_tac x="n" in exI, simp)
+ apply (rule_tac x="Suc n" in exI, simp)
+ apply (erule impE)+
+ apply (rule_tac x="Suc n" in exI, simp)
+ apply (rule_tac x="n" in exI, simp)
+ apply (erule impE)+
+ apply (rule_tac x="Suc n" in exI, simp)
+ apply clarsimp
+ done
+
+text{* polypow is a power function and preserves normal forms *}
+lemma polypow[simp]: "Ipoly bs (polypow n p) = ((Ipoly bs p :: 'a::{ring_char_0,division_by_zero,field})) ^ n"
+proof(induct n rule: polypow.induct)
+ case 1 thus ?case by simp
+next
+ case (2 n)
+ let ?q = "polypow ((Suc n) div 2) p"
+ let ?d = "polymul(?q,?q)"
+ have "odd (Suc n) \<or> even (Suc n)" by simp
+ moreover
+ {assume odd: "odd (Suc n)"
+ have th: "(Suc (Suc (Suc (0\<Colon>nat)) * (Suc n div Suc (Suc (0\<Colon>nat))))) = Suc n div 2 + Suc n div 2 + 1" by arith
+ from odd have "Ipoly bs (p ^\<^sub>p Suc n) = Ipoly bs (polymul(p, ?d))" by (simp add: Let_def)
+ also have "\<dots> = (Ipoly bs p) * (Ipoly bs p)^(Suc n div 2)*(Ipoly bs p)^(Suc n div 2)"
+ using "2.hyps" by simp
+ also have "\<dots> = (Ipoly bs p) ^ (Suc n div 2 + Suc n div 2 + 1)"
+ apply (simp only: power_add power_one_right) by simp
+ also have "\<dots> = (Ipoly bs p) ^ (Suc (Suc (Suc (0\<Colon>nat)) * (Suc n div Suc (Suc (0\<Colon>nat)))))"
+ by (simp only: th)
+ finally have ?case
+ using odd_nat_div_two_times_two_plus_one[OF odd, symmetric] by simp }
+ moreover
+ {assume even: "even (Suc n)"
+ have th: "(Suc (Suc (0\<Colon>nat))) * (Suc n div Suc (Suc (0\<Colon>nat))) = Suc n div 2 + Suc n div 2" by arith
+ from even have "Ipoly bs (p ^\<^sub>p Suc n) = Ipoly bs ?d" by (simp add: Let_def)
+ also have "\<dots> = (Ipoly bs p) ^ (Suc n div 2 + Suc n div 2)"
+ using "2.hyps" apply (simp only: power_add) by simp
+ finally have ?case using even_nat_div_two_times_two[OF even] by (simp only: th)}
+ ultimately show ?case by blast
+qed
+
+lemma polypow_normh:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "isnpolyh p n \<Longrightarrow> isnpolyh (polypow k p) n"
+proof (induct k arbitrary: n rule: polypow.induct)
+ case (2 k n)
+ let ?q = "polypow (Suc k div 2) p"
+ let ?d = "polymul (?q,?q)"
+ from prems have th1:"isnpolyh ?q n" and th2: "isnpolyh p n" by blast+
+ from polymul_normh[OF th1 th1] have dn: "isnpolyh ?d n" by simp
+ from polymul_normh[OF th2 dn] have on: "isnpolyh (polymul(p,?d)) n" by simp
+ from dn on show ?case by (simp add: Let_def)
+qed auto
+
+lemma polypow_norm:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "isnpoly p \<Longrightarrow> isnpoly (polypow k p)"
+ by (simp add: polypow_normh isnpoly_def)
+
+text{* Finally the whole normalization*}
+
+lemma polynate[simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{ring_char_0,division_by_zero,field})"
+by (induct p rule:polynate.induct, auto)
+
+lemma polynate_norm[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "isnpoly (polynate p)"
+ by (induct p rule: polynate.induct, simp_all add: polyadd_norm polymul_norm polysub_norm polyneg_norm polypow_norm) (simp_all add: isnpoly_def)
+
+text{* shift1 *}
+
+
+lemma shift1: "Ipoly bs (shift1 p) = Ipoly bs (Mul (Bound 0) p)"
+by (simp add: shift1_def polymul)
+
+lemma shift1_isnpoly:
+ assumes pn: "isnpoly p" and pnz: "p \<noteq> 0\<^sub>p" shows "isnpoly (shift1 p) "
+ using pn pnz by (simp add: shift1_def isnpoly_def )
+
+lemma shift1_nz[simp]:"shift1 p \<noteq> 0\<^sub>p"
+ by (simp add: shift1_def)
+lemma funpow_shift1_isnpoly:
+ "\<lbrakk> isnpoly p ; p \<noteq> 0\<^sub>p\<rbrakk> \<Longrightarrow> isnpoly (funpow n shift1 p)"
+ by (induct n arbitrary: p, auto simp add: shift1_isnpoly)
+
+lemma funpow_isnpolyh:
+ assumes f: "\<And> p. isnpolyh p n \<Longrightarrow> isnpolyh (f p) n "and np: "isnpolyh p n"
+ shows "isnpolyh (funpow k f p) n"
+ using f np by (induct k arbitrary: p, auto)
+
+lemma funpow_shift1: "(Ipoly bs (funpow n shift1 p) :: 'a :: {ring_char_0,division_by_zero,field}) = Ipoly bs (Mul (Pw (Bound 0) n) p)"
+ by (induct n arbitrary: p, simp_all add: shift1_isnpoly shift1 power_Suc )
+
+lemma shift1_isnpolyh: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> isnpolyh (shift1 p) 0"
+ using isnpolyh_mono[where n="n0" and n'="0" and p="p"] by (simp add: shift1_def)
+
+lemma funpow_shift1_1:
+ "(Ipoly bs (funpow n shift1 p) :: 'a :: {ring_char_0,division_by_zero,field}) = Ipoly bs (funpow n shift1 1\<^sub>p *\<^sub>p p)"
+ by (simp add: funpow_shift1)
+
+lemma poly_cmul[simp]: "Ipoly bs (poly_cmul c p) = Ipoly bs (Mul (C c) p)"
+by (induct p arbitrary: n0 rule: poly_cmul.induct, auto simp add: ring_simps)
+
+lemma behead:
+ assumes np: "isnpolyh p n"
+ shows "Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = (Ipoly bs p :: 'a :: {ring_char_0,division_by_zero,field})"
+ using np
+proof (induct p arbitrary: n rule: behead.induct)
+ case (1 c p n) hence pn: "isnpolyh p n" by simp
+ from prems(2)[OF pn]
+ have th:"Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = Ipoly bs p" .
+ then show ?case using "1.hyps" apply (simp add: Let_def,cases "behead p = 0\<^sub>p")
+ by (simp_all add: th[symmetric] ring_simps power_Suc)
+qed (auto simp add: Let_def)
+
+lemma behead_isnpolyh:
+ assumes np: "isnpolyh p n" shows "isnpolyh (behead p) n"
+ using np by (induct p rule: behead.induct, auto simp add: Let_def isnpolyh_mono)
+
+subsection{* Miscilanious lemmas about indexes, decrementation, substitution etc ... *}
+lemma isnpolyh_polybound0: "isnpolyh p (Suc n) \<Longrightarrow> polybound0 p"
+proof(induct p arbitrary: n rule: polybound0.induct, auto)
+ case (goal1 c n p n')
+ hence "n = Suc (n - 1)" by simp
+ hence "isnpolyh p (Suc (n - 1))" using `isnpolyh p n` by simp
+ with prems(2) show ?case by simp
+qed
+
+lemma isconstant_polybound0: "isnpolyh p n0 \<Longrightarrow> isconstant p \<longleftrightarrow> polybound0 p"
+by (induct p arbitrary: n0 rule: isconstant.induct, auto simp add: isnpolyh_polybound0)
+
+lemma decrpoly_zero[simp]: "decrpoly p = 0\<^sub>p \<longleftrightarrow> p = 0\<^sub>p" by (induct p, auto)
+
+lemma decrpoly_normh: "isnpolyh p n0 \<Longrightarrow> polybound0 p \<Longrightarrow> isnpolyh (decrpoly p) (n0 - 1)"
+ apply (induct p arbitrary: n0, auto)
+ apply (atomize)
+ apply (erule_tac x = "Suc nat" in allE)
+ apply auto
+ done
+
+lemma head_polybound0: "isnpolyh p n0 \<Longrightarrow> polybound0 (head p)"
+ by (induct p arbitrary: n0 rule: head.induct, auto intro: isnpolyh_polybound0)
+
+lemma polybound0_I:
+ assumes nb: "polybound0 a"
+ shows "Ipoly (b#bs) a = Ipoly (b'#bs) a"
+using nb
+by (induct a rule: polybound0.induct) auto
+lemma polysubst0_I:
+ shows "Ipoly (b#bs) (polysubst0 a t) = Ipoly ((Ipoly (b#bs) a)#bs) t"
+ by (induct t) simp_all
+
+lemma polysubst0_I':
+ assumes nb: "polybound0 a"
+ shows "Ipoly (b#bs) (polysubst0 a t) = Ipoly ((Ipoly (b'#bs) a)#bs) t"
+ by (induct t) (simp_all add: polybound0_I[OF nb, where b="b" and b'="b'"])
+
+lemma decrpoly: assumes nb: "polybound0 t"
+ shows "Ipoly (x#bs) t = Ipoly bs (decrpoly t)"
+ using nb by (induct t rule: decrpoly.induct, simp_all)
+
+lemma polysubst0_polybound0: assumes nb: "polybound0 t"
+ shows "polybound0 (polysubst0 t a)"
+using nb by (induct a rule: polysubst0.induct, auto)
+
+lemma degree0_polybound0: "isnpolyh p n \<Longrightarrow> degree p = 0 \<Longrightarrow> polybound0 p"
+ by (induct p arbitrary: n rule: degree.induct, auto simp add: isnpolyh_polybound0)
+
+fun maxindex :: "poly \<Rightarrow> nat" where
+ "maxindex (Bound n) = n + 1"
+| "maxindex (CN c n p) = max (n + 1) (max (maxindex c) (maxindex p))"
+| "maxindex (Add p q) = max (maxindex p) (maxindex q)"
+| "maxindex (Sub p q) = max (maxindex p) (maxindex q)"
+| "maxindex (Mul p q) = max (maxindex p) (maxindex q)"
+| "maxindex (Neg p) = maxindex p"
+| "maxindex (Pw p n) = maxindex p"
+| "maxindex (C x) = 0"
+
+definition wf_bs :: "'a list \<Rightarrow> poly \<Rightarrow> bool" where
+ "wf_bs bs p = (length bs \<ge> maxindex p)"
+
+lemma wf_bs_coefficients: "wf_bs bs p \<Longrightarrow> \<forall> c \<in> set (coefficients p). wf_bs bs c"
+proof(induct p rule: coefficients.induct)
+ case (1 c p)
+ show ?case
+ proof
+ fix x assume xc: "x \<in> set (coefficients (CN c 0 p))"
+ hence "x = c \<or> x \<in> set (coefficients p)" by simp
+ moreover
+ {assume "x = c" hence "wf_bs bs x" using "1.prems" unfolding wf_bs_def by simp}
+ moreover
+ {assume H: "x \<in> set (coefficients p)"
+ from "1.prems" have "wf_bs bs p" unfolding wf_bs_def by simp
+ with "1.hyps" H have "wf_bs bs x" by blast }
+ ultimately show "wf_bs bs x" by blast
+ qed
+qed simp_all
+
+lemma maxindex_coefficients: " \<forall>c\<in> set (coefficients p). maxindex c \<le> maxindex p"
+by (induct p rule: coefficients.induct, auto)
+
+lemma length_exists: "\<exists>xs. length xs = n" by (rule exI[where x="replicate n x"], simp)
+
+lemma wf_bs_I: "wf_bs bs p ==> Ipoly (bs@bs') p = Ipoly bs p"
+ unfolding wf_bs_def by (induct p, auto simp add: nth_append)
+
+lemma take_maxindex_wf: assumes wf: "wf_bs bs p"
+ shows "Ipoly (take (maxindex p) bs) p = Ipoly bs p"
+proof-
+ let ?ip = "maxindex p"
+ let ?tbs = "take ?ip bs"
+ from wf have "length ?tbs = ?ip" unfolding wf_bs_def by simp
+ hence wf': "wf_bs ?tbs p" unfolding wf_bs_def by simp
+ have eq: "bs = ?tbs @ (drop ?ip bs)" by simp
+ from wf_bs_I[OF wf', of "drop ?ip bs"] show ?thesis using eq by simp
+qed
+
+lemma decr_maxindex: "polybound0 p \<Longrightarrow> maxindex (decrpoly p) = maxindex p - 1"
+ by (induct p, auto)
+
+lemma wf_bs_insensitive: "length bs = length bs' \<Longrightarrow> wf_bs bs p = wf_bs bs' p"
+ unfolding wf_bs_def by simp
+
+lemma wf_bs_insensitive': "wf_bs (x#bs) p = wf_bs (y#bs) p"
+ unfolding wf_bs_def by simp
+
+
+
+lemma wf_bs_coefficients': "\<forall>c \<in> set (coefficients p). wf_bs bs c \<Longrightarrow> wf_bs (x#bs) p"
+by(induct p rule: coefficients.induct, auto simp add: wf_bs_def)
+lemma coefficients_Nil[simp]: "coefficients p \<noteq> []"
+ by (induct p rule: coefficients.induct, simp_all)
+
+
+lemma coefficients_head: "last (coefficients p) = head p"
+ by (induct p rule: coefficients.induct, auto)
+
+lemma wf_bs_decrpoly: "wf_bs bs (decrpoly p) \<Longrightarrow> wf_bs (x#bs) p"
+ unfolding wf_bs_def by (induct p rule: decrpoly.induct, auto)
+
+lemma length_le_list_ex: "length xs \<le> n \<Longrightarrow> \<exists> ys. length (xs @ ys) = n"
+ apply (rule exI[where x="replicate (n - length xs) z"])
+ by simp
+lemma isnpolyh_Suc_const:"isnpolyh p (Suc n) \<Longrightarrow> isconstant p"
+by (cases p, auto) (case_tac "nat", simp_all)
+
+lemma wf_bs_polyadd: "wf_bs bs p \<and> wf_bs bs q \<longrightarrow> wf_bs bs (p +\<^sub>p q)"
+ unfolding wf_bs_def
+ apply (induct p q rule: polyadd.induct)
+ apply (auto simp add: Let_def)
+ done
+
+lemma wf_bs_polyul: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p *\<^sub>p q)"
+
+ unfolding wf_bs_def
+ apply (induct p q arbitrary: bs rule: polymul.induct)
+ apply (simp_all add: wf_bs_polyadd)
+ apply clarsimp
+ apply (rule wf_bs_polyadd[unfolded wf_bs_def, rule_format])
+ apply auto
+ done
+
+lemma wf_bs_polyneg: "wf_bs bs p \<Longrightarrow> wf_bs bs (~\<^sub>p p)"
+ unfolding wf_bs_def by (induct p rule: polyneg.induct, auto)
+
+lemma wf_bs_polysub: "wf_bs bs p \<Longrightarrow> wf_bs bs q \<Longrightarrow> wf_bs bs (p -\<^sub>p q)"
+ unfolding polysub_def split_def fst_conv snd_conv using wf_bs_polyadd wf_bs_polyneg by blast
+
+subsection{* Canonicity of polynomial representation, see lemma isnpolyh_unique*}
+
+definition "polypoly bs p = map (Ipoly bs) (coefficients p)"
+definition "polypoly' bs p = map ((Ipoly bs o decrpoly)) (coefficients p)"
+definition "poly_nate bs p = map ((Ipoly bs o decrpoly)) (coefficients (polynate p))"
+
+lemma coefficients_normh: "isnpolyh p n0 \<Longrightarrow> \<forall> q \<in> set (coefficients p). isnpolyh q n0"
+proof (induct p arbitrary: n0 rule: coefficients.induct)
+ case (1 c p n0)
+ have cp: "isnpolyh (CN c 0 p) n0" by fact
+ hence norm: "isnpolyh c 0" "isnpolyh p 0" "p \<noteq> 0\<^sub>p" "n0 = 0"
+ by (auto simp add: isnpolyh_mono[where n'=0])
+ from "1.hyps"[OF norm(2)] norm(1) norm(4) show ?case by simp
+qed auto
+
+lemma coefficients_isconst:
+ "isnpolyh p n \<Longrightarrow> \<forall>q\<in>set (coefficients p). isconstant q"
+ by (induct p arbitrary: n rule: coefficients.induct,
+ auto simp add: isnpolyh_Suc_const)
+
+lemma polypoly_polypoly':
+ assumes np: "isnpolyh p n0"
+ shows "polypoly (x#bs) p = polypoly' bs p"
+proof-
+ let ?cf = "set (coefficients p)"
+ from coefficients_normh[OF np] have cn_norm: "\<forall> q\<in> ?cf. isnpolyh q n0" .
+ {fix q assume q: "q \<in> ?cf"
+ from q cn_norm have th: "isnpolyh q n0" by blast
+ from coefficients_isconst[OF np] q have "isconstant q" by blast
+ with isconstant_polybound0[OF th] have "polybound0 q" by blast}
+ hence "\<forall>q \<in> ?cf. polybound0 q" ..
+ hence "\<forall>q \<in> ?cf. Ipoly (x#bs) q = Ipoly bs (decrpoly q)"
+ using polybound0_I[where b=x and bs=bs and b'=y] decrpoly[where x=x and bs=bs]
+ by auto
+
+ thus ?thesis unfolding polypoly_def polypoly'_def by simp
+qed
+
+lemma polypoly_poly:
+ assumes np: "isnpolyh p n0" shows "Ipoly (x#bs) p = poly (polypoly (x#bs) p) x"
+ using np
+by (induct p arbitrary: n0 bs rule: coefficients.induct, auto simp add: polypoly_def)
+
+lemma polypoly'_poly:
+ assumes np: "isnpolyh p n0" shows "\<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup> = poly (polypoly' bs p) x"
+ using polypoly_poly[OF np, simplified polypoly_polypoly'[OF np]] .
+
+
+lemma polypoly_poly_polybound0:
+ assumes np: "isnpolyh p n0" and nb: "polybound0 p"
+ shows "polypoly bs p = [Ipoly bs p]"
+ using np nb unfolding polypoly_def
+ by (cases p, auto, case_tac nat, auto)
+
+lemma head_isnpolyh: "isnpolyh p n0 \<Longrightarrow> isnpolyh (head p) n0"
+ by (induct p rule: head.induct, auto)
+
+lemma headn_nz[simp]: "isnpolyh p n0 \<Longrightarrow> (headn p m = 0\<^sub>p) = (p = 0\<^sub>p)"
+ by (cases p,auto)
+
+lemma head_eq_headn0: "head p = headn p 0"
+ by (induct p rule: head.induct, simp_all)
+
+lemma head_nz[simp]: "isnpolyh p n0 \<Longrightarrow> (head p = 0\<^sub>p) = (p = 0\<^sub>p)"
+ by (simp add: head_eq_headn0)
+
+lemma isnpolyh_zero_iff:
+ assumes nq: "isnpolyh p n0" and eq :"\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a::{ring_char_0,power,division_by_zero,field})"
+ shows "p = 0\<^sub>p"
+using nq eq
+proof (induct n\<equiv>"maxindex p" arbitrary: p n0 rule: nat_less_induct)
+ fix n p n0
+ assume H: "\<forall>m<n. \<forall>p n0. isnpolyh p n0 \<longrightarrow>
+ (\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)) \<longrightarrow> m = maxindex p \<longrightarrow> p = 0\<^sub>p"
+ and np: "isnpolyh p n0" and zp: "\<forall>bs. wf_bs bs p \<longrightarrow> \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" and n: "n = maxindex p"
+ {assume nz: "n = 0"
+ then obtain c where "p = C c" using n np by (cases p, auto)
+ with zp np have "p = 0\<^sub>p" unfolding wf_bs_def by simp}
+ moreover
+ {assume nz: "n \<noteq> 0"
+ let ?h = "head p"
+ let ?hd = "decrpoly ?h"
+ let ?ihd = "maxindex ?hd"
+ from head_isnpolyh[OF np] head_polybound0[OF np] have h:"isnpolyh ?h n0" "polybound0 ?h"
+ by simp_all
+ hence nhd: "isnpolyh ?hd (n0 - 1)" using decrpoly_normh by blast
+
+ from maxindex_coefficients[of p] coefficients_head[of p, symmetric]
+ have mihn: "maxindex ?h \<le> n" unfolding n by auto
+ with decr_maxindex[OF h(2)] nz have ihd_lt_n: "?ihd < n" by auto
+ {fix bs:: "'a list" assume bs: "wf_bs bs ?hd"
+ let ?ts = "take ?ihd bs"
+ let ?rs = "drop ?ihd bs"
+ have ts: "wf_bs ?ts ?hd" using bs unfolding wf_bs_def by simp
+ have bs_ts_eq: "?ts@ ?rs = bs" by simp
+ from wf_bs_decrpoly[OF ts] have tsh: " \<forall>x. wf_bs (x#?ts) ?h" by simp
+ from ihd_lt_n have "ALL x. length (x#?ts) \<le> n" by simp
+ with length_le_list_ex obtain xs where xs:"length ((x#?ts) @ xs) = n" by blast
+ hence "\<forall> x. wf_bs ((x#?ts) @ xs) p" using n unfolding wf_bs_def by simp
+ with zp have "\<forall> x. Ipoly ((x#?ts) @ xs) p = 0" by blast
+ hence "\<forall> x. Ipoly (x#(?ts @ xs)) p = 0" by simp
+ with polypoly_poly[OF np, where ?'a = 'a] polypoly_polypoly'[OF np, where ?'a = 'a]
+ have "\<forall>x. poly (polypoly' (?ts @ xs) p) x = poly [] x" by simp
+ hence "poly (polypoly' (?ts @ xs) p) = poly []" by (auto intro: ext)
+ hence "\<forall> c \<in> set (coefficients p). Ipoly (?ts @ xs) (decrpoly c) = 0"
+ thm poly_zero
+ using poly_zero[where ?'a='a] by (simp add: polypoly'_def list_all_iff)
+ with coefficients_head[of p, symmetric]
+ have th0: "Ipoly (?ts @ xs) ?hd = 0" by simp
+ from bs have wf'': "wf_bs ?ts ?hd" unfolding wf_bs_def by simp
+ with th0 wf_bs_I[of ?ts ?hd xs] have "Ipoly ?ts ?hd = 0" by simp
+ with wf'' wf_bs_I[of ?ts ?hd ?rs] bs_ts_eq have "\<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = 0" by simp }
+ then have hdz: "\<forall>bs. wf_bs bs ?hd \<longrightarrow> \<lparr>?hd\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" by blast
+
+ from H[rule_format, OF ihd_lt_n nhd] hdz have "?hd = 0\<^sub>p" by blast
+ hence "?h = 0\<^sub>p" by simp
+ with head_nz[OF np] have "p = 0\<^sub>p" by simp}
+ ultimately show "p = 0\<^sub>p" by blast
+qed
+
+lemma isnpolyh_unique:
+ assumes np:"isnpolyh p n0" and nq: "isnpolyh q n1"
+ shows "(\<forall>bs. \<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (\<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup> :: 'a::{ring_char_0,power,division_by_zero,field})) \<longleftrightarrow> p = q"
+proof(auto)
+ assume H: "\<forall>bs. (\<lparr>p\<rparr>\<^sub>p\<^bsup>bs\<^esup> ::'a)= \<lparr>q\<rparr>\<^sub>p\<^bsup>bs\<^esup>"
+ hence "\<forall>bs.\<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup>= (0::'a)" by simp
+ hence "\<forall>bs. wf_bs bs (p -\<^sub>p q) \<longrightarrow> \<lparr>p -\<^sub>p q\<rparr>\<^sub>p\<^bsup>bs\<^esup> = (0::'a)"
+ using wf_bs_polysub[where p=p and q=q] by auto
+ with isnpolyh_zero_iff[OF polysub_normh[OF np nq]] polysub_0[OF np nq]
+ show "p = q" by blast
+qed
+
+
+text{* consequenses of unicity on the algorithms for polynomial normalization *}
+
+lemma polyadd_commute: assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" and nq: "isnpolyh q n1" shows "p +\<^sub>p q = q +\<^sub>p p"
+ using isnpolyh_unique[OF polyadd_normh[OF np nq] polyadd_normh[OF nq np]] by simp
+
+lemma zero_normh: "isnpolyh 0\<^sub>p n" by simp
+lemma one_normh: "isnpolyh 1\<^sub>p n" by simp
+lemma polyadd_0[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" shows "p +\<^sub>p 0\<^sub>p = p" and "0\<^sub>p +\<^sub>p p = p"
+ using isnpolyh_unique[OF polyadd_normh[OF np zero_normh] np]
+ isnpolyh_unique[OF polyadd_normh[OF zero_normh np] np] by simp_all
+
+lemma polymul_1[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" shows "p *\<^sub>p 1\<^sub>p = p" and "1\<^sub>p *\<^sub>p p = p"
+ using isnpolyh_unique[OF polymul_normh[OF np one_normh] np]
+ isnpolyh_unique[OF polymul_normh[OF one_normh np] np] by simp_all
+lemma polymul_0[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" shows "p *\<^sub>p 0\<^sub>p = 0\<^sub>p" and "0\<^sub>p *\<^sub>p p = 0\<^sub>p"
+ using isnpolyh_unique[OF polymul_normh[OF np zero_normh] zero_normh]
+ isnpolyh_unique[OF polymul_normh[OF zero_normh np] zero_normh] by simp_all
+
+lemma polymul_commute:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np:"isnpolyh p n0" and nq: "isnpolyh q n1"
+ shows "p *\<^sub>p q = q *\<^sub>p p"
+using isnpolyh_unique[OF polymul_normh[OF np nq] polymul_normh[OF nq np], where ?'a = "'a\<Colon>{ring_char_0,power,division_by_zero,field}"] by simp
+
+declare polyneg_polyneg[simp]
+
+lemma isnpolyh_polynate_id[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np:"isnpolyh p n0" shows "polynate p = p"
+ using isnpolyh_unique[where ?'a= "'a::{ring_char_0,division_by_zero,field}", OF polynate_norm[of p, unfolded isnpoly_def] np] polynate[where ?'a = "'a::{ring_char_0,division_by_zero,field}"] by simp
+
+lemma polynate_idempotent[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "polynate (polynate p) = polynate p"
+ using isnpolyh_polynate_id[OF polynate_norm[of p, unfolded isnpoly_def]] .
+
+lemma poly_nate_polypoly': "poly_nate bs p = polypoly' bs (polynate p)"
+ unfolding poly_nate_def polypoly'_def ..
+lemma poly_nate_poly: shows "poly (poly_nate bs p) = (\<lambda>x:: 'a ::{ring_char_0,division_by_zero,field}. \<lparr>p\<rparr>\<^sub>p\<^bsup>x # bs\<^esup>)"
+ using polypoly'_poly[OF polynate_norm[unfolded isnpoly_def], symmetric, of bs p]
+ unfolding poly_nate_polypoly' by (auto intro: ext)
+
+subsection{* heads, degrees and all that *}
+lemma degree_eq_degreen0: "degree p = degreen p 0"
+ by (induct p rule: degree.induct, simp_all)
+
+lemma degree_polyneg: assumes n: "isnpolyh p n"
+ shows "degree (polyneg p) = degree p"
+ using n
+ by (induct p arbitrary: n rule: polyneg.induct, simp_all) (case_tac na, auto)
+
+lemma degree_polyadd:
+ assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1"
+ shows "degree (p +\<^sub>p q) \<le> max (degree p) (degree q)"
+using degreen_polyadd[OF np nq, where m= "0"] degree_eq_degreen0 by simp
+
+
+lemma degree_polysub: assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1"
+ shows "degree (p -\<^sub>p q) \<le> max (degree p) (degree q)"
+proof-
+ from nq have nq': "isnpolyh (~\<^sub>p q) n1" using polyneg_normh by simp
+ from degree_polyadd[OF np nq'] show ?thesis by (simp add: polysub_def degree_polyneg[OF nq])
+qed
+
+lemma degree_polysub_samehead:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and h: "head p = head q"
+ and d: "degree p = degree q"
+ shows "degree (p -\<^sub>p q) < degree p \<or> (p -\<^sub>p q = 0\<^sub>p)"
+unfolding polysub_def split_def fst_conv snd_conv
+using np nq h d
+proof(induct p q rule:polyadd.induct)
+ case (1 a b a' b') thus ?case by (simp add: Nsub_def Nsub0[simplified Nsub_def])
+next
+ case (2 a b c' n' p')
+ let ?c = "(a,b)"
+ from prems have "degree (C ?c) = degree (CN c' n' p')" by simp
+ hence nz:"n' > 0" by (cases n', auto)
+ hence "head (CN c' n' p') = CN c' n' p'" by (cases n', auto)
+ with prems show ?case by simp
+next
+ case (3 c n p a' b')
+ let ?c' = "(a',b')"
+ from prems have "degree (C ?c') = degree (CN c n p)" by simp
+ hence nz:"n > 0" by (cases n, auto)
+ hence "head (CN c n p) = CN c n p" by (cases n, auto)
+ with prems show ?case by simp
+next
+ case (4 c n p c' n' p')
+ hence H: "isnpolyh (CN c n p) n0" "isnpolyh (CN c' n' p') n1"
+ "head (CN c n p) = head (CN c' n' p')" "degree (CN c n p) = degree (CN c' n' p')" by simp+
+ hence degc: "degree c = 0" and degc': "degree c' = 0" by simp_all
+ hence degnc: "degree (~\<^sub>p c) = 0" and degnc': "degree (~\<^sub>p c') = 0"
+ using H(1-2) degree_polyneg by auto
+ from H have cnh: "isnpolyh c (Suc n)" and c'nh: "isnpolyh c' (Suc n')" by simp+
+ from degree_polysub[OF cnh c'nh, simplified polysub_def] degc degc' have degcmc': "degree (c +\<^sub>p ~\<^sub>pc') = 0" by simp
+ from H have pnh: "isnpolyh p n" and p'nh: "isnpolyh p' n'" by auto
+ have "n = n' \<or> n < n' \<or> n > n'" by arith
+ moreover
+ {assume nn': "n = n'"
+ have "n = 0 \<or> n >0" by arith
+ moreover {assume nz: "n = 0" hence ?case using prems by (auto simp add: Let_def degcmc')}
+ moreover {assume nz: "n > 0"
+ with nn' H(3) have cc':"c = c'" and pp': "p = p'" by (cases n, auto)+
+ hence ?case using polysub_same_0[OF p'nh, simplified polysub_def split_def fst_conv snd_conv] polysub_same_0[OF c'nh, simplified polysub_def split_def fst_conv snd_conv] using nn' prems by (simp add: Let_def)}
+ ultimately have ?case by blast}
+ moreover
+ {assume nn': "n < n'" hence n'p: "n' > 0" by simp
+ hence headcnp':"head (CN c' n' p') = CN c' n' p'" by (cases n', simp_all)
+ have degcnp': "degree (CN c' n' p') = 0" and degcnpeq: "degree (CN c n p) = degree (CN c' n' p')" using prems by (cases n', simp_all)
+ hence "n > 0" by (cases n, simp_all)
+ hence headcnp: "head (CN c n p) = CN c n p" by (cases n, auto)
+ from H(3) headcnp headcnp' nn' have ?case by auto}
+ moreover
+ {assume nn': "n > n'" hence np: "n > 0" by simp
+ hence headcnp:"head (CN c n p) = CN c n p" by (cases n, simp_all)
+ from prems have degcnpeq: "degree (CN c' n' p') = degree (CN c n p)" by simp
+ from np have degcnp: "degree (CN c n p) = 0" by (cases n, simp_all)
+ with degcnpeq have "n' > 0" by (cases n', simp_all)
+ hence headcnp': "head (CN c' n' p') = CN c' n' p'" by (cases n', auto)
+ from H(3) headcnp headcnp' nn' have ?case by auto}
+ ultimately show ?case by blast
+qed auto
+
+lemma shift1_head : "isnpolyh p n0 \<Longrightarrow> head (shift1 p) = head p"
+by (induct p arbitrary: n0 rule: head.induct, simp_all add: shift1_def)
+
+lemma funpow_shift1_head: "isnpolyh p n0 \<Longrightarrow> p\<noteq> 0\<^sub>p \<Longrightarrow> head (funpow k shift1 p) = head p"
+proof(induct k arbitrary: n0 p)
+ case (Suc k n0 p) hence "isnpolyh (shift1 p) 0" by (simp add: shift1_isnpolyh)
+ with prems have "head (funpow k shift1 (shift1 p)) = head (shift1 p)"
+ and "head (shift1 p) = head p" by (simp_all add: shift1_head)
+ thus ?case by simp
+qed auto
+
+lemma shift1_degree: "degree (shift1 p) = 1 + degree p"
+ by (simp add: shift1_def)
+lemma funpow_shift1_degree: "degree (funpow k shift1 p) = k + degree p "
+ by (induct k arbitrary: p, auto simp add: shift1_degree)
+
+lemma funpow_shift1_nz: "p \<noteq> 0\<^sub>p \<Longrightarrow> funpow n shift1 p \<noteq> 0\<^sub>p"
+ by (induct n arbitrary: p, simp_all add: funpow_def)
+
+lemma head_isnpolyh_Suc[simp]: "isnpolyh p (Suc n) \<Longrightarrow> head p = p"
+ by (induct p arbitrary: n rule: degree.induct, auto)
+lemma headn_0[simp]: "isnpolyh p n \<Longrightarrow> m < n \<Longrightarrow> headn p m = p"
+ by (induct p arbitrary: n rule: degreen.induct, auto)
+lemma head_isnpolyh_Suc': "n > 0 \<Longrightarrow> isnpolyh p n \<Longrightarrow> head p = p"
+ by (induct p arbitrary: n rule: degree.induct, auto)
+lemma head_head[simp]: "isnpolyh p n0 \<Longrightarrow> head (head p) = head p"
+ by (induct p rule: head.induct, auto)
+
+lemma polyadd_eq_const_degree:
+ "\<lbrakk> isnpolyh p n0 ; isnpolyh q n1 ; polyadd (p,q) = C c\<rbrakk> \<Longrightarrow> degree p = degree q"
+ using polyadd_eq_const_degreen degree_eq_degreen0 by simp
+
+lemma polyadd_head: assumes np: "isnpolyh p n0" and nq: "isnpolyh q n1"
+ and deg: "degree p \<noteq> degree q"
+ shows "head (p +\<^sub>p q) = (if degree p < degree q then head q else head p)"
+using np nq deg
+apply(induct p q arbitrary: n0 n1 rule: polyadd.induct,simp_all)
+apply (case_tac n', simp, simp)
+apply (case_tac n, simp, simp)
+apply (case_tac n, case_tac n', simp add: Let_def)
+apply (case_tac "pa +\<^sub>p p' = 0\<^sub>p")
+apply (clarsimp simp add: polyadd_eq_const_degree)
+apply clarsimp
+apply (erule_tac impE,blast)
+apply (erule_tac impE,blast)
+apply clarsimp
+apply simp
+apply (case_tac n', simp_all)
+done
+
+lemma polymul_head_polyeq:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "\<lbrakk>isnpolyh p n0; isnpolyh q n1 ; p \<noteq> 0\<^sub>p ; q \<noteq> 0\<^sub>p \<rbrakk> \<Longrightarrow> head (p *\<^sub>p q) = head p *\<^sub>p head q"
+proof (induct p q arbitrary: n0 n1 rule: polymul.induct)
+ case (2 a b c' n' p' n0 n1)
+ hence "isnpolyh (head (CN c' n' p')) n1" "isnormNum (a,b)" by (simp_all add: head_isnpolyh)
+ thus ?case using prems by (cases n', auto)
+next
+ case (3 c n p a' b' n0 n1)
+ hence "isnpolyh (head (CN c n p)) n0" "isnormNum (a',b')" by (simp_all add: head_isnpolyh)
+ thus ?case using prems by (cases n, auto)
+next
+ case (4 c n p c' n' p' n0 n1)
+ hence norm: "isnpolyh p n" "isnpolyh c (Suc n)" "isnpolyh p' n'" "isnpolyh c' (Suc n')"
+ "isnpolyh (CN c n p) n" "isnpolyh (CN c' n' p') n'"
+ by simp_all
+ have "n < n' \<or> n' < n \<or> n = n'" by arith
+ moreover
+ {assume nn': "n < n'" hence ?case
+ thm prems
+ using norm
+ prems(6)[rule_format, OF nn' norm(1,6)]
+ prems(7)[rule_format, OF nn' norm(2,6)] by (simp, cases n, simp,cases n', simp_all)}
+ moreover {assume nn': "n'< n"
+ hence stupid: "n' < n \<and> \<not> n < n'" by simp
+ hence ?case using norm prems(4) [rule_format, OF stupid norm(5,3)]
+ prems(5)[rule_format, OF stupid norm(5,4)]
+ by (simp,cases n',simp,cases n,auto)}
+ moreover {assume nn': "n' = n"
+ hence stupid: "\<not> n' < n \<and> \<not> n < n'" by simp
+ from nn' polymul_normh[OF norm(5,4)]
+ have ncnpc':"isnpolyh (CN c n p *\<^sub>p c') n" by (simp add: min_def)
+ from nn' polymul_normh[OF norm(5,3)] norm
+ have ncnpp':"isnpolyh (CN c n p *\<^sub>p p') n" by simp
+ from nn' ncnpp' polymul_eq0_iff[OF norm(5,3)] norm(6)
+ have ncnpp0':"isnpolyh (CN 0\<^sub>p n (CN c n p *\<^sub>p p')) n" by simp
+ from polyadd_normh[OF ncnpc' ncnpp0']
+ have nth: "isnpolyh ((CN c n p *\<^sub>p c') +\<^sub>p (CN 0\<^sub>p n (CN c n p *\<^sub>p p'))) n"
+ by (simp add: min_def)
+ {assume np: "n > 0"
+ with nn' head_isnpolyh_Suc'[OF np nth]
+ head_isnpolyh_Suc'[OF np norm(5)] head_isnpolyh_Suc'[OF np norm(6)[simplified nn']]
+ have ?case by simp}
+ moreover
+ {moreover assume nz: "n = 0"
+ from polymul_degreen[OF norm(5,4), where m="0"]
+ polymul_degreen[OF norm(5,3), where m="0"] nn' nz degree_eq_degreen0
+ norm(5,6) degree_npolyhCN[OF norm(6)]
+ have dth:"degree (CN c 0 p *\<^sub>p c') < degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))" by simp
+ hence dth':"degree (CN c 0 p *\<^sub>p c') \<noteq> degree (CN 0\<^sub>p 0 (CN c 0 p *\<^sub>p p'))" by simp
+ from polyadd_head[OF ncnpc'[simplified nz] ncnpp0'[simplified nz] dth'] dth
+ have ?case using norm prems(2)[rule_format, OF stupid norm(5,3)]
+ prems(3)[rule_format, OF stupid norm(5,4)] nn' nz by simp }
+ ultimately have ?case by (cases n) auto}
+ ultimately show ?case by blast
+qed simp_all
+
+lemma degree_coefficients: "degree p = length (coefficients p) - 1"
+ by(induct p rule: degree.induct, auto)
+
+lemma degree_head[simp]: "degree (head p) = 0"
+ by (induct p rule: head.induct, auto)
+
+lemma degree_CN: "isnpolyh p n \<Longrightarrow> degree (CN c n p) \<le> 1+ degree p"
+ by (cases n, simp_all)
+lemma degree_CN': "isnpolyh p n \<Longrightarrow> degree (CN c n p) \<ge> degree p"
+ by (cases n, simp_all)
+
+lemma polyadd_different_degree: "\<lbrakk>isnpolyh p n0 ; isnpolyh q n1; degree p \<noteq> degree q\<rbrakk> \<Longrightarrow> degree (polyadd(p,q)) = max (degree p) (degree q)"
+ using polyadd_different_degreen degree_eq_degreen0 by simp
+
+lemma degreen_polyneg: "isnpolyh p n0 \<Longrightarrow> degreen (~\<^sub>p p) m = degreen p m"
+ by (induct p arbitrary: n0 rule: polyneg.induct, auto)
+
+lemma degree_polymul:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" and nq: "isnpolyh q n1"
+ shows "degree (p *\<^sub>p q) \<le> degree p + degree q"
+ using polymul_degreen[OF np nq, where m="0"] degree_eq_degreen0 by simp
+
+lemma polyneg_degree: "isnpolyh p n \<Longrightarrow> degree (polyneg p) = degree p"
+ by (induct p arbitrary: n rule: degree.induct, auto)
+
+lemma polyneg_head: "isnpolyh p n \<Longrightarrow> head(polyneg p) = polyneg (head p)"
+ by (induct p arbitrary: n rule: degree.induct, auto)
+
+subsection {* Correctness of polynomial pseudo division *}
+
+lemma polydivide_aux_real_domintros:
+ assumes call1: "\<lbrakk>s \<noteq> 0\<^sub>p; \<not> degree s < n; a = head s\<rbrakk>
+ \<Longrightarrow> polydivide_aux_dom (a, n, p, k, s -\<^sub>p funpow (degree s - n) shift1 p)"
+ and call2 : "\<lbrakk>s \<noteq> 0\<^sub>p; \<not> degree s < n; a \<noteq> head s\<rbrakk>
+ \<Longrightarrow> polydivide_aux_dom(a, n, p,Suc k, a *\<^sub>p s -\<^sub>p (head s *\<^sub>p funpow (degree s - n) shift1 p))"
+ shows "polydivide_aux_dom (a, n, p, k, s)"
+proof (rule accpI, erule polydivide_aux_rel.cases)
+ fix y aa ka na pa sa x xa xb
+ assume eqs: "y = (aa, na, pa, ka, sa -\<^sub>p xb)" "(a, n, p, k, s) = (aa, na, pa, ka, sa)"
+ and \<Gamma>1': "sa \<noteq> 0\<^sub>p" "x = head sa" "xa = degree sa" "\<not> xa < na"
+ "xb = funpow (xa - na) shift1 pa" "aa = x"
+
+ hence \<Gamma>1: "s \<noteq> 0\<^sub>p" "a = head s" "xa = degree s" "\<not> degree s < n" "\<not> xa < na"
+ "xb = funpow (xa - na) shift1 pa" "aa = x" by auto
+
+ with call1 have "polydivide_aux_dom (a, n, p, k, s -\<^sub>p funpow (degree s - n) shift1 p)"
+ by auto
+ with eqs \<Gamma>1 show "polydivide_aux_dom y" by auto
+next
+ fix y aa ka na pa sa x xa xb
+ assume eqs: "y = (aa, na, pa, Suc ka, aa *\<^sub>p sa -\<^sub>p (x *\<^sub>p xb))"
+ "(a, n, p, k, s) =(aa, na, pa, ka, sa)"
+ and \<Gamma>2': "sa \<noteq> 0\<^sub>p" "x = head sa" "xa = degree sa" "\<not> xa < na"
+ "xb = funpow (xa - na) shift1 pa" "aa \<noteq> x"
+ hence \<Gamma>2: "s \<noteq> 0\<^sub>p" "a \<noteq> head s" "xa = degree s" "\<not> degree s < n"
+ "xb = funpow (xa - na) shift1 pa" "aa \<noteq> x" by auto
+ with call2 have "polydivide_aux_dom (a, n, p, Suc k, a *\<^sub>p s -\<^sub>p (head s *\<^sub>p funpow (degree s - n) shift1 p))" by auto
+ with eqs \<Gamma>2' show "polydivide_aux_dom y" by auto
+qed
+
+lemma polydivide_aux_properties:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" and ns: "isnpolyh s n1"
+ and ap: "head p = a" and ndp: "degree p = n" and pnz: "p \<noteq> 0\<^sub>p"
+ shows "polydivide_aux_dom (a,n,p,k,s) \<and>
+ (polydivide_aux (a,n,p,k,s) = (k',r) \<longrightarrow> (k' \<ge> k) \<and> (degree r = 0 \<or> degree r < degree p)
+ \<and> (\<exists>nr. isnpolyh r nr) \<and> (\<exists>q n1. isnpolyh q n1 \<and> ((polypow (k' - k) a) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))"
+ using ns
+proof(induct d\<equiv>"degree s" arbitrary: s k k' r n1 rule: nat_less_induct)
+ fix d s k k' r n1
+ let ?D = "polydivide_aux_dom"
+ let ?dths = "?D (a, n, p, k, s)"
+ let ?qths = "\<exists>q n1. isnpolyh q n1 \<and> (a ^\<^sub>p (k' - k) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)"
+ let ?qrths = "polydivide_aux (a, n, p, k, s) = (k', r) \<longrightarrow> k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p)
+ \<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths"
+ let ?ths = "?dths \<and> ?qrths"
+ let ?b = "head s"
+ let ?p' = "funpow (d - n) shift1 p"
+ let ?xdn = "funpow (d - n) shift1 1\<^sub>p"
+ let ?akk' = "a ^\<^sub>p (k' - k)"
+ assume H: "\<forall>m<d. \<forall>x xa xb xc xd.
+ isnpolyh x xd \<longrightarrow>
+ m = degree x \<longrightarrow> ?D (a, n, p, xa, x) \<and>
+ (polydivide_aux (a, n, p, xa, x) = (xb, xc) \<longrightarrow>
+ xa \<le> xb \<and> (degree xc = 0 \<or> degree xc < degree p) \<and>
+ (\<exists> nr. isnpolyh xc nr) \<and>
+ (\<exists>q n1. isnpolyh q n1 \<and> a ^\<^sub>p xb - xa *\<^sub>p x = p *\<^sub>p q +\<^sub>p xc))"
+ and ns: "isnpolyh s n1" and ds: "d = degree s"
+ from np have np0: "isnpolyh p 0"
+ using isnpolyh_mono[where n="n0" and n'="0" and p="p"] by simp
+ have np': "isnpolyh ?p' 0" using funpow_shift1_isnpoly[OF np0[simplified isnpoly_def[symmetric]] pnz, where n="d -n"] isnpoly_def by simp
+ have headp': "head ?p' = head p" using funpow_shift1_head[OF np pnz] by simp
+ from funpow_shift1_isnpoly[where p="1\<^sub>p"] have nxdn: "isnpolyh ?xdn 0" by (simp add: isnpoly_def)
+ from polypow_normh [OF head_isnpolyh[OF np0], where k="k' - k"] ap
+ have nakk':"isnpolyh ?akk' 0" by blast
+ {assume sz: "s = 0\<^sub>p"
+ hence dom: ?dths apply - apply (rule polydivide_aux_real_domintros) apply simp_all done
+ from polydivide_aux.psimps[OF dom] sz
+ have ?qrths using np apply clarsimp by (rule exI[where x="0\<^sub>p"], simp)
+ hence ?ths using dom by blast}
+ moreover
+ {assume sz: "s \<noteq> 0\<^sub>p"
+ {assume dn: "d < n"
+ with sz ds have dom:"?dths" by - (rule polydivide_aux_real_domintros,simp_all)
+ from polydivide_aux.psimps[OF dom] sz dn ds
+ have "?qrths" using ns ndp np by auto (rule exI[where x="0\<^sub>p"],simp)
+ with dom have ?ths by blast}
+ moreover
+ {assume dn': "\<not> d < n" hence dn: "d \<ge> n" by arith
+ have degsp': "degree s = degree ?p'"
+ using ds dn ndp funpow_shift1_degree[where k = "d - n" and p="p"] by simp
+ {assume ba: "?b = a"
+ hence headsp': "head s = head ?p'" using ap headp' by simp
+ have nr: "isnpolyh (s -\<^sub>p ?p') 0" using polysub_normh[OF ns np'] by simp
+ from ds degree_polysub_samehead[OF ns np' headsp' degsp']
+ have "degree (s -\<^sub>p ?p') < d \<or> s -\<^sub>p ?p' = 0\<^sub>p" by simp
+ moreover
+ {assume deglt:"degree (s -\<^sub>p ?p') < d"
+ from H[rule_format, OF deglt nr,simplified]
+ have domsp: "?D (a, n, p, k, s -\<^sub>p ?p')" by blast
+ have dom: ?dths apply (rule polydivide_aux_real_domintros)
+ using ba ds dn' domsp by simp_all
+ from polydivide_aux.psimps[OF dom] sz dn' ba ds
+ have eq: "polydivide_aux (a,n,p,k,s) = polydivide_aux (a,n,p,k, s -\<^sub>p ?p')"
+ by (simp add: Let_def)
+ {assume h1: "polydivide_aux (a, n, p, k, s) = (k', r)"
+ from H[rule_format, OF deglt nr, where xa = "k" and xb="k'" and xc="r", simplified]
+ trans[OF eq[symmetric] h1]
+ have kk': "k \<le> k'" and nr:"\<exists>nr. isnpolyh r nr" and dr: "degree r = 0 \<or> degree r < degree p"
+ and q1:"\<exists>q nq. isnpolyh q nq \<and> (a ^\<^sub>p k' - k *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r)" by auto
+ from q1 obtain q n1 where nq: "isnpolyh q n1"
+ and asp:"a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p') = p *\<^sub>p q +\<^sub>p r" by blast
+ from nr obtain nr where nr': "isnpolyh r nr" by blast
+ from polymul_normh[OF nakk' ns] have nakks': "isnpolyh (a ^\<^sub>p (k' - k) *\<^sub>p s) 0" by simp
+ from polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]
+ have nq': "isnpolyh (?akk' *\<^sub>p ?xdn +\<^sub>p q) 0" by simp
+ from polyadd_normh[OF polymul_normh[OF np
+ polyadd_normh[OF polymul_normh[OF nakk' nxdn] nq]] nr']
+ have nqr': "isnpolyh (p *\<^sub>p (?akk' *\<^sub>p ?xdn +\<^sub>p q) +\<^sub>p r) 0" by simp
+ from asp have "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) =
+ Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
+ hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a^\<^sub>p (k' - k)*\<^sub>p s) =
+ Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs ?p' + Ipoly bs p * Ipoly bs q + Ipoly bs r"
+ by (simp add: ring_simps)
+ hence " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (d - n) shift1 1\<^sub>p *\<^sub>p p)
+ + Ipoly bs p * Ipoly bs q + Ipoly bs r"
+ by (auto simp only: funpow_shift1_1)
+ hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (d - n) shift1 1\<^sub>p)
+ + Ipoly bs q) + Ipoly bs r" by (simp add: ring_simps)
+ hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r)" by simp
+ with isnpolyh_unique[OF nakks' nqr']
+ have "a ^\<^sub>p (k' - k) *\<^sub>p s =
+ p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q) +\<^sub>p r" by blast
+ hence ?qths using nq'
+ apply (rule_tac x="(a^\<^sub>p (k' - k)) *\<^sub>p (funpow (d - n) shift1 1\<^sub>p) +\<^sub>p q" in exI)
+ apply (rule_tac x="0" in exI) by simp
+ with kk' nr dr have "k \<le> k' \<and> (degree r = 0 \<or> degree r < degree p) \<and> (\<exists>nr. isnpolyh r nr) \<and> ?qths"
+ by blast } hence ?qrths by blast
+ with dom have ?ths by blast}
+ moreover
+ {assume spz:"s -\<^sub>p ?p' = 0\<^sub>p"
+ hence domsp: "?D (a, n, p, k, s -\<^sub>p ?p')"
+ apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
+ have dom: ?dths apply (rule polydivide_aux_real_domintros)
+ using ba ds dn' domsp by simp_all
+ from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{ring_char_0,division_by_zero,field}"]
+ have " \<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs ?p'" by simp
+ hence "\<forall>(bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp
+ by (simp only: funpow_shift1_1) simp
+ hence sp': "s = ?xdn *\<^sub>p p" using isnpolyh_unique[OF ns polymul_normh[OF nxdn np]] by blast
+ {assume h1: "polydivide_aux (a,n,p,k,s) = (k',r)"
+ from polydivide_aux.psimps[OF dom] sz dn' ba ds
+ have eq: "polydivide_aux (a,n,p,k,s) = polydivide_aux (a,n,p,k, s -\<^sub>p ?p')"
+ by (simp add: Let_def)
+ also have "\<dots> = (k,0\<^sub>p)" using polydivide_aux.psimps[OF domsp] spz by simp
+ finally have "(k',r) = (k,0\<^sub>p)" using h1 by simp
+ with sp' ns np nxdn polyadd_0(1)[OF polymul_normh[OF np nxdn]]
+ polyadd_0(2)[OF polymul_normh[OF np nxdn]] have ?qrths
+ apply auto
+ apply (rule exI[where x="?xdn"])
+ apply auto
+ apply (rule polymul_commute)
+ apply simp_all
+ done}
+ with dom have ?ths by blast}
+ ultimately have ?ths by blast }
+ moreover
+ {assume ba: "?b \<noteq> a"
+ from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns]
+ polymul_normh[OF head_isnpolyh[OF ns] np']]
+ have nth: "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by(simp add: min_def)
+ have nzths: "a *\<^sub>p s \<noteq> 0\<^sub>p" "?b *\<^sub>p ?p' \<noteq> 0\<^sub>p"
+ using polymul_eq0_iff[OF head_isnpolyh[OF np0, simplified ap] ns]
+ polymul_eq0_iff[OF head_isnpolyh[OF ns] np']head_nz[OF np0] ap pnz sz head_nz[OF ns]
+ funpow_shift1_nz[OF pnz] by simp_all
+ from polymul_head_polyeq[OF head_isnpolyh[OF np] ns] head_nz[OF np] sz ap head_head[OF np] pnz
+ polymul_head_polyeq[OF head_isnpolyh[OF ns] np'] head_nz [OF ns] sz funpow_shift1_nz[OF pnz, where n="d - n"]
+ have hdth: "head (a *\<^sub>p s) = head (?b *\<^sub>p ?p')"
+ using head_head[OF ns] funpow_shift1_head[OF np pnz]
+ polymul_commute[OF head_isnpolyh[OF np] head_isnpolyh[OF ns]]
+ by (simp add: ap)
+ from polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
+ head_nz[OF np] pnz sz ap[symmetric]
+ funpow_shift1_nz[OF pnz, where n="d - n"]
+ polymul_degreen[OF head_isnpolyh[OF ns] np', where m="0"] head_nz[OF ns]
+ ndp ds[symmetric] dn
+ have degth: "degree (a *\<^sub>p s) = degree (?b *\<^sub>p ?p') "
+ by (simp add: degree_eq_degreen0[symmetric] funpow_shift1_degree)
+ {assume dth: "degree ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) < d"
+ have th: "?D (a, n, p, Suc k, (a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))"
+ using H[rule_format, OF dth nth, simplified] by blast
+ have dom: ?dths
+ using ba ds dn' th apply simp apply (rule polydivide_aux_real_domintros)
+ using ba ds dn' by simp_all
+ from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np] ns] polymul_normh[OF head_isnpolyh[OF ns]np']]
+ ap have nasbp': "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" by simp
+ {assume h1:"polydivide_aux (a,n,p,k,s) = (k', r)"
+ from h1 polydivide_aux.psimps[OF dom] sz dn' ba ds
+ have eq:"polydivide_aux (a,n,p,Suc k,(a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = (k',r)"
+ by (simp add: Let_def)
+ with H[rule_format, OF dth nasbp', simplified, where xa="Suc k" and xb="k'" and xc="r"]
+ obtain q nq nr where kk': "Suc k \<le> k'" and nr: "isnpolyh r nr" and nq: "isnpolyh q nq"
+ and dr: "degree r = 0 \<or> degree r < degree p"
+ and qr: "a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) = p *\<^sub>p q +\<^sub>p r" by auto
+ from kk' have kk'':"Suc (k' - Suc k) = k' - k" by arith
+ {fix bs:: "'a::{ring_char_0,division_by_zero,field} list"
+
+ from qr isnpolyh_unique[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k", simplified ap] nasbp', symmetric]
+ have "Ipoly bs (a ^\<^sub>p (k' - Suc k) *\<^sub>p ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p'))) = Ipoly bs (p *\<^sub>p q +\<^sub>p r)" by simp
+ hence "Ipoly bs a ^ (Suc (k' - Suc k)) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?p' + Ipoly bs r"
+ by (simp add: ring_simps power_Suc)
+ hence "Ipoly bs a ^ (k' - k) * Ipoly bs s = Ipoly bs p * Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn * Ipoly bs p + Ipoly bs r"
+ by (simp add:kk'' funpow_shift1_1[where n="d - n" and p="p"])
+ hence "Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = Ipoly bs p * (Ipoly bs q + Ipoly bs a ^ (k' - Suc k) * Ipoly bs ?b * Ipoly bs ?xdn) + Ipoly bs r"
+ by (simp add: ring_simps)}
+ hence ieq:"\<forall>(bs :: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) =
+ Ipoly bs (p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r)" by auto
+ let ?q = "q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)"
+ from polyadd_normh[OF nq polymul_normh[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - Suc k"] head_isnpolyh[OF ns], simplified ap ] nxdn]]
+ have nqw: "isnpolyh ?q 0" by simp
+ from ieq isnpolyh_unique[OF polymul_normh[OF polypow_normh[OF head_isnpolyh[OF np], where k="k' - k"] ns, simplified ap] polyadd_normh[OF polymul_normh[OF np nqw] nr]]
+ have asth: "(a ^\<^sub>p (k' - k) *\<^sub>p s) = p *\<^sub>p (q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn)) +\<^sub>p r" by blast
+ from dr kk' nr h1 asth nqw have ?qrths apply simp
+ apply (rule conjI)
+ apply (rule exI[where x="nr"], simp)
+ apply (rule exI[where x="(q +\<^sub>p (a ^\<^sub>p (k' - Suc k) *\<^sub>p ?b *\<^sub>p ?xdn))"], simp)
+ apply (rule exI[where x="0"], simp)
+ done}
+ hence ?qrths by blast
+ with dom have ?ths by blast}
+ moreover
+ {assume spz: "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p"
+ hence domsp: "?D (a, n, p, Suc k, a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p'))"
+ apply (simp) by (rule polydivide_aux_real_domintros, simp_all)
+ have dom: ?dths using sz ba dn' ds domsp
+ by - (rule polydivide_aux_real_domintros, simp_all)
+ {fix bs :: "'a::{ring_char_0,division_by_zero,field} list"
+ from isnpolyh_unique[OF nth, where ?'a="'a" and q="0\<^sub>p",simplified,symmetric] spz
+ have "Ipoly bs (a*\<^sub>p s) = Ipoly bs ?b * Ipoly bs ?p'" by simp
+ hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (?b *\<^sub>p ?xdn) * Ipoly bs p"
+ by (simp add: funpow_shift1_1[where n="d - n" and p="p"])
+ hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" by simp
+ }
+ hence hth: "\<forall> (bs:: 'a::{ring_char_0,division_by_zero,field} list). Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" ..
+ from hth
+ have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)"
+ using isnpolyh_unique[where ?'a = "'a::{ring_char_0,division_by_zero,field}", OF polymul_normh[OF head_isnpolyh[OF np] ns]
+ polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]],
+ simplified ap] by simp
+ {assume h1: "polydivide_aux (a,n,p,k,s) = (k', r)"
+ from h1 sz ds ba dn' spz polydivide_aux.psimps[OF dom] polydivide_aux.psimps[OF domsp]
+ have "(k',r) = (Suc k, 0\<^sub>p)" by (simp add: Let_def)
+ with h1 np head_isnpolyh[OF np, simplified ap] ns polymul_normh[OF head_isnpolyh[OF ns] nxdn]
+ polymul_normh[OF np polymul_normh[OF head_isnpolyh[OF ns] nxdn]] asq
+ have ?qrths apply (clarsimp simp add: Let_def)
+ apply (rule exI[where x="?b *\<^sub>p ?xdn"]) apply simp
+ apply (rule exI[where x="0"], simp)
+ done}
+ hence ?qrths by blast
+ with dom have ?ths by blast}
+ ultimately have ?ths using degree_polysub_samehead[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] polymul_normh[OF head_isnpolyh[OF ns] np'] hdth degth] polymul_degreen[OF head_isnpolyh[OF np] ns, where m="0"]
+ head_nz[OF np] pnz sz ap[symmetric] ds[symmetric]
+ by (simp add: degree_eq_degreen0[symmetric]) blast }
+ ultimately have ?ths by blast
+ }
+ ultimately have ?ths by blast}
+ ultimately show ?ths by blast
+qed
+
+lemma polydivide_properties:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ and np: "isnpolyh p n0" and ns: "isnpolyh s n1" and pnz: "p \<noteq> 0\<^sub>p"
+ shows "(\<exists> k r. polydivide s p = (k,r) \<and> (\<exists>nr. isnpolyh r nr) \<and> (degree r = 0 \<or> degree r < degree p)
+ \<and> (\<exists>q n1. isnpolyh q n1 \<and> ((polypow k (head p)) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))"
+proof-
+ have trv: "head p = head p" "degree p = degree p" by simp_all
+ from polydivide_aux_properties[OF np ns trv pnz, where k="0"]
+ have d: "polydivide_aux_dom (head p, degree p, p, 0, s)" by blast
+ from polydivide_def[where s="s" and p="p"] polydivide_aux.psimps[OF d]
+ have ex: "\<exists> k r. polydivide s p = (k,r)" by auto
+ then obtain k r where kr: "polydivide s p = (k,r)" by blast
+ from trans[OF meta_eq_to_obj_eq[OF polydivide_def[where s="s" and p="p"], symmetric] kr]
+ polydivide_aux_properties[OF np ns trv pnz, where k="0" and k'="k" and r="r"]
+ have "(degree r = 0 \<or> degree r < degree p) \<and>
+ (\<exists>nr. isnpolyh r nr) \<and> (\<exists>q n1. isnpolyh q n1 \<and> head p ^\<^sub>p k - 0 *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)" by blast
+ with kr show ?thesis
+ apply -
+ apply (rule exI[where x="k"])
+ apply (rule exI[where x="r"])
+ apply simp
+ done
+qed
+
+subsection{* More about polypoly and pnormal etc *}
+
+definition "isnonconstant p = (\<not> isconstant p)"
+
+lemma last_map: "xs \<noteq> [] ==> last (map f xs) = f (last xs)" by (induct xs, auto)
+
+lemma isnonconstant_pnormal_iff: assumes nc: "isnonconstant p"
+ shows "pnormal (polypoly bs p) \<longleftrightarrow> Ipoly bs (head p) \<noteq> 0"
+proof
+ let ?p = "polypoly bs p"
+ assume H: "pnormal ?p"
+ have csz: "coefficients p \<noteq> []" using nc by (cases p, auto)
+
+ from coefficients_head[of p] last_map[OF csz, of "Ipoly bs"]
+ pnormal_last_nonzero[OF H]
+ show "Ipoly bs (head p) \<noteq> 0" by (simp add: polypoly_def)
+next
+ assume h: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+ let ?p = "polypoly bs p"
+ have csz: "coefficients p \<noteq> []" using nc by (cases p, auto)
+ hence pz: "?p \<noteq> []" by (simp add: polypoly_def)
+ hence lg: "length ?p > 0" by simp
+ from h coefficients_head[of p] last_map[OF csz, of "Ipoly bs"]
+ have lz: "last ?p \<noteq> 0" by (simp add: polypoly_def)
+ from pnormal_last_length[OF lg lz] show "pnormal ?p" .
+qed
+
+lemma isnonconstant_coefficients_length: "isnonconstant p \<Longrightarrow> length (coefficients p) > 1"
+ unfolding isnonconstant_def
+ apply (cases p, simp_all)
+ apply (case_tac nat, auto)
+ done
+lemma isnonconstant_nonconstant: assumes inc: "isnonconstant p"
+ shows "nonconstant (polypoly bs p) \<longleftrightarrow> Ipoly bs (head p) \<noteq> 0"
+proof
+ let ?p = "polypoly bs p"
+ assume nc: "nonconstant ?p"
+ from isnonconstant_pnormal_iff[OF inc, of bs] nc
+ show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" unfolding nonconstant_def by blast
+next
+ let ?p = "polypoly bs p"
+ assume h: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+ from isnonconstant_pnormal_iff[OF inc, of bs] h
+ have pn: "pnormal ?p" by blast
+ {fix x assume H: "?p = [x]"
+ from H have "length (coefficients p) = 1" unfolding polypoly_def by auto
+ with isnonconstant_coefficients_length[OF inc] have False by arith}
+ thus "nonconstant ?p" using pn unfolding nonconstant_def by blast
+qed
+
+lemma pnormal_length: "p\<noteq>[] \<Longrightarrow> pnormal p \<longleftrightarrow> length (pnormalize p) = length p"
+ unfolding pnormal_def
+ apply (induct p rule: pnormalize.induct, simp_all)
+ apply (case_tac "p=[]", simp_all)
+ done
+
+lemma degree_degree: assumes inc: "isnonconstant p"
+ shows "degree p = Polynomial_List.degree (polypoly bs p) \<longleftrightarrow> \<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+proof
+ let ?p = "polypoly bs p"
+ assume H: "degree p = Polynomial_List.degree ?p"
+ from isnonconstant_coefficients_length[OF inc] have pz: "?p \<noteq> []"
+ unfolding polypoly_def by auto
+ from H degree_coefficients[of p] isnonconstant_coefficients_length[OF inc]
+ have lg:"length (pnormalize ?p) = length ?p"
+ unfolding Polynomial_List.degree_def polypoly_def by simp
+ hence "pnormal ?p" using pnormal_length[OF pz] by blast
+ with isnonconstant_pnormal_iff[OF inc]
+ show "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0" by blast
+next
+ let ?p = "polypoly bs p"
+ assume H: "\<lparr>head p\<rparr>\<^sub>p\<^bsup>bs\<^esup> \<noteq> 0"
+ with isnonconstant_pnormal_iff[OF inc] have "pnormal ?p" by blast
+ with degree_coefficients[of p] isnonconstant_coefficients_length[OF inc]
+ show "degree p = Polynomial_List.degree ?p"
+ unfolding polypoly_def pnormal_def Polynomial_List.degree_def by auto
+qed
+
+section{* Swaps ; Division by a certain variable *}
+consts swap:: "nat \<Rightarrow> nat \<Rightarrow> poly \<Rightarrow> poly"
+primrec
+ "swap n m (C x) = C x"
+ "swap n m (Bound k) = Bound (if k = n then m else if k=m then n else k)"
+ "swap n m (Neg t) = Neg (swap n m t)"
+ "swap n m (Add s t) = Add (swap n m s) (swap n m t)"
+ "swap n m (Sub s t) = Sub (swap n m s) (swap n m t)"
+ "swap n m (Mul s t) = Mul (swap n m s) (swap n m t)"
+ "swap n m (Pw t k) = Pw (swap n m t) k"
+ "swap n m (CN c k p) = CN (swap n m c) (if k = n then m else if k=m then n else k)
+ (swap n m p)"
+
+lemma swap: assumes nbs: "n < length bs" and mbs: "m < length bs"
+ shows "Ipoly bs (swap n m t) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t"
+proof (induct t)
+ case (Bound k) thus ?case using nbs mbs by simp
+next
+ case (CN c k p) thus ?case using nbs mbs by simp
+qed simp_all
+lemma swap_swap_id[simp]: "swap n m (swap m n t) = t"
+ by (induct t,simp_all)
+
+lemma swap_commute: "swap n m p = swap m n p" by (induct p, simp_all)
+
+lemma swap_same_id[simp]: "swap n n t = t"
+ by (induct t, simp_all)
+
+definition "swapnorm n m t = polynate (swap n m t)"
+
+lemma swapnorm: assumes nbs: "n < length bs" and mbs: "m < length bs"
+ shows "((Ipoly bs (swapnorm n m t) :: 'a\<Colon>{ring_char_0,division_by_zero,field})) = Ipoly ((bs[n:= bs!m])[m:= bs!n]) t"
+ using swap[OF prems] swapnorm_def by simp
+
+lemma swapnorm_isnpoly[simp]:
+ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero,field})"
+ shows "isnpoly (swapnorm n m p)"
+ unfolding swapnorm_def by simp
+
+definition "polydivideby n s p =
+ (let ss = swapnorm 0 n s ; sp = swapnorm 0 n p ; h = head sp; (k,r) = polydivide ss sp
+ in (k,swapnorm 0 n h,swapnorm 0 n r))"
+
+lemma swap_nz [simp]: " (swap n m p = 0\<^sub>p) = (p = 0\<^sub>p)" by (induct p, simp_all)
+
+consts isweaknpoly :: "poly \<Rightarrow> bool"
+recdef isweaknpoly "measure size"
+ "isweaknpoly (C c) = True"
+ "isweaknpoly (CN c n p) \<longleftrightarrow> isweaknpoly c \<and> isweaknpoly p"
+ "isweaknpoly p = False"
+
+lemma isnpolyh_isweaknpoly: "isnpolyh p n0 \<Longrightarrow> isweaknpoly p"
+ by (induct p arbitrary: n0, auto)
+
+lemma swap_isweanpoly: "isweaknpoly p \<Longrightarrow> isweaknpoly (swap n m p)"
+ by (induct p, auto)
+
+end
\ No newline at end of file