Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit ad25cac

Browse files
committed
refactor(data/polynomial/derivative): change polynomial.derivative to be a linear_map (#5198)
Refactors polynomial.derivative to be a linear_map by default
1 parent 240f6eb commit ad25cac

File tree

2 files changed

+42
-47
lines changed

2 files changed

+42
-47
lines changed

src/data/polynomial/derivative.lean

Lines changed: 42 additions & 43 deletions
Original file line numberDiff line numberDiff line change
@@ -6,8 +6,10 @@ Authors: Chris Hughes, Johannes Hölzl, Scott Morrison, Jens Wagemaker
66
import data.polynomial.eval
77

88
/-!
9-
# Theory of univariate polynomials
9+
# The derivative map on polynomials
1010
11+
## Main definitions
12+
* `polynomial.derivative`: The formal derivative of polynomials, expressed as a linear map.
1113
1214
-/
1315

@@ -27,12 +29,22 @@ section semiring
2729
variables [semiring R]
2830

2931
/-- `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
3143

3244
lemma coeff_derivative (p : polynomial R) (n : ℕ) :
3345
coeff (derivative p) n = coeff p (n + 1) * (n + 1) :=
3446
begin
35-
rw [derivative],
47+
rw [derivative_apply],
3648
simp only [coeff_X_pow, coeff_sum, coeff_C_mul],
3749
rw [finsupp.sum, finset.sum_eq_single (n + 1)],
3850
simp only [nat.add_succ_sub_one, add_zero, mul_one, if_true, eq_self_iff_true], norm_cast,
@@ -43,17 +55,17 @@ begin
4355
{ intros _ H, rw [nat.succ_sub_one b, if_neg (mt (congr_arg nat.succ) H.symm), mul_zero] } }
4456
end
4557

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
4860

4961
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 _ _))
5163

5264
lemma derivative_C_mul_X_pow (a : R) (n : ℕ) : derivative (C a * X ^ n) = C (a * n) * X^(n - 1) :=
5365
by rw [C_mul_X_pow_eq_monomial, C_mul_X_pow_eq_monomial, derivative_monomial]
5466

5567
@[simp] lemma derivative_C {a : R} : derivative (C a) = 0 :=
56-
(derivative_monomial a 0).trans $ by simp
68+
by simp [derivative_apply]
5769

5870
@[simp] lemma derivative_X : derivative (X : polynomial R) = 1 :=
5971
(derivative_monomial _ _).trans $ by simp
@@ -63,38 +75,32 @@ derivative_C
6375

6476
@[simp] lemma derivative_add {f g : polynomial R} :
6577
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
7479

7580
@[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
7883

7984
@[simp] lemma derivative_sub {R : Type*} [ring R] (f g : polynomial R) :
8085
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
8587

8688
@[simp] lemma derivative_sum {s : finset ι} {f : ι → polynomial R} :
8789
derivative (∑ b in s, f b) = ∑ b in s, derivative (f b) :=
88-
(derivative_hom R).map_sum f s
90+
derivative.map_sum
8991

9092
@[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 _ _
9294

9395
end semiring
9496

9597
section comm_semiring
9698
variables [comm_semiring R]
9799

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+
98104
@[simp] lemma derivative_mul {f g : polynomial R} :
99105
derivative (f * g) = derivative f * g + f * derivative g :=
100106
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 : ℕ)
118124
conv { to_rhs, congr,
119125
{ rw [← sum_C_mul_X_eq g] },
120126
{ 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]
123128
end
124129

125130
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
171176
theorem degree_derivative_le {p : polynomial R} : p.derivative.degree ≤ p.degree :=
172177
if H : p = 0 then le_of_eq $ by rw [H, derivative_zero] else le_of_lt $ degree_derivative_lt H
173178

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-
180179
end comm_semiring
181180

182181
section domain
@@ -190,23 +189,23 @@ by { rw [nat.cast_eq_zero], simp only [nat.succ_ne_zero, or_false] }
190189

191190
@[simp] lemma degree_derivative_eq [char_zero R] (p : polynomial R) (hp : 0 < nat_degree p) :
192191
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 _,
203202
rw [mem_support_derivative, nat.sub_add_cancel, mem_support_iff],
204203
{ show ¬ leading_coeff p = 0,
205204
rw [leading_coeff_eq_zero],
206205
assume h, rw [h, nat_degree_zero] at hp,
207206
exact lt_irrefl 0 (lt_of_le_of_lt (zero_le _) hp), },
208-
exact hp
209-
end
207+
exact hp }
208+
end
210209

211210
theorem nat_degree_eq_zero_of_derivative_eq_zero [char_zero R] {f : polynomial R} (h : f.derivative = 0) :
212211
f.nat_degree = 0 :=

src/data/polynomial/identities.lean

Lines changed: 0 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -63,10 +63,6 @@ private lemma poly_binom_aux3 (f : polynomial R) (x y : R) : f.eval (x + y) =
6363
f.sum (λ e a, (a *(poly_binom_aux1 x y e a).val)*y^2) :=
6464
by rw poly_binom_aux2; simp [left_distrib, finsupp.sum_add, mul_assoc]
6565

66-
lemma derivative_eval (p : polynomial R) (x : R) :
67-
p.derivative.eval x = p.sum (λ n a, (a * n)*x^(n-1)) :=
68-
by simp only [derivative, eval_sum, eval_pow, eval_C, eval_X, eval_nat_cast, eval_mul]
69-
7066
def binom_expansion (f : polynomial R) (x y : R) :
7167
{k : R // f.eval (x + y) = f.eval x + (f.derivative.eval x) * y + k * y^2} :=
7268
begin

0 commit comments

Comments
 (0)