src/HOL/Computational_Algebra/Polynomial_FPS.thy
author wenzelm
Sat, 09 Mar 2024 16:52:08 +0100
changeset 79833 d71af537a6e9
parent 69597 ff784d5a5bfb
child 81107 ad5fc948e053
permissions -rw-r--r--
obsolete;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
65435
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     1
(*  Title:      HOL/Computational_Algebra/Polynomial_FPS.thy
65366
10ca63a18e56 proper imports;
wenzelm
parents: 64911
diff changeset
     2
    Author:     Manuel Eberl, TU München
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
     3
*)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
     4
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
     5
section \<open>Converting polynomials to formal power series\<close>
65366
10ca63a18e56 proper imports;
wenzelm
parents: 64911
diff changeset
     6
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
     7
theory Polynomial_FPS
65366
10ca63a18e56 proper imports;
wenzelm
parents: 64911
diff changeset
     8
  imports Polynomial Formal_Power_Series
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
     9
begin
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    10
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    11
context
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    12
  includes fps_notation
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    13
begin
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    14
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    15
definition fps_of_poly where
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    16
  "fps_of_poly p = Abs_fps (coeff p)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    17
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    18
lemma fps_of_poly_eq_iff: "fps_of_poly p = fps_of_poly q \<longleftrightarrow> p = q"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    19
  by (simp add: fps_of_poly_def poly_eq_iff fps_eq_iff)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    20
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    21
lemma fps_of_poly_nth [simp]: "fps_of_poly p $ n = coeff p n"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    22
  by (simp add: fps_of_poly_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    23
  
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    24
lemma fps_of_poly_const: "fps_of_poly [:c:] = fps_const c"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    25
proof (subst fps_eq_iff, clarify)
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    26
  fix n :: nat show "fps_of_poly [:c:] $ n = fps_const c $ n"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    27
    by (cases n) (auto simp: fps_of_poly_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    28
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    29
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    30
lemma fps_of_poly_0 [simp]: "fps_of_poly 0 = 0"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    31
  by (subst fps_const_0_eq_0 [symmetric], subst fps_of_poly_const [symmetric]) simp
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    32
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    33
lemma fps_of_poly_1 [simp]: "fps_of_poly 1 = 1"
65486
d801126a14cb more systematic treatment of polynomial 1
haftmann
parents: 65435
diff changeset
    34
  by (simp add: fps_eq_iff)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    35
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    36
lemma fps_of_poly_1' [simp]: "fps_of_poly [:1:] = 1"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    37
  by (subst fps_const_1_eq_1 [symmetric], subst fps_of_poly_const [symmetric])
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    38
     (simp add: one_poly_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    39
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    40
lemma fps_of_poly_numeral [simp]: "fps_of_poly (numeral n) = numeral n"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    41
  by (simp add: numeral_fps_const fps_of_poly_const [symmetric] numeral_poly)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    42
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    43
lemma fps_of_poly_numeral' [simp]: "fps_of_poly [:numeral n:] = numeral n"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    44
  by (simp add: numeral_fps_const fps_of_poly_const [symmetric] numeral_poly)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    45
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    46
lemma fps_of_poly_fps_X [simp]: "fps_of_poly [:0, 1:] = fps_X"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    47
  by (auto simp add: fps_of_poly_def fps_eq_iff coeff_pCons split: nat.split)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    48
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    49
lemma fps_of_poly_add: "fps_of_poly (p + q) = fps_of_poly p + fps_of_poly q"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    50
  by (simp add: fps_of_poly_def plus_poly.rep_eq fps_plus_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    51
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    52
lemma fps_of_poly_diff: "fps_of_poly (p - q) = fps_of_poly p - fps_of_poly q"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    53
  by (simp add: fps_of_poly_def minus_poly.rep_eq fps_minus_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    54
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    55
lemma fps_of_poly_uminus: "fps_of_poly (-p) = -fps_of_poly p"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    56
  by (simp add: fps_of_poly_def uminus_poly.rep_eq fps_uminus_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    57
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    58
lemma fps_of_poly_mult: "fps_of_poly (p * q) = fps_of_poly p * fps_of_poly q"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    59
  by (simp add: fps_of_poly_def fps_times_def fps_eq_iff coeff_mult atLeast0AtMost)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    60
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    61
lemma fps_of_poly_smult: 
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    62
  "fps_of_poly (smult c p) = fps_const c * fps_of_poly p"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    63
  using fps_of_poly_mult[of "[:c:]" p] by (simp add: fps_of_poly_mult fps_of_poly_const)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    64
  
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
    65
lemma fps_of_poly_sum: "fps_of_poly (sum f A) = sum (\<lambda>x. fps_of_poly (f x)) A"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    66
  by (cases "finite A", induction rule: finite_induct) (simp_all add: fps_of_poly_add)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    67
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63539
diff changeset
    68
lemma fps_of_poly_sum_list: "fps_of_poly (sum_list xs) = sum_list (map fps_of_poly xs)"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    69
  by (induction xs) (simp_all add: fps_of_poly_add)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    70
  
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
    71
lemma fps_of_poly_prod: "fps_of_poly (prod f A) = prod (\<lambda>x. fps_of_poly (f x)) A"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    72
  by (cases "finite A", induction rule: finite_induct) (simp_all add: fps_of_poly_mult)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    73
  
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63539
diff changeset
    74
lemma fps_of_poly_prod_list: "fps_of_poly (prod_list xs) = prod_list (map fps_of_poly xs)"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    75
  by (induction xs) (simp_all add: fps_of_poly_mult)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    76
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    77
lemma fps_of_poly_pCons: 
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    78
  "fps_of_poly (pCons (c :: 'a :: semiring_1) p) = fps_const c + fps_of_poly p * fps_X"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    79
  by (subst fps_mult_fps_X_commute [symmetric], intro fps_ext) 
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    80
     (auto simp: fps_of_poly_def coeff_pCons split: nat.split)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    81
  
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    82
lemma fps_of_poly_pderiv: "fps_of_poly (pderiv p) = fps_deriv (fps_of_poly p)"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    83
  by (intro fps_ext) (simp add: fps_of_poly_nth coeff_pderiv)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    84
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    85
lemma fps_of_poly_power: "fps_of_poly (p ^ n) = fps_of_poly p ^ n"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    86
  by (induction n) (simp_all add: fps_of_poly_mult)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    87
  
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    88
lemma fps_of_poly_monom: "fps_of_poly (monom (c :: 'a :: comm_ring_1) n) = fps_const c * fps_X ^ n"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    89
  by (intro fps_ext) simp_all
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    90
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
    91
lemma fps_of_poly_monom': "fps_of_poly (monom (1 :: 'a :: comm_ring_1) n) = fps_X ^ n"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    92
  by (simp add: fps_of_poly_monom)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    93
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    94
lemma fps_of_poly_div:
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    95
  assumes "(q :: 'a :: field poly) dvd p"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    96
  shows   "fps_of_poly (p div q) = fps_of_poly p / fps_of_poly q"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    97
proof (cases "q = 0")
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
    98
  case False
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
    99
  from False fps_of_poly_eq_iff[of q 0] have nz: "fps_of_poly q \<noteq> 0" by simp 
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   100
  from assms have "p = (p div q) * q" by simp
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   101
  also have "fps_of_poly \<dots> = fps_of_poly (p div q) * fps_of_poly q" 
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   102
    by (simp add: fps_of_poly_mult)
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   103
  also from nz have "\<dots> / fps_of_poly q = fps_of_poly (p div q)"
64240
eabf80376aab more standardized names
haftmann
parents: 63882
diff changeset
   104
    by (intro nonzero_mult_div_cancel_right) (auto simp: fps_of_poly_0)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   105
  finally show ?thesis ..
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   106
qed simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   107
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   108
lemma fps_of_poly_divide_numeral:
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   109
  "fps_of_poly (smult (inverse (numeral c :: 'a :: field)) p) = fps_of_poly p / numeral c"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   110
proof -
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   111
  have "smult (inverse (numeral c)) p = [:inverse (numeral c):] * p" by simp
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   112
  also have "fps_of_poly \<dots> = fps_of_poly p / numeral c"
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   113
    by (subst fps_of_poly_mult) (simp add: numeral_fps_const fps_of_poly_pCons)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   114
  finally show ?thesis by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   115
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   116
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   117
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   118
lemma subdegree_fps_of_poly:
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   119
  assumes "p \<noteq> 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   120
  defines "n \<equiv> Polynomial.order 0 p"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   121
  shows   "subdegree (fps_of_poly p) = n"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   122
proof (rule subdegreeI)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   123
  from assms have "monom 1 n dvd p" by (simp add: monom_1_dvd_iff)
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   124
  thus zero: "fps_of_poly p $ i = 0" if "i < n" for i
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   125
    using that by (simp add: monom_1_dvd_iff')
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   126
    
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   127
  from assms have "\<not>monom 1 (Suc n) dvd p"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   128
    by (auto simp: monom_1_dvd_iff simp del: power_Suc)
63539
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63319
diff changeset
   129
  then obtain k where k: "k \<le> n" "fps_of_poly p $ k \<noteq> 0" 
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   130
    by (auto simp: monom_1_dvd_iff' less_Suc_eq_le)
63539
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63319
diff changeset
   131
  with zero[of k] have "k = n" by linarith
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63319
diff changeset
   132
  with k show "fps_of_poly p $ n \<noteq> 0" by simp
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   133
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   134
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   135
lemma fps_of_poly_dvd:
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   136
  assumes "p dvd q"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   137
  shows   "fps_of_poly (p :: 'a :: field poly) dvd fps_of_poly q"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   138
proof (cases "p = 0 \<or> q = 0")
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   139
  case False
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   140
  with assms fps_of_poly_eq_iff[of p 0] fps_of_poly_eq_iff[of q 0] show ?thesis
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   141
    by (auto simp: fps_dvd_iff subdegree_fps_of_poly dvd_imp_order_le)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   142
qed (insert assms, auto)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   143
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   144
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   145
lemmas fps_of_poly_simps =
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
   146
  fps_of_poly_0 fps_of_poly_1 fps_of_poly_numeral fps_of_poly_const fps_of_poly_fps_X
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   147
  fps_of_poly_add fps_of_poly_diff fps_of_poly_uminus fps_of_poly_mult fps_of_poly_smult
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   148
  fps_of_poly_sum fps_of_poly_sum_list fps_of_poly_prod fps_of_poly_prod_list
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   149
  fps_of_poly_pCons fps_of_poly_pderiv fps_of_poly_power fps_of_poly_monom
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   150
  fps_of_poly_divide_numeral
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   151
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   152
lemma fps_of_poly_pcompose:
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   153
  assumes "coeff q 0 = (0 :: 'a :: idom)"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   154
  shows   "fps_of_poly (pcompose p q) = fps_compose (fps_of_poly p) (fps_of_poly q)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   155
  using assms by (induction p rule: pCons_induct)
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   156
                 (auto simp: pcompose_pCons fps_of_poly_simps fps_of_poly_pCons 
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   157
                             fps_compose_add_distrib fps_compose_mult_distrib)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   158
  
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   159
lemmas reify_fps_atom =
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
   160
  fps_of_poly_0 fps_of_poly_1' fps_of_poly_numeral' fps_of_poly_const fps_of_poly_fps_X
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   161
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   162
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   163
text \<open>
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   164
  The following simproc can reduce the equality of two polynomial FPSs two equality of the
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   165
  respective polynomials. A polynomial FPS is one that only has finitely many non-zero 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   166
  coefficients and can therefore be written as \<^term>\<open>fps_of_poly p\<close> for some 
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64272
diff changeset
   167
  polynomial \<open>p\<close>.
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   168
  
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   169
  This may sound trivial, but it covers a number of annoying side conditions like 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   170
  \<^term>\<open>1 + fps_X \<noteq> 0\<close> that would otherwise not be solved automatically.
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   171
\<close>
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   172
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   173
ML \<open>
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   174
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   175
(* TODO: Support for division *)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   176
signature POLY_FPS = sig
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   177
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   178
val reify_conv : conv
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   179
val eq_conv : conv
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   180
val eq_simproc : cterm -> thm option
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   181
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   182
end
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   183
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   184
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   185
structure Poly_Fps = struct
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   186
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   187
fun const_binop_conv s conv ct =
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   188
  case Thm.term_of ct of
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   189
    (Const (s', _) $ _ $ _) => 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   190
      if s = s' then 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   191
        Conv.binop_conv conv ct 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   192
      else 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   193
        raise CTERM ("const_binop_conv", [ct])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   194
  | _ => raise CTERM ("const_binop_conv", [ct])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   195
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   196
fun reify_conv ct = 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   197
  let
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   198
    val rewr = Conv.rewrs_conv o map (fn thm => thm RS @{thm eq_reflection})
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   199
    val un = Conv.arg_conv reify_conv
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   200
    val bin = Conv.binop_conv reify_conv
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   201
  in
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   202
    case Thm.term_of ct of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   203
      (Const (\<^const_name>\<open>fps_of_poly\<close>, _) $ _) => ct |> Conv.all_conv
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   204
    | (Const (\<^const_name>\<open>Groups.plus\<close>, _) $ _ $ _) => ct |> (
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   205
        bin then_conv rewr @{thms fps_of_poly_add [symmetric]})
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   206
    | (Const (\<^const_name>\<open>Groups.uminus\<close>, _) $ _) => ct |> (
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   207
        un then_conv rewr @{thms fps_of_poly_uminus [symmetric]})
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   208
    | (Const (\<^const_name>\<open>Groups.minus\<close>, _) $ _ $ _) => ct |> (
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   209
        bin then_conv rewr @{thms fps_of_poly_diff [symmetric]})
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   210
    | (Const (\<^const_name>\<open>Groups.times\<close>, _) $ _ $ _) => ct |> (
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   211
        bin then_conv rewr @{thms fps_of_poly_mult [symmetric]})
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   212
    | (Const (\<^const_name>\<open>Rings.divide\<close>, _) $ _ $ (Const (\<^const_name>\<open>Num.numeral\<close>, _) $ _))
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   213
        => ct |> (Conv.fun_conv (Conv.arg_conv reify_conv)
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   214
             then_conv rewr @{thms fps_of_poly_divide_numeral [symmetric]})
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   215
    | (Const (\<^const_name>\<open>Power.power\<close>, _) $ Const (\<^const_name>\<open>fps_X\<close>,_) $ _) => ct |> (
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   216
        rewr @{thms fps_of_poly_monom' [symmetric]}) 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   217
    | (Const (\<^const_name>\<open>Power.power\<close>, _) $ _ $ _) => ct |> (
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   218
        Conv.fun_conv (Conv.arg_conv reify_conv) 
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   219
        then_conv rewr @{thms fps_of_poly_power [symmetric]})
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   220
    | _ => ct |> (
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   221
        rewr @{thms reify_fps_atom [symmetric]})
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   222
  end
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   223
    
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   224
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   225
fun eq_conv ct =
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   226
  case Thm.term_of ct of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
   227
    (Const (\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ _) => ct |> (
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   228
      Conv.binop_conv reify_conv
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   229
      then_conv Conv.rewr_conv @{thm fps_of_poly_eq_iff[THEN eq_reflection]})
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   230
  | _ => raise CTERM ("poly_fps_eq_conv", [ct])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   231
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   232
val eq_simproc = try eq_conv
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   233
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   234
end
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   235
\<close> 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   236
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   237
simproc_setup poly_fps_eq ("(f :: 'a fps) = g") = \<open>K (K Poly_Fps.eq_simproc)\<close>
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   238
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
   239
lemma fps_of_poly_linear: "fps_of_poly [:a,1 :: 'a :: field:] = fps_X + fps_const a"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   240
  by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   241
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
   242
lemma fps_of_poly_linear': "fps_of_poly [:1,a :: 'a :: field:] = 1 + fps_const a * fps_X"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   243
  by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   244
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   245
lemma fps_of_poly_cutoff [simp]: 
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   246
  "fps_of_poly (poly_cutoff n p) = fps_cutoff n (fps_of_poly p)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   247
  by (simp add: fps_eq_iff coeff_poly_cutoff)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   248
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   249
lemma fps_of_poly_shift [simp]: "fps_of_poly (poly_shift n p) = fps_shift n (fps_of_poly p)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   250
  by (simp add: fps_eq_iff coeff_poly_shift)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   251
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   252
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   253
definition poly_subdegree :: "'a::zero poly \<Rightarrow> nat" where
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   254
  "poly_subdegree p = subdegree (fps_of_poly p)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   255
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   256
lemma coeff_less_poly_subdegree:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   257
  "k < poly_subdegree p \<Longrightarrow> coeff p k = 0"
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   258
  unfolding poly_subdegree_def using nth_less_subdegree_zero[of k "fps_of_poly p"] by simp
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   259
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   260
(* TODO: Move ? *)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   261
definition prefix_length :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat" where
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   262
  "prefix_length P xs = length (takeWhile P xs)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   263
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   264
primrec prefix_length_aux :: "('a \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> nat" where
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   265
  "prefix_length_aux P acc [] = acc"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   266
| "prefix_length_aux P acc (x#xs) = (if P x then prefix_length_aux P (Suc acc) xs else acc)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   267
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   268
lemma prefix_length_aux_correct: "prefix_length_aux P acc xs = prefix_length P xs + acc"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   269
  by (induction xs arbitrary: acc) (simp_all add: prefix_length_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   270
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   271
lemma prefix_length_code [code]: "prefix_length P xs = prefix_length_aux P 0 xs"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   272
  by (simp add: prefix_length_aux_correct)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   273
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   274
lemma prefix_length_le_length: "prefix_length P xs \<le> length xs"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   275
  by (induction xs) (simp_all add: prefix_length_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   276
  
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   277
lemma prefix_length_less_length: "(\<exists>x\<in>set xs. \<not>P x) \<Longrightarrow> prefix_length P xs < length xs"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   278
  by (induction xs) (simp_all add: prefix_length_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   279
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   280
lemma nth_prefix_length:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   281
  "(\<exists>x\<in>set xs. \<not>P x) \<Longrightarrow> \<not>P (xs ! prefix_length P xs)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   282
  by (induction xs) (simp_all add: prefix_length_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   283
  
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   284
lemma nth_less_prefix_length:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   285
  "n < prefix_length P xs \<Longrightarrow> P (xs ! n)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   286
  by (induction xs arbitrary: n) 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   287
     (auto simp: prefix_length_def nth_Cons split: if_splits nat.splits)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   288
(* END TODO *)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   289
  
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66480
diff changeset
   290
lemma poly_subdegree_code [code]: "poly_subdegree p = prefix_length ((=) 0) (coeffs p)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   291
proof (cases "p = 0")
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   292
  case False
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   293
  note [simp] = this
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66480
diff changeset
   294
  define n where "n = prefix_length ((=) 0) (coeffs p)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   295
  from False have "\<exists>k. coeff p k \<noteq> 0" by (auto simp: poly_eq_iff)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   296
  hence ex: "\<exists>x\<in>set (coeffs p). x \<noteq> 0" by (auto simp: coeffs_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   297
  hence n_less: "n < length (coeffs p)" and nonzero: "coeffs p ! n \<noteq> 0" 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   298
    unfolding n_def by (auto intro!: prefix_length_less_length nth_prefix_length)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   299
  show ?thesis unfolding poly_subdegree_def
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   300
  proof (intro subdegreeI)
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   301
    from n_less have "fps_of_poly p $ n = coeffs p ! n"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   302
      by (subst coeffs_nth) (simp_all add: degree_eq_length_coeffs)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66480
diff changeset
   303
    with nonzero show "fps_of_poly p $ prefix_length ((=) 0) (coeffs p) \<noteq> 0"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   304
      unfolding n_def by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   305
  next
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66480
diff changeset
   306
    fix k assume A: "k < prefix_length ((=) 0) (coeffs p)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   307
    also have "\<dots> \<le> length (coeffs p)" by (rule prefix_length_le_length)
63319
bc8793d7bd21 fps_from_poly → fps_of_poly
eberlm
parents: 63317
diff changeset
   308
    finally show "fps_of_poly p $ k = 0"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   309
      using nth_less_prefix_length[OF A]
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   310
      by (simp add: coeffs_nth degree_eq_length_coeffs)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   311
  qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   312
qed (simp_all add: poly_subdegree_def prefix_length_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   313
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents:
diff changeset
   314
end
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
   315
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 65486
diff changeset
   316
end