# HG changeset patch # User krauss # Date 1293311938 -3600 # Node ID aae9f912cca8f6bbf46f3a3cf5131bd29465e1e2 # Parent 7eba049f7310b148ec1da3584c85406dbd3f25ec dropped duplicate unused lemmas; spelling diff -r 7eba049f7310 -r aae9f912cca8 src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy --- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Sat Dec 25 22:18:55 2010 +0100 +++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Sat Dec 25 22:18:58 2010 +0100 @@ -321,7 +321,7 @@ lemma polyadd_norm: "\ isnpoly p ; isnpoly q\ \ 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*} +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\ \ @@ -598,9 +598,8 @@ qed - +text{* polyneg is a negation and preserves normal forms *} -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) @@ -615,7 +614,8 @@ using isnpoly_def polyneg_normh by simp -text{* polysub is a substraction and preserves normalform *} +text{* polysub is a substraction and preserves normal forms *} + lemma polysub[simp]: "Ipoly bs (polysub (p,q)) = (Ipoly bs p) - (Ipoly bs q)" by (simp add: polysub_def polyneg polyadd) lemma polysub_normh: "\ n0 n1. \ isnpolyh p n0 ; isnpolyh q n1\ \ isnpolyh (polysub(p,q)) (min n0 n1)" @@ -651,6 +651,7 @@ done text{* polypow is a power function and preserves normal forms *} + lemma polypow[simp]: "Ipoly bs (polypow n p) = ((Ipoly bs p :: 'a::{field_char_0, field_inverse_zero})) ^ n" proof(induct n rule: polypow.induct) case 1 thus ?case by simp @@ -699,7 +700,7 @@ shows "isnpoly p \ isnpoly (polypow k p)" by (simp add: polypow_normh isnpoly_def) -text{* Finally the whole normalization*} +text{* Finally the whole normalization *} lemma polynate[simp]: "Ipoly bs (polynate p) = (Ipoly bs p :: 'a ::{field_char_0, field_inverse_zero})" by (induct p rule:polynate.induct, auto) @@ -759,7 +760,7 @@ 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 ... *} +subsection{* Miscellaneous lemmas about indexes, decrementation, substitution etc ... *} lemma isnpolyh_polybound0: "isnpolyh p (Suc n) \ polybound0 p" proof(induct p arbitrary: n rule: poly.induct, auto) case (goal1 c n p n') @@ -841,8 +842,6 @@ lemma maxindex_coefficients: " \c\ set (coefficients p). maxindex c \ maxindex p" by (induct p rule: coefficients.induct, auto) -lemma length_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) @@ -1038,7 +1037,7 @@ qed -text{* consequenses of unicity on the algorithms for polynomial normalization *} +text{* consequences of unicity on the algorithms for polynomial normalization *} lemma polyadd_commute: assumes "SORT_CONSTRAINT('a::{field_char_0, field_inverse_zero})" and np: "isnpolyh p n0" and nq: "isnpolyh q n1" shows "p +\<^sub>p q = q +\<^sub>p p" @@ -1525,8 +1524,6 @@ definition "isnonconstant p = (\ isconstant p)" -lemma last_map: "xs \ [] ==> last (map f xs) = f (last xs)" by (induct xs, auto) - lemma isnonconstant_pnormal_iff: assumes nc: "isnonconstant p" shows "pnormal (polypoly bs p) \ Ipoly bs (head p) \ 0" proof