Skip to content

Commit

Permalink
feat(algebra/polynomial/big_operators): add a lemma, reduce assumptio…
Browse files Browse the repository at this point in the history
…ns, golf (#13264)
  • Loading branch information
negiizhao committed Apr 25, 2022
1 parent 0d16bb4 commit 54d1ddd
Show file tree
Hide file tree
Showing 2 changed files with 24 additions and 10 deletions.
32 changes: 23 additions & 9 deletions src/algebra/polynomial/big_operators.lean
Expand Up @@ -277,7 +277,23 @@ by simpa using multiset_prod_X_sub_C_coeff_card_pred (s.1.map f) (by simpa using
end comm_ring

section no_zero_divisors
variables [comm_ring R] [no_zero_divisors R] (f : ι → R[X]) (t : multiset R[X])

section semiring
variables [semiring R] [no_zero_divisors R]

/--
The degree of a product of polynomials is equal to
the sum of the degrees, where the degree of the zero polynomial is ⊥.
`[nontrivial R]` is needed, otherwise for `l = []` we have `⊥` in the LHS and `0` in the RHS.
-/
lemma degree_list_prod [nontrivial R] (l : list R[X]) :
l.prod.degree = (l.map degree).sum :=
map_list_prod (@degree_monoid_hom R _ _ _) l

end semiring

section comm_semiring
variables [comm_semiring R] [no_zero_divisors R] (f : ι → R[X]) (t : multiset R[X])

/--
The degree of a product of polynomials is equal to
Expand All @@ -295,9 +311,8 @@ begin
intros x hx, simp [h x hx]
end

lemma nat_degree_multiset_prod (s : multiset R[X])
(h : (0 : R[X]) ∉ s) :
nat_degree s.prod = (s.map nat_degree).sum :=
lemma nat_degree_multiset_prod (h : (0 : R[X]) ∉ t) :
nat_degree t.prod = (t.map nat_degree).sum :=
begin
nontriviality R,
rw nat_degree_multiset_prod',
Expand All @@ -312,17 +327,14 @@ the sum of the degrees, where the degree of the zero polynomial is ⊥.
-/
lemma degree_multiset_prod [nontrivial R] :
t.prod.degree = (t.map (λ f, degree f)).sum :=
begin
refine multiset.induction_on t _ (λ a t ht, _), { simp },
{ rw [multiset.prod_cons, degree_mul, ht, map_cons, multiset.sum_cons] }
end
map_multiset_prod (@degree_monoid_hom R _ _ _) _

/--
The degree of a product of polynomials is equal to
the sum of the degrees, where the degree of the zero polynomial is ⊥.
-/
lemma degree_prod [nontrivial R] : (∏ i in s, f i).degree = ∑ i in s, (f i).degree :=
by simpa using degree_multiset_prod (s.1.map f)
map_prod (@degree_monoid_hom R _ _ _) _ _

/--
The leading coefficient of a product of polynomials is equal to
Expand All @@ -346,5 +358,7 @@ lemma leading_coeff_prod :
(∏ i in s, f i).leading_coeff = ∏ i in s, (f i).leading_coeff :=
by simpa using leading_coeff_multiset_prod (s.1.map f)

end comm_semiring

end no_zero_divisors
end polynomial
2 changes: 1 addition & 1 deletion src/data/polynomial/degree/definitions.lean
Expand Up @@ -1189,7 +1189,7 @@ def degree_monoid_hom [nontrivial R] : R[X] →* multiplicative (with_bot ℕ) :

@[simp] lemma degree_pow [nontrivial R] (p : R[X]) (n : ℕ) :
degree (p ^ n) = n • (degree p) :=
map_pow (degree_monoid_hom : R[X] →* _) _ _
map_pow (@degree_monoid_hom R _ _ _) _ _

@[simp] lemma leading_coeff_mul (p q : R[X]) : leading_coeff (p * q) =
leading_coeff p * leading_coeff q :=
Expand Down

0 comments on commit 54d1ddd

Please sign in to comment.