@@ -6,8 +6,10 @@ Authors: Chris Hughes, Johannes Hölzl, Scott Morrison, Jens Wagemaker
6
6
import data.polynomial.eval
7
7
8
8
/-!
9
- # Theory of univariate polynomials
9
+ # The derivative map on polynomials
10
10
11
+ ## Main definitions
12
+ * `polynomial.derivative`: The formal derivative of polynomials, expressed as a linear map.
11
13
12
14
-/
13
15
@@ -27,12 +29,22 @@ section semiring
27
29
variables [semiring R]
28
30
29
31
/-- `derivative p` is the formal derivative of the polynomial `p` -/
30
- def derivative (p : polynomial R) : polynomial R := p.sum (λn a, C (a * n) * X^(n - 1 ))
32
+ def derivative : polynomial R →ₗ[R] polynomial R :=
33
+ finsupp.total ℕ (polynomial R) R (λ n, C ↑n * X^(n - 1 ))
34
+
35
+ lemma derivative_apply (p : polynomial R) :
36
+ derivative p = p.sum (λn a, C (a * n) * X^(n - 1 )) :=
37
+ begin
38
+ rw [derivative, total_apply],
39
+ apply congr rfl,
40
+ ext,
41
+ simp [mul_assoc, coeff_C_mul],
42
+ end
31
43
32
44
lemma coeff_derivative (p : polynomial R) (n : ℕ) :
33
45
coeff (derivative p) n = coeff p (n + 1 ) * (n + 1 ) :=
34
46
begin
35
- rw [derivative ],
47
+ rw [derivative_apply ],
36
48
simp only [coeff_X_pow, coeff_sum, coeff_C_mul],
37
49
rw [finsupp.sum, finset.sum_eq_single (n + 1 )],
38
50
simp only [nat.add_succ_sub_one, add_zero, mul_one, if_true, eq_self_iff_true], norm_cast,
@@ -43,17 +55,17 @@ begin
43
55
{ intros _ H, rw [nat.succ_sub_one b, if_neg (mt (congr_arg nat.succ) H.symm), mul_zero] } }
44
56
end
45
57
46
- @[simp] lemma derivative_zero : derivative (0 : polynomial R) = 0 :=
47
- finsupp.sum_zero_index
58
+ lemma derivative_zero : derivative (0 : polynomial R) = 0 :=
59
+ derivative.map_zero
48
60
49
61
lemma derivative_monomial (a : R) (n : ℕ) : derivative (monomial n a) = monomial (n - 1 ) (a * n) :=
50
- (sum_single_index $ by simp).trans (C_mul_X_pow_eq_monomial _ _)
62
+ (derivative_apply _).trans (( sum_single_index $ by simp).trans (C_mul_X_pow_eq_monomial _ _) )
51
63
52
64
lemma derivative_C_mul_X_pow (a : R) (n : ℕ) : derivative (C a * X ^ n) = C (a * n) * X^(n - 1 ) :=
53
65
by rw [C_mul_X_pow_eq_monomial, C_mul_X_pow_eq_monomial, derivative_monomial]
54
66
55
67
@[simp] lemma derivative_C {a : R} : derivative (C a) = 0 :=
56
- (derivative_monomial a 0 ).trans $ by simp
68
+ by simp [derivative_apply]
57
69
58
70
@[simp] lemma derivative_X : derivative (X : polynomial R) = 1 :=
59
71
(derivative_monomial _ _).trans $ by simp
@@ -63,38 +75,32 @@ derivative_C
63
75
64
76
@[simp] lemma derivative_add {f g : polynomial R} :
65
77
derivative (f + g) = derivative f + derivative g :=
66
- by refine finsupp.sum_add_index _ _; intros;
67
- simp only [add_mul, zero_mul, C_0, C_add, C_mul]
68
-
69
- /-- The formal derivative of polynomials, as additive homomorphism. -/
70
- def derivative_hom (R : Type *) [semiring R] : polynomial R →+ polynomial R :=
71
- { to_fun := derivative,
72
- map_zero' := derivative_zero,
73
- map_add' := λ p q, derivative_add }
78
+ derivative.map_add f g
74
79
75
80
@[simp] lemma derivative_neg {R : Type *} [ring R] (f : polynomial R) :
76
- derivative (-f) = -derivative f :=
77
- (derivative_hom R) .map_neg f
81
+ derivative (-f) = - derivative f :=
82
+ linear_map .map_neg derivative f
78
83
79
84
@[simp] lemma derivative_sub {R : Type *} [ring R] (f g : polynomial R) :
80
85
derivative (f - g) = derivative f - derivative g :=
81
- (derivative_hom R).map_sub f g
82
-
83
- instance : is_add_monoid_hom (derivative : polynomial R → polynomial R) :=
84
- (derivative_hom R).is_add_monoid_hom
86
+ linear_map.map_sub derivative f g
85
87
86
88
@[simp] lemma derivative_sum {s : finset ι} {f : ι → polynomial R} :
87
89
derivative (∑ b in s, f b) = ∑ b in s, derivative (f b) :=
88
- (derivative_hom R) .map_sum f s
90
+ derivative .map_sum
89
91
90
92
@[simp] lemma derivative_smul (r : R) (p : polynomial R) : derivative (r • p) = r • derivative p :=
91
- by { ext, simp only [coeff_derivative, mul_assoc, coeff_smul], }
93
+ derivative.map_smul _ _
92
94
93
95
end semiring
94
96
95
97
section comm_semiring
96
98
variables [comm_semiring R]
97
99
100
+ lemma derivative_eval (p : polynomial R) (x : R) :
101
+ p.derivative.eval x = p.sum (λ n a, (a * n)*x^(n-1 )) :=
102
+ by simp only [derivative_apply, eval_sum, eval_pow, eval_C, eval_X, eval_nat_cast, eval_mul]
103
+
98
104
@[simp] lemma derivative_mul {f g : polynomial R} :
99
105
derivative (f * g) = derivative f * g + f * derivative g :=
100
106
calc derivative (f * g) = f.sum (λn a, g.sum (λm b, C ((a * b) * (n + m : ℕ)) * X^((n + m) - 1 ))) :
@@ -118,8 +124,7 @@ calc derivative (f * g) = f.sum (λn a, g.sum (λm b, C ((a * b) * (n + m : ℕ)
118
124
conv { to_rhs, congr,
119
125
{ rw [← sum_C_mul_X_eq g] },
120
126
{ rw [← sum_C_mul_X_eq f] } },
121
- unfold derivative finsupp.sum,
122
- simp only [sum_add_distrib, finset.mul_sum, finset.sum_mul]
127
+ simp only [finsupp.sum, sum_add_distrib, finset.mul_sum, finset.sum_mul, derivative_apply]
123
128
end
124
129
125
130
theorem derivative_pow_succ (p : polynomial R) (n : ℕ) :
@@ -171,12 +176,6 @@ with_bot.some_lt_some.1 $ by { rw [nat_degree, option.get_or_else_of_ne_none $ m
171
176
theorem degree_derivative_le {p : polynomial R} : p.derivative.degree ≤ p.degree :=
172
177
if H : p = 0 then le_of_eq $ by rw [H, derivative_zero] else le_of_lt $ degree_derivative_lt H
173
178
174
- /-- The formal derivative of polynomials, as linear homomorphism. -/
175
- def derivative_lhom (R : Type *) [comm_ring R] : polynomial R →ₗ[R] polynomial R :=
176
- { to_fun := derivative,
177
- map_add' := λ p q, derivative_add,
178
- map_smul' := λ r p, derivative_smul r p }
179
-
180
179
end comm_semiring
181
180
182
181
section domain
@@ -190,23 +189,23 @@ by { rw [nat.cast_eq_zero], simp only [nat.succ_ne_zero, or_false] }
190
189
191
190
@[simp] lemma degree_derivative_eq [char_zero R] (p : polynomial R) (hp : 0 < nat_degree p) :
192
191
degree (derivative p) = (nat_degree p - 1 : ℕ) :=
193
- le_antisymm
194
- (le_trans (degree_sum_le _ _) $ sup_le $ assume n hn ,
195
- have n ≤ nat_degree p, begin
196
- rw [← with_bot.coe_le_coe, ← degree_eq_nat_degree] ,
197
- { refine le_degree_of_ne_zero _, simpa only [mem_support_iff] using hn } ,
198
- { assume h, simpa only [h, support_zero] using hn }
199
- end ,
200
- le_trans (degree_C_mul_X_pow_le _ _) $ with_bot.coe_le_coe.2 $ nat.sub_le_sub_right this _)
201
- begin
202
- refine le_sup _,
192
+ begin
193
+ have h0 : p ≠ 0 ,
194
+ { contrapose! hp,
195
+ simp [hp] } ,
196
+ apply le_antisymm ,
197
+ { rw derivative_apply,
198
+ apply le_trans (degree_sum_le _ _) (sup_le (λ n hn, _)) ,
199
+ apply le_trans (degree_C_mul_X_pow_le _ _) ( with_bot.coe_le_coe.2 ( nat.sub_le_sub_right _ _)),
200
+ apply le_nat_degree_of_mem_supp _ hn },
201
+ { refine le_sup _,
203
202
rw [mem_support_derivative, nat.sub_add_cancel, mem_support_iff],
204
203
{ show ¬ leading_coeff p = 0 ,
205
204
rw [leading_coeff_eq_zero],
206
205
assume h, rw [h, nat_degree_zero] at hp,
207
206
exact lt_irrefl 0 (lt_of_le_of_lt (zero_le _) hp), },
208
- exact hp
209
- end
207
+ exact hp }
208
+ end
210
209
211
210
theorem nat_degree_eq_zero_of_derivative_eq_zero [char_zero R] {f : polynomial R} (h : f.derivative = 0 ) :
212
211
f.nat_degree = 0 :=
0 commit comments