# HG changeset patch # User wenzelm # Date 1375218977 -7200 # Node ID bcaa5bbf7e6b928cf7cada4e77283c92d70ad03d # Parent 0b98561d07907ae2572bc99864bed3fcb3a73485 tuned proofs; diff -r 0b98561d0790 -r bcaa5bbf7e6b src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy --- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Tue Jul 30 22:43:11 2013 +0200 +++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Tue Jul 30 23:16:17 2013 +0200 @@ -8,7 +8,7 @@ imports Complex_Main "~~/src/HOL/Library/Abstract_Rat" Polynomial_List begin -subsection{* Datatype of polynomial expressions *} +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 @@ -36,7 +36,7 @@ | "polybound0 (Bound n) = (n>0)" | "polybound0 (Neg a) = polybound0 a" | "polybound0 (Add a b) = (polybound0 a \ polybound0 b)" -| "polybound0 (Sub a b) = (polybound0 a \ polybound0 b)" +| "polybound0 (Sub a b) = (polybound0 a \ polybound0 b)" | "polybound0 (Mul a b) = (polybound0 a \ polybound0 b)" | "polybound0 (Pw p n) = (polybound0 p)" | "polybound0 (CN c n p) = (n \ 0 \ polybound0 c \ polybound0 p)" @@ -47,13 +47,13 @@ | "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 (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))" -fun decrpoly:: "poly \ poly" +fun decrpoly:: "poly \ poly" where "decrpoly (Bound n) = Bound (n - 1)" | "decrpoly (Neg a) = Neg (decrpoly a)" @@ -117,12 +117,12 @@ fun polyadd :: "poly \ poly \ poly" (infixl "+\<^sub>p" 60) where "polyadd (C c) (C c') = C (c+\<^sub>Nc')" -| "polyadd (C c) (CN c' n' p') = CN (polyadd (C c) c') n' p'" +| "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" | "polyadd (CN c n p) (CN c' n' p') = (if n < n' then CN (polyadd c (CN c' n' p')) n p else if n'p then cc' else CN cc' n pp')))" | "polyadd a b = Add a b" @@ -140,13 +140,13 @@ fun polymul :: "poly \ poly \ poly" (infixl "*\<^sub>p" 60) where "polymul (C c) (C c') = C (c*\<^sub>Nc')" -| "polymul (C c) (CN c' n' p') = +| "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') = +| "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') = +| "polymul (CN c n p) (CN c' n' p') = (if np n' (polymul (CN c n p) p')))" | "polymul a b = Mul a b" @@ -157,7 +157,7 @@ fun polypow :: "nat \ poly \ poly" where "polypow 0 = (\p. (1)\<^sub>p)" -| "polypow n = (\p. let q = polypow (n div 2) p ; d = polymul q q in +| "polypow n = (\p. let q = polypow (n div 2) p ; d = polymul q q in if even n then d else polymul p d)" abbreviation poly_pow :: "poly \ nat \ poly" (infixl "^\<^sub>p" 60) @@ -196,13 +196,15 @@ partial_function (tailrec) polydivide_aux :: "poly \ nat \ poly \ nat \ poly \ nat \ poly" where - "polydivide_aux a n p k s = + "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')))))))" + 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')))))))" definition polydivide :: "poly \ poly \ (nat \ poly)" where "polydivide s p \ polydivide_aux (head p) (degree p) p 0 s" @@ -234,9 +236,9 @@ Ipoly_syntax :: "poly \ 'a list \'a::{field_char_0, field_inverse_zero, power}" ("\_\\<^sub>p\<^bsup>_\<^esup>") where "\p\\<^sub>p\<^bsup>bs\<^esup> \ Ipoly bs p" -lemma Ipoly_CInt: "Ipoly bs (C (i,1)) = of_int i" +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" +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 @@ -258,49 +260,52 @@ text{* polyadd preserves normal forms *} -lemma polyadd_normh: "\isnpolyh p n0 ; isnpolyh q n1\ +lemma polyadd_normh: "\isnpolyh p n0 ; isnpolyh q n1\ \ isnpolyh (polyadd p q) (min n0 n1)" -proof(induct p q arbitrary: n0 n1 rule: polyadd.induct) +proof (induct p q arbitrary: n0 n1 rule: polyadd.induct) case (2 ab c' n' p' n0 n1) - from 2 have th1: "isnpolyh (C ab) (Suc n')" by simp + from 2 have th1: "isnpolyh (C ab) (Suc n')" by simp from 2(3) have th2: "isnpolyh c' (Suc n')" and nplen1: "n' \ n1" by simp_all with isnpolyh_mono have cp: "isnpolyh c' (Suc n')" by simp with 2(1)[OF th1 th2] have th3:"isnpolyh (C ab +\<^sub>p c') (Suc n')" by simp - from nplen1 have n01len1: "min n0 n1 \ n'" by simp + from nplen1 have n01len1: "min n0 n1 \ n'" by simp thus ?case using 2 th3 by simp next case (3 c' n' p' ab n1 n0) - from 3 have th1: "isnpolyh (C ab) (Suc n')" by simp + from 3 have th1: "isnpolyh (C ab) (Suc n')" by simp from 3(2) have th2: "isnpolyh c' (Suc n')" and nplen1: "n' \ n1" by simp_all with isnpolyh_mono have cp: "isnpolyh c' (Suc n')" by simp with 3(1)[OF th2 th1] have th3:"isnpolyh (c' +\<^sub>p C ab) (Suc n')" by simp - from nplen1 have n01len1: "min n0 n1 \ n'" by simp + from nplen1 have n01len1: "min n0 n1 \ n'" by simp thus ?case using 3 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 4 have nc': "isnpolyh c' (Suc n')" and np': "isnpolyh p' n'" by simp_all + from 4 have nc': "isnpolyh c' (Suc n')" and np': "isnpolyh p' n'" by simp_all from 4 have ngen0: "n \ n0" by simp - from 4 have n'gen1: "n' \ n1" by simp + from 4 have n'gen1: "n' \ n1" by simp have "n < n' \ n' < n \ n = n'" by auto - moreover {assume eq: "n = n'" - with "4.hyps"(3)[OF nc nc'] + moreover { + assume eq: "n = n'" + with "4.hyps"(3)[OF 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 "4.hyps"(4)[OF np np'] have npp': "isnpolyh (p +\<^sub>p p') n" by simp have minle: "min n0 n1 \ n'" using ngen0 n'gen1 eq by simp - from minle npp' ncc'n01 4 eq ngen0 n'gen1 ncc' have ?case by (simp add: Let_def)} - moreover {assume lt: "n < n'" + from minle npp' ncc'n01 4 eq ngen0 n'gen1 ncc' have ?case by (simp add: Let_def) } + moreover { + assume lt: "n < n'" have "min n0 n1 \ n0" by simp - with 4 lt have th1:"min n0 n1 \ n" by auto + with 4 lt have th1:"min n0 n1 \ n" by auto from 4 have th21: "isnpolyh c (Suc n)" by simp from 4 have th22: "isnpolyh (CN c' n' p') n'" by simp from lt have th23: "min (Suc n) n' = Suc n" by arith from "4.hyps"(1)[OF th21 th22] have "isnpolyh (polyadd c (CN c' n' p')) (Suc n)" using th23 by simp - with 4 lt th1 have ?case by simp } - moreover {assume gt: "n' < n" hence gt': "n' < n \ \ n < n'" by simp + with 4 lt th1 have ?case by simp } + moreover { + assume gt: "n' < n" hence gt': "n' < n \ \ n < n'" by simp have "min n0 n1 \ n1" by simp with 4 gt have th1:"min n0 n1 \ n'" by auto from 4 have th21: "isnpolyh c' (Suc n')" by simp_all @@ -308,8 +313,8 @@ from gt have th23: "min n (Suc n') = Suc n'" by arith from "4.hyps"(2)[OF th22 th21] have "isnpolyh (polyadd (CN c n p) c') (Suc n')" using th23 by simp - with 4 gt th1 have ?case by simp} - ultimately show ?case by blast + with 4 gt 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" @@ -321,8 +326,8 @@ text{* The degree of addition and other general lemmas needed for the normal form of polymul *} -lemma polyadd_different_degreen: - "\isnpolyh p n0 ; isnpolyh q n1; degreen p m \ degreen q m ; m \ min n0 n1\ \ +lemma polyadd_different_degreen: + "\isnpolyh p n0 ; isnpolyh q n1; degreen p m \ degreen q m ; m \ min n0 n1\ \ 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) @@ -362,11 +367,13 @@ shows "degreen (p +\<^sub>p q) m \ 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 + 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 + 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) + case (4 c n p c' n' p' n0 n1 m) have "n' = n \ n < n' \ n' < n" by arith thus ?case proof (elim disjE) @@ -376,21 +383,21 @@ qed simp_all qed auto -lemma polyadd_eq_const_degreen: "\ isnpolyh p n0 ; isnpolyh q n1 ; polyadd p q = C c\ +lemma polyadd_eq_const_degreen: "\ isnpolyh p n0 ; isnpolyh q n1 ; polyadd p q = C c\ \ 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) - {assume nn': "n' < n" hence ?case using 4 by simp} - moreover - {assume nn':"\ n' < n" hence "n < n' \ n = n'" by arith - moreover {assume "n < n'" with 4 have ?case by simp } - moreover {assume eq: "n = n'" hence ?case using 4 + case (4 c n p c' n' p' m n0 n1 x) + { assume nn': "n' < n" hence ?case using 4 by simp } + moreover + { assume nn':"\ n' < n" hence "n < n' \ n = n'" by arith + moreover { assume "n < n'" with 4 have ?case by simp } + moreover { assume eq: "n = n'" hence ?case using 4 apply (cases "p +\<^sub>p p' = 0\<^sub>p") apply (auto simp add: Let_def) apply blast done - } - ultimately have ?case by blast} + } + ultimately have ?case by blast } ultimately show ?case by blast qed simp_all @@ -399,37 +406,37 @@ and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and m: "m \ min n0 n1" - shows "isnpolyh (p *\<^sub>p q) (min n0 n1)" - and "(p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \ q = 0\<^sub>p)" + shows "isnpolyh (p *\<^sub>p q) (min n0 n1)" + and "(p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \ q = 0\<^sub>p)" and "degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \ 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 c c' n' p') - { case (1 n0 n1) + case (2 c c' n' p') + { case (1 n0 n1) with "2.hyps"(4-6)[of n' n' n'] and "2.hyps"(1-3)[of "Suc n'" "Suc n'" n'] show ?case by (auto simp add: min_def) next - case (2 n0 n1) thus ?case by auto + case (2 n0 n1) thus ?case by auto next - case (3 n0 n1) thus ?case using "2.hyps" by auto } + case (3 n0 n1) thus ?case using "2.hyps" by auto } next case (3 c n p c') - { case (1 n0 n1) + { case (1 n0 n1) with "3.hyps"(4-6)[of n n n] "3.hyps"(1-3)[of "Suc n" "Suc n" n] show ?case by (auto simp add: min_def) next case (2 n0 n1) thus ?case by auto next - case (3 n0 n1) thus ?case using "3.hyps" by auto } + 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'" { case (1 n0 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 np': "isnpolyh p' n'" and nc': "isnpolyh c' (Suc n')" and nn0: "n \ n0" and nn1:"n' \ n1" by simp_all @@ -462,23 +469,24 @@ let ?d2 = "degreen ?cnp' m" let ?eq = "?d = (if ?cnp = 0\<^sub>p \ ?cnp' = 0\<^sub>p then 0 else ?d1 + ?d2)" have "n' n < n' \ n' = n" by auto - moreover + moreover {assume "n' < n \ n < n'" - with "4.hyps"(3,6,18) np np' m + with "4.hyps"(3,6,18) np np' m have ?eq by auto } moreover - {assume nn': "n' = n" hence nn:"\ n' < n \ \ n < n'" by arith + { assume nn': "n' = n" + hence nn: "\ n' < n \ \ n < n'" by arith from "4.hyps"(16,18)[of n n' n] "4.hyps"(13,14)[of n "Suc n'" 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 c' = 0\<^sub>p) = (c' = 0\<^sub>p)" "?cnp *\<^sub>p p' \ 0\<^sub>p" by (auto simp add: min_def) - {assume mn: "m = n" + { assume mn: "m = n" from "4.hyps"(17,18)[OF norm(1,4), of n] "4.hyps"(13,15)[OF norm(1,2), of n] norm nn' mn - have degs: "degreen (?cnp *\<^sub>p c') n = + 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 @@ -487,31 +495,31 @@ by simp have nmin: "n \ 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 + 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"(16-18)[OF norm(1,4), of n] "4.hyps"(13-15)[OF norm(1,2), of n] mn norm m nn' deg - have ?eq by simp} + have ?eq by simp } moreover - {assume mn: "m \ n" hence mn': "m < n" using m np by auto - from nn' m np have max1: "m \ max n n" by simp - hence min1: "m \ min n n" by simp + { assume mn: "m \ n" hence mn': "m < n" using m np by auto + from nn' m np have max1: "m \ max n n" by simp + hence min1: "m \ min n n" by simp hence min2: "m \ min n (Suc n)" by simp from "4.hyps"(16-18)[OF norm(1,4) min1] "4.hyps"(13-15)[OF 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 + have "degreen (?cnp *\<^sub>p c' +\<^sub>p CN 0\<^sub>p n (?cnp *\<^sub>p p')) m \ max (degreen (?cnp *\<^sub>p c') m) (degreen (CN 0\<^sub>p n (?cnp *\<^sub>p p')) m)" using mn nn' np np' by simp with "4.hyps"(16-18)[OF norm(1,4) min1] "4.hyps"(13-15)[OF norm(1,2) min2] degreen_0[OF norm(3) mn'] - have ?eq using nn' mn np np' by clarsimp} - ultimately have ?eq by blast} - ultimately show ?eq by blast} + have ?eq using nn' mn np np' by clarsimp } + ultimately have ?eq by blast } + ultimately show ?eq by blast } { case (2 n0 n1) - hence np: "isnpolyh ?cnp n0" and np': "isnpolyh ?cnp' n1" + hence np: "isnpolyh ?cnp n0" and np': "isnpolyh ?cnp' n1" and m: "m \ min n0 n1" by simp_all hence mn: "m \ n" by simp let ?c0p = "CN 0\<^sub>p n (?cnp *\<^sub>p p')" @@ -522,32 +530,32 @@ 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' \ 0\<^sub>p" + "(?cnp *\<^sub>p c' = 0\<^sub>p) = (c' = 0\<^sub>p)" + "?cnp *\<^sub>p p' \ 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" + 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} + 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: field_simps) -lemma polymul_normh: +lemma polymul_normh: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "\isnpolyh p n0 ; isnpolyh q n1\ \ isnpolyh (p *\<^sub>p q) (min n0 n1)" - using polymul_properties(1) by blast + using polymul_properties(1) by blast -lemma polymul_eq0_iff: +lemma polymul_eq0_iff: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "\ isnpolyh p n0 ; isnpolyh q n1\ \ (p *\<^sub>p q = 0\<^sub>p) = (p = 0\<^sub>p \ q = 0\<^sub>p) " - using polymul_properties(2) by blast + using polymul_properties(2) by blast lemma polymul_degreen: (* FIXME duplicate? *) assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" @@ -555,7 +563,7 @@ degreen (p *\<^sub>p q) m = (if (p = 0\<^sub>p \ q = 0\<^sub>p) then 0 else degreen p m + degreen q m)" using polymul_properties(3) by blast -lemma polymul_norm: +lemma polymul_norm: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "\ isnpoly p; isnpoly q\ \ isnpoly (polymul p q)" using polymul_normh[of "p" "0" "q" "0"] isnpoly_def by simp @@ -567,14 +575,14 @@ by (induct p arbitrary: n0) auto lemma monic_eqI: - assumes np: "isnpolyh p n0" + assumes np: "isnpolyh p n0" shows "INum (headconst p) * Ipoly bs (fst (monic p)) = (Ipoly bs p ::'a::{field_char_0, field_inverse_zero, power})" 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 \ 0\<^sub>p" - {assume hz: "INum ?h = (0::'a)" + { 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} @@ -602,18 +610,19 @@ lemma polysub[simp]: "Ipoly bs (polysub p q) = (Ipoly bs p) - (Ipoly bs q)" by (simp add: polysub_def) -lemma polysub_normh: "\ n0 n1. \ isnpolyh p n0 ; isnpolyh q n1\ \ isnpolyh (polysub p q) (min n0 n1)" +lemma polysub_normh: + "\n0 n1. \ isnpolyh p n0 ; isnpolyh q n1\ \ isnpolyh (polysub p q) (min n0 n1)" by (simp add: polysub_def polyneg_normh polyadd_normh) lemma polysub_norm: "\ isnpoly p; isnpoly q\ \ isnpoly (polysub p q)" - using polyadd_norm polyneg_norm by (simp add: polysub_def) + using polyadd_norm polyneg_norm by (simp add: polysub_def) lemma polysub_same_0[simp]: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "isnpolyh p n0 \ 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: +lemma polysub_0: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "\ isnpolyh p n0 ; isnpolyh q n1\ \ (p -\<^sub>p q = 0\<^sub>p) = (p = q)" unfolding polysub_def split_def fst_conv snd_conv @@ -631,8 +640,8 @@ let ?q = "polypow ((Suc n) div 2) p" let ?d = "polymul ?q ?q" have "odd (Suc n) \ even (Suc n)" by simp - moreover - {assume odd: "odd (Suc n)" + moreover + { assume odd: "odd (Suc n)" have th: "(Suc (Suc (Suc (0\nat)) * (Suc n div Suc (Suc (0\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) @@ -642,10 +651,10 @@ by (simp only: power_add power_one_right) simp also have "\ = (Ipoly bs p) ^ (Suc (Suc (Suc (0\nat)) * (Suc n div Suc (Suc (0\nat)))))" by (simp only: th) - finally have ?case + finally have ?case using odd_nat_div_two_times_two_plus_one[OF odd, symmetric] by simp } - moreover - {assume even: "even (Suc n)" + moreover + { assume even: "even (Suc n)" have th: "(Suc (Suc (0\nat))) * (Suc n div Suc (Suc (0\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) @@ -655,7 +664,7 @@ ultimately show ?case by blast qed -lemma polypow_normh: +lemma polypow_normh: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "isnpolyh p n \ isnpolyh (polypow k p) n" proof (induct k arbitrary: n rule: polypow.induct) @@ -666,9 +675,9 @@ 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 +qed auto -lemma polypow_norm: +lemma polypow_norm: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "isnpoly p \ isnpoly (polypow k p)" by (simp add: polypow_normh isnpoly_def) @@ -679,7 +688,7 @@ "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field_char_0, field_inverse_zero})" by (induct p rule:polynate.induct) auto -lemma polynate_norm[simp]: +lemma polynate_norm[simp]: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "isnpoly (polynate p)" by (induct p rule: polynate.induct) @@ -692,7 +701,7 @@ lemma shift1: "Ipoly bs (shift1 p) = Ipoly bs (Mul (Bound 0) p)" by (simp add: shift1_def) -lemma shift1_isnpoly: +lemma shift1_isnpoly: assumes pn: "isnpoly p" and pnz: "p \ 0\<^sub>p" shows "isnpoly (shift1 p) " @@ -700,11 +709,11 @@ lemma shift1_nz[simp]:"shift1 p \ 0\<^sub>p" by (simp add: shift1_def) -lemma funpow_shift1_isnpoly: +lemma funpow_shift1_isnpoly: "\ isnpoly p ; p \ 0\<^sub>p\ \ isnpoly (funpow n shift1 p)" by (induct n arbitrary: p) (auto simp add: shift1_isnpoly funpow_swap1) -lemma funpow_isnpolyh: +lemma funpow_isnpolyh: assumes f: "\ p. isnpolyh p n \ isnpolyh (f p) n" and np: "isnpolyh p n" shows "isnpolyh (funpow k f p) n" @@ -718,7 +727,7 @@ lemma shift1_isnpolyh: "isnpolyh p n0 \ p\ 0\<^sub>p \ 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: +lemma funpow_shift1_1: "(Ipoly bs (funpow n shift1 p) :: 'a :: {field_char_0, field_inverse_zero}) = Ipoly bs (funpow n shift1 (1)\<^sub>p *\<^sub>p p)" by (simp add: funpow_shift1) @@ -733,8 +742,8 @@ using np proof (induct p arbitrary: n rule: behead.induct) case (1 c p n) hence pn: "isnpolyh p n" by simp - from 1(1)[OF pn] - have th:"Ipoly bs (Add (Mul (head p) (Pw (Bound 0) (degree p))) (behead p)) = Ipoly bs p" . + from 1(1)[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") apply (simp_all add: th[symmetric] field_simps) @@ -778,7 +787,7 @@ assumes nb: "polybound0 a" shows "Ipoly (b#bs) a = Ipoly (b'#bs) a" using nb - by (induct a rule: poly.induct) auto + by (induct a rule: poly.induct) auto lemma polysubst0_I: "Ipoly (b#bs) (polysubst0 a t) = Ipoly ((Ipoly (b#bs) a)#bs) t" by (induct t) simp_all @@ -816,15 +825,15 @@ lemma wf_bs_coefficients: "wf_bs bs p \ \ c \ set (coefficients p). wf_bs bs c" proof (induct p rule: coefficients.induct) - case (1 c p) - show ?case + case (1 c p) + show ?case proof fix x assume xc: "x \ set (coefficients (CN c 0 p))" hence "x = c \ x \ set (coefficients p)" by simp - moreover + moreover {assume "x = c" hence "wf_bs bs x" using "1.prems" unfolding wf_bs_def by simp} - moreover - {assume H: "x \ set (coefficients p)" + moreover + {assume H: "x \ 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 @@ -838,7 +847,7 @@ unfolding wf_bs_def by (induct p) (auto simp add: nth_append) lemma take_maxindex_wf: - assumes wf: "wf_bs bs p" + assumes wf: "wf_bs bs p" shows "Ipoly (take (maxindex p) bs) p = Ipoly bs p" proof- let ?ip = "maxindex p" @@ -885,14 +894,14 @@ done lemma wf_bs_polyadd: "wf_bs bs p \ wf_bs bs q \ wf_bs bs (p +\<^sub>p q)" - unfolding wf_bs_def + 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 \ wf_bs bs q \ wf_bs bs (p *\<^sub>p q)" - unfolding wf_bs_def - apply (induct p q arbitrary: bs rule: polymul.induct) + 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]) @@ -918,12 +927,12 @@ have cp: "isnpolyh (CN c 0 p) n0" by fact hence norm: "isnpolyh c 0" "isnpolyh p 0" "p \ 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 + from "1.hyps"[OF norm(2)] norm(1) norm(4) show ?case by simp qed auto lemma coefficients_isconst: "isnpolyh p n \ \q\set (coefficients p). isconstant q" - by (induct p arbitrary: n rule: coefficients.induct) + by (induct p arbitrary: n rule: coefficients.induct) (auto simp add: isnpolyh_Suc_const) lemma polypoly_polypoly': @@ -940,17 +949,17 @@ hence "\q \ ?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 + + 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 + using np by (induct p arbitrary: n0 bs rule: coefficients.induct) (auto simp add: polypoly_def) -lemma polypoly'_poly: +lemma polypoly'_poly: assumes np: "isnpolyh p n0" shows "\p\\<^sub>p\<^bsup>x # bs\<^esup> = poly (polypoly' bs p) x" using polypoly_poly[OF np, simplified polypoly_polypoly'[OF np]] . @@ -959,14 +968,14 @@ 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 + using np nb unfolding polypoly_def apply (cases p) apply auto apply (case_tac nat) apply auto done -lemma head_isnpolyh: "isnpolyh p n0 \ isnpolyh (head p) n0" +lemma head_isnpolyh: "isnpolyh p n0 \ isnpolyh (head p) n0" by (induct p rule: head.induct) auto lemma headn_nz[simp]: "isnpolyh p n0 \ (headn p m = 0\<^sub>p) = (p = 0\<^sub>p)" @@ -978,7 +987,7 @@ lemma head_nz[simp]: "isnpolyh p n0 \ (head p = 0\<^sub>p) = (p = 0\<^sub>p)" by (simp add: head_eq_headn0) -lemma isnpolyh_zero_iff: +lemma isnpolyh_zero_iff: assumes nq: "isnpolyh p n0" and eq :"\bs. wf_bs bs p \ \p\\<^sub>p\<^bsup>bs\<^esup> = (0::'a::{field_char_0, field_inverse_zero, power})" shows "p = 0\<^sub>p" @@ -994,10 +1003,10 @@ 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" + 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 \ maxindex p" by auto with decr_maxindex[OF h(2)] nz have ihd_lt_n: "?ihd < maxindex p" by auto @@ -1023,21 +1032,21 @@ 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 "\?hd\\<^sub>p\<^bsup>bs\<^esup> = 0" by simp } then have hdz: "\bs. wf_bs bs ?hd \ \?hd\\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" by blast - + from less(1)[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: +lemma isnpolyh_unique: assumes np:"isnpolyh p n0" and nq: "isnpolyh q n1" shows "(\bs. \p\\<^sub>p\<^bsup>bs\<^esup> = (\q\\<^sub>p\<^bsup>bs\<^esup> :: 'a::{field_char_0, field_inverse_zero, power})) \ p = q" proof(auto) assume H: "\bs. (\p\\<^sub>p\<^bsup>bs\<^esup> ::'a)= \q\\<^sub>p\<^bsup>bs\<^esup>" hence "\bs.\p -\<^sub>p q\\<^sub>p\<^bsup>bs\<^esup>= (0::'a)" by simp - hence "\bs. wf_bs bs (p -\<^sub>p q) \ \p -\<^sub>p q\\<^sub>p\<^bsup>bs\<^esup> = (0::'a)" + hence "\bs. wf_bs bs (p -\<^sub>p q) \ \p -\<^sub>p q\\<^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 @@ -1056,28 +1065,28 @@ lemma zero_normh: "isnpolyh 0\<^sub>p n" by simp lemma one_normh: "isnpolyh (1)\<^sub>p n" by simp -lemma polyadd_0[simp]: +lemma polyadd_0[simp]: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" 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] + 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]: +lemma polymul_1[simp]: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" 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] + 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]: +lemma polymul_0[simp]: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" 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] + 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: +lemma polymul_commute: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" and np:"isnpolyh p n0" and nq: "isnpolyh q n1" @@ -1086,15 +1095,15 @@ by simp declare polyneg_polyneg [simp] - -lemma isnpolyh_polynate_id [simp]: + +lemma isnpolyh_polynate_id [simp]: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" and np:"isnpolyh p n0" shows "polynate p = p" using isnpolyh_unique[where ?'a= "'a::{field_char_0, field_inverse_zero}", OF polynate_norm[of p, unfolded isnpoly_def] np] polynate[where ?'a = "'a::{field_char_0, field_inverse_zero}"] by simp -lemma polynate_idempotent[simp]: +lemma polynate_idempotent[simp]: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "polynate (polynate p) = polynate p" using isnpolyh_polynate_id[OF polynate_norm[of p, unfolded isnpoly_def]] . @@ -1137,34 +1146,34 @@ from degree_polyadd[OF np nq'] show ?thesis by (simp add: polysub_def degree_polyneg[OF nq]) qed -lemma degree_polysub_samehead: +lemma degree_polysub_samehead: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" - and np: "isnpolyh p n0" and nq: "isnpolyh q n1" and h: "head p = head q" + 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 \ (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 c c') - thus ?case by (simp add: Nsub_def Nsub0[simplified Nsub_def]) + thus ?case by (simp add: Nsub_def Nsub0[simplified Nsub_def]) next - case (2 c c' n' p') + case (2 c c' n' p') from 2 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 2 show ?case by simp next - case (3 c n p c') + case (3 c n p c') hence "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 3 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" + 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" + 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 @@ -1178,10 +1187,10 @@ 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] - using nn' 4 by (simp add: Let_def)} + using nn' 4 by (simp add: Let_def) } ultimately have ?case by blast} moreover - {assume nn': "n < n'" hence n'p: "n' > 0" by simp + {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 4 nn' by (cases n', simp_all) @@ -1189,7 +1198,7 @@ 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 + {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 4 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 @@ -1198,7 +1207,7 @@ from H(3) headcnp headcnp' nn' have ?case by auto} ultimately show ?case by blast qed auto - + lemma shift1_head : "isnpolyh p n0 \ head (shift1 p) = head p" by (induct p arbitrary: n0 rule: head.induct) (simp_all add: shift1_def) @@ -1210,7 +1219,7 @@ case (Suc k n0 p) hence "isnpolyh (shift1 p) 0" by (simp add: shift1_isnpolyh) with Suc have "head (funpow k shift1 (shift1 p)) = head (shift1 p)" - and "head (shift1 p) = head p" by (simp_all add: shift1_head) + and "head (shift1 p) = head p" by (simp_all add: shift1_head) thus ?case by (simp add: funpow_swap1) qed @@ -1231,7 +1240,7 @@ lemma head_head[simp]: "isnpolyh p n0 \ head (head p) = head p" by (induct p rule: head.induct) auto -lemma polyadd_eq_const_degree: +lemma polyadd_eq_const_degree: "isnpolyh p n0 \ isnpolyh q n1 \ polyadd p q = C c \ degree p = degree q" using polyadd_eq_const_degreen degree_eq_degreen0 by simp @@ -1255,15 +1264,15 @@ apply (metis degree.simps(9) gr0_conv_Suc nat_less_le order_le_less_trans) done -lemma polymul_head_polyeq: +lemma polymul_head_polyeq: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" shows "\isnpolyh p n0; isnpolyh q n1 ; p \ 0\<^sub>p ; q \ 0\<^sub>p \ \ head (p *\<^sub>p q) = head p *\<^sub>p head q" proof (induct p q arbitrary: n0 n1 rule: polymul.induct) case (2 c c' n' p' n0 n1) hence "isnpolyh (head (CN c' n' p')) n1" "isnormNum c" by (simp_all add: head_isnpolyh) thus ?case using 2 by (cases n') auto -next - case (3 c n p c' n0 n1) +next + case (3 c n p c' n0 n1) hence "isnpolyh (head (CN c n p)) n0" "isnormNum c'" by (simp_all add: head_isnpolyh) thus ?case using 3 by (cases n) auto next @@ -1272,8 +1281,8 @@ "isnpolyh (CN c n p) n" "isnpolyh (CN c' n' p') n'" by simp_all have "n < n' \ n' < n \ n = n'" by arith - moreover - {assume nn': "n < n'" hence ?case + moreover + {assume nn': "n < n'" hence ?case using norm "4.hyps"(2)[OF norm(1,6)] "4.hyps"(1)[OF norm(2,6)] apply simp apply (cases n) @@ -1283,7 +1292,7 @@ done } moreover {assume nn': "n'< n" hence ?case - using norm "4.hyps"(6) [OF norm(5,3)] "4.hyps"(5)[OF norm(5,4)] + using norm "4.hyps"(6) [OF norm(5,3)] "4.hyps"(5)[OF norm(5,4)] apply simp apply (cases n') apply simp @@ -1291,14 +1300,14 @@ apply auto done } moreover {assume nn': "n' = n" - from nn' polymul_normh[OF norm(5,4)] + 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 + 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" + 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] @@ -1314,7 +1323,7 @@ from polyadd_head[OF ncnpc'[simplified nz] ncnpp0'[simplified nz] dth'] dth have ?case using norm "4.hyps"(6)[OF norm(5,3)] "4.hyps"(5)[OF norm(5,4)] nn' nz by simp } - ultimately have ?case by (cases n) auto} + ultimately have ?case by (cases n) auto} ultimately show ?case by blast qed simp_all @@ -1359,25 +1368,29 @@ and ns: "isnpolyh s n1" and ap: "head p = a" and ndp: "degree p = n" and pnz: "p \ 0\<^sub>p" - shows "(polydivide_aux a n p k s = (k',r) \ (k' \ k) \ (degree r = 0 \ degree r < degree p) + shows "(polydivide_aux a n p k s = (k',r) \ (k' \ k) \ (degree r = 0 \ degree r < degree p) \ (\nr. isnpolyh r nr) \ (\q n1. isnpolyh q n1 \ ((polypow (k' - k) a) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)))" using ns proof (induct "degree s" arbitrary: s k k' r n1 rule: less_induct) case less let ?qths = "\q n1. isnpolyh q n1 \ (a ^\<^sub>p (k' - k) *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)" - let ?ths = "polydivide_aux a n p k s = (k', r) \ k \ k' \ (degree r = 0 \ degree r < degree p) + let ?ths = "polydivide_aux a n p k s = (k', r) \ k \ k' \ (degree r = 0 \ degree r < degree p) \ (\nr. isnpolyh r nr) \ ?qths" let ?b = "head s" let ?p' = "funpow (degree s - n) shift1 p" let ?xdn = "funpow (degree s - n) shift1 (1)\<^sub>p" let ?akk' = "a ^\<^sub>p (k' - k)" note ns = `isnpolyh s n1` - 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="degree s - 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 + 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="degree s - 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 ?ths using np polydivide_aux.simps @@ -1386,67 +1399,82 @@ apply simp done } moreover - {assume sz: "s \ 0\<^sub>p" - {assume dn: "degree s < n" + { assume sz: "s \ 0\<^sub>p" + { assume dn: "degree s < n" hence "?ths" using ns ndp np polydivide_aux.simps apply auto apply (rule exI[where x="0\<^sub>p"]) apply simp done } - moreover - {assume dn': "\ degree s < n" hence dn: "degree s \ n" by arith - have degsp': "degree s = degree ?p'" + moreover + { assume dn': "\ degree s < n" hence dn: "degree s \ n" by arith + have degsp': "degree s = degree ?p'" using dn ndp funpow_shift1_degree[where k = "degree s - 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 + { 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 degree_polysub_samehead[OF ns np' headsp' degsp'] have "degree (s -\<^sub>p ?p') < degree s \ s -\<^sub>p ?p' = 0\<^sub>p" by simp - moreover - {assume deglt:"degree (s -\<^sub>p ?p') < degree s" + moreover + { assume deglt:"degree (s -\<^sub>p ?p') < degree s" from polydivide_aux.simps sz dn' ba 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 less(1)[OF deglt nr, of k k' r] - trans[OF eq[symmetric] h1] - have kk': "k \ k'" and nr:"\nr. isnpolyh r nr" and dr: "degree r = 0 \ degree r < degree p" - and q1:"\q nq. isnpolyh q nq \ (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 + { assume h1: "polydivide_aux a n p k s = (k', r)" + from less(1)[OF deglt nr, of k k' r] trans[OF eq[symmetric] h1] + have kk': "k \ k'" + and nr:"\nr. isnpolyh r nr" + and dr: "degree r = 0 \ degree r < degree p" + and q1: "\q nq. isnpolyh q nq \ (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 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 + 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 "\ (bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a^\<^sub>p (k' - k) *\<^sub>p (s -\<^sub>p ?p')) = + have nqr': "isnpolyh (p *\<^sub>p (?akk' *\<^sub>p ?xdn +\<^sub>p q) +\<^sub>p r) 0" + by simp + from asp have "\ (bs:: 'a::{field_char_0, field_inverse_zero} 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 " \(bs:: 'a::{field_char_0, field_inverse_zero} 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" + hence " \(bs:: 'a::{field_char_0, field_inverse_zero} 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: field_simps) - hence " \(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = - Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - 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 "\(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = - Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - n) shift1 (1)\<^sub>p) - + Ipoly bs q) + Ipoly bs r" by (simp add: field_simps) - hence "\(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = - Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - n) shift1 (1)\<^sub>p) +\<^sub>p q) +\<^sub>p r)" by simp + hence " \(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = + Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - 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 "\(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = + Ipoly bs p * (Ipoly bs (a^\<^sub>p (k' - k)) * Ipoly bs (funpow (degree s - n) shift1 (1)\<^sub>p) + + Ipoly bs q) + Ipoly bs r" + by (simp add: field_simps) + hence "\(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a ^\<^sub>p (k' - k) *\<^sub>p s) = + Ipoly bs (p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - 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 (degree s - n) shift1 (1)\<^sub>p) +\<^sub>p q) +\<^sub>p r" by blast + have "a ^\<^sub>p (k' - k) *\<^sub>p s = + p *\<^sub>p ((a^\<^sub>p (k' - k)) *\<^sub>p (funpow (degree s - 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 (degree s - n) shift1 (1)\<^sub>p) +\<^sub>p q" in exI) - apply (rule_tac x="0" in exI) by simp + apply (rule_tac x="0" in exI) + apply simp + done with kk' nr dr have "k \ k' \ (degree r = 0 \ degree r < degree p) \ (\nr. isnpolyh r nr) \ ?qths" - by blast } hence ?ths by blast } - moreover - {assume spz:"s -\<^sub>p ?p' = 0\<^sub>p" + by blast + } + hence ?ths by blast + } + moreover + { assume spz:"s -\<^sub>p ?p' = 0\<^sub>p" from spz isnpolyh_unique[OF polysub_normh[OF ns np'], where q="0\<^sub>p", symmetric, where ?'a = "'a::{field_char_0, field_inverse_zero}"] - have " \(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs s = Ipoly bs ?p'" by simp + have " \(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs s = Ipoly bs ?p'" + by simp hence "\(bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs s = Ipoly bs (?xdn *\<^sub>p p)" using np nxdn apply simp @@ -1455,134 +1483,162 @@ done 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)" + { assume h1: "polydivide_aux a n p k s = (k',r)" from polydivide_aux.simps sz dn' ba 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 "\ = (k,0\<^sub>p)" using polydivide_aux.simps spz by simp + also have "\ = (k,0\<^sub>p)" + using polydivide_aux.simps spz by simp finally have "(k',r) = (k,0\<^sub>p)" using h1 by simp with sp'[symmetric] ns np nxdn polyadd_0(1)[OF polymul_normh[OF np nxdn]] polyadd_0(2)[OF polymul_normh[OF np nxdn]] have ?ths apply auto - apply (rule exI[where x="?xdn"]) + apply (rule exI[where x="?xdn"]) apply (auto simp add: polymul_commute[of p]) - done} } - ultimately have ?ths by blast } + done + } + } + ultimately have ?ths by blast + } moreover - {assume ba: "?b \ a" - from polysub_normh[OF polymul_normh[OF head_isnpolyh[OF np0, simplified ap] ns] + { assume ba: "?b \ 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 nth: "isnpolyh ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) 0" + by (simp add: min_def) have nzths: "a *\<^sub>p s \ 0\<^sub>p" "?b *\<^sub>p ?p' \ 0\<^sub>p" - using polymul_eq0_iff[OF head_isnpolyh[OF np0, simplified ap] ns] + 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 + 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="degree s - n"] - have hdth: "head (a *\<^sub>p s) = head (?b *\<^sub>p ?p')" + 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="degree s - n"] - polymul_degreen[OF head_isnpolyh[OF ns] np', where m="0"] head_nz[OF ns] + polymul_degreen[OF head_isnpolyh[OF ns] np', where m="0"] head_nz[OF ns] ndp dn - have degth: "degree (a *\<^sub>p s) = degree (?b *\<^sub>p ?p') " + 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')) < degree s" - 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)" + { assume dth: "degree ((a *\<^sub>p s) -\<^sub>p (?b *\<^sub>p ?p')) < degree s" + 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.simps sz dn' ba 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 less(1)[OF dth nasbp', of "Suc k" k' r] - obtain q nq nr where kk': "Suc k \ k'" and nr: "isnpolyh r nr" and nq: "isnpolyh q nq" + obtain q nq nr where kk': "Suc k \ k'" + and nr: "isnpolyh r nr" + and nq: "isnpolyh q nq" and dr: "degree r = 0 \ 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 + 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::{field_char_0, field_inverse_zero} 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: field_simps) - 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="degree s - 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: field_simps)} - hence ieq:"\(bs :: 'a::{field_char_0, field_inverse_zero} 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 + { + fix bs:: "'a::{field_char_0, field_inverse_zero} 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: field_simps) + 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="degree s - 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: field_simps) + } + hence ieq:"\(bs :: 'a::{field_char_0, field_inverse_zero} 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 + 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 ?ths apply simp + 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 ?ths + 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 ?ths by blast } - moreover - {assume spz: "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p" - {fix bs :: "'a::{field_char_0, field_inverse_zero} list" + done + } + hence ?ths by blast + } + moreover + { assume spz: "a *\<^sub>p s -\<^sub>p (?b *\<^sub>p ?p') = 0\<^sub>p" + { + fix bs :: "'a::{field_char_0, field_inverse_zero} 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="degree s - n" and p="p"]) - hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" by simp - } - hence hth: "\ (bs:: 'a::{field_char_0, field_inverse_zero} list). Ipoly bs (a*\<^sub>p s) = + 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="degree s - n" and p="p"]) + hence "Ipoly bs (a*\<^sub>p s) = Ipoly bs (p *\<^sub>p (?b *\<^sub>p ?xdn))" + by simp + } + hence hth: "\ (bs:: 'a::{field_char_0, field_inverse_zero} 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::{field_char_0, field_inverse_zero}", OF polymul_normh[OF head_isnpolyh[OF np] ns] + from hth have asq: "a *\<^sub>p s = p *\<^sub>p (?b *\<^sub>p ?xdn)" + using isnpolyh_unique[where ?'a = "'a::{field_char_0, field_inverse_zero}", 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 ba dn' spz polydivide_aux.simps polydivide_aux.simps - 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 ?ths - 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 ?ths by blast } + { assume h1: "polydivide_aux a n p k s = (k', r)" + from h1 sz ba dn' spz polydivide_aux.simps polydivide_aux.simps + 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 ?ths + 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 ?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] - by (simp add: degree_eq_degreen0[symmetric]) blast } + by (simp add: degree_eq_degreen0[symmetric]) blast + } ultimately have ?ths by blast } - ultimately have ?ths by blast } + ultimately have ?ths by blast + } ultimately show ?ths by blast qed -lemma polydivide_properties: +lemma polydivide_properties: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" - and np: "isnpolyh p n0" and ns: "isnpolyh s n1" and pnz: "p \ 0\<^sub>p" - shows "(\ k r. polydivide s p = (k,r) \ (\nr. isnpolyh r nr) \ (degree r = 0 \ degree r < degree p) - \ (\q n1. isnpolyh q n1 \ ((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_def[where s="s" and p="p"] - have ex: "\ k r. polydivide s p = (k,r)" by auto - then obtain k r where kr: "polydivide s p = (k,r)" by blast + and np: "isnpolyh p n0" and ns: "isnpolyh s n1" and pnz: "p \ 0\<^sub>p" + shows "\k r. polydivide s p = (k,r) \ + (\nr. isnpolyh r nr) \ (degree r = 0 \ degree r < degree p) \ + (\q n1. isnpolyh q n1 \ ((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_def[where s="s" and p="p"] have ex: "\ 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 \ degree r < degree p) \ - (\nr. isnpolyh r nr) \ (\q n1. isnpolyh q n1 \ head p ^\<^sub>p k - 0 *\<^sub>p s = p *\<^sub>p q +\<^sub>p r)" by blast - with kr show ?thesis + (\nr. isnpolyh r nr) \ (\q n1. isnpolyh q n1 \ 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"]) @@ -1596,23 +1652,23 @@ definition "isnonconstant p = (\ isconstant p)" lemma isnonconstant_pnormal_iff: - assumes nc: "isnonconstant p" - shows "pnormal (polypoly bs p) \ Ipoly bs (head p) \ 0" + assumes nc: "isnonconstant p" + shows "pnormal (polypoly bs p) \ Ipoly bs (head p) \ 0" proof - let ?p = "polypoly bs p" + let ?p = "polypoly bs p" assume H: "pnormal ?p" have csz: "coefficients p \ []" using nc by (cases p) auto - - from coefficients_head[of p] last_map[OF csz, of "Ipoly bs"] + + from coefficients_head[of p] last_map[OF csz, of "Ipoly bs"] pnormal_last_nonzero[OF H] show "Ipoly bs (head p) \ 0" by (simp add: polypoly_def) next assume h: "\head p\\<^sub>p\<^bsup>bs\<^esup> \ 0" let ?p = "polypoly bs p" have csz: "coefficients p \ []" using nc by (cases p) auto - hence pz: "?p \ []" by (simp add: polypoly_def) + hence pz: "?p \ []" 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"] + from h coefficients_head[of p] last_map[OF csz, of "Ipoly bs"] have lz: "last ?p \ 0" by (simp add: polypoly_def) from pnormal_last_length[OF lg lz] show "pnormal ?p" . qed @@ -1638,10 +1694,10 @@ assume h: "\head p\\<^sub>p\<^bsup>bs\<^esup> \ 0" from isnonconstant_pnormal_iff[OF inc, of bs] h have pn: "pnormal ?p" by blast - {fix x assume H: "?p = [x]" + { 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 + 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\[] \ pnormal p \ length (pnormalize p) = length p" @@ -1655,29 +1711,29 @@ assumes inc: "isnonconstant p" shows "degree p = Polynomial_List.degree (polypoly bs p) \ \head p\\<^sub>p\<^bsup>bs\<^esup> \ 0" proof - let ?p = "polypoly bs p" + let ?p = "polypoly bs p" assume H: "degree p = Polynomial_List.degree ?p" from isnonconstant_coefficients_length[OF inc] have pz: "?p \ []" 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] + hence "pnormal ?p" using pnormal_length[OF pz] by blast + with isnonconstant_pnormal_iff[OF inc] show "\head p\\<^sub>p\<^bsup>bs\<^esup> \ 0" by blast next - let ?p = "polypoly bs p" + let ?p = "polypoly bs p" assume H: "\head p\\<^sub>p\<^bsup>bs\<^esup> \ 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" + 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 *} +section {* Swaps ; Division by a certain variable *} -primrec swap:: "nat \ nat \ poly \ poly" where +primrec swap :: "nat \ nat \ poly \ poly" where "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)" @@ -1685,8 +1741,8 @@ | "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)" +| "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" @@ -1694,10 +1750,10 @@ 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 + thus ?case using nbs mbs by simp next case (CN c k p) - thus ?case using nbs mbs by simp + thus ?case using nbs mbs by simp qed simp_all lemma swap_swap_id [simp]: "swap n m (swap m n t) = t" @@ -1723,9 +1779,9 @@ shows "isnpoly (swapnorm n m p)" unfolding swapnorm_def by simp -definition "polydivideby n s p = +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))" + 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 @@ -1736,10 +1792,10 @@ | "isweaknpoly (CN c n p) \ isweaknpoly c \ isweaknpoly p" | "isweaknpoly p = False" -lemma isnpolyh_isweaknpoly: "isnpolyh p n0 \ isweaknpoly p" +lemma isnpolyh_isweaknpoly: "isnpolyh p n0 \ isweaknpoly p" by (induct p arbitrary: n0) auto -lemma swap_isweanpoly: "isweaknpoly p \ isweaknpoly (swap n m p)" +lemma swap_isweanpoly: "isweaknpoly p \ isweaknpoly (swap n m p)" by (induct p) auto end \ No newline at end of file