Skip to content

Commit 12f3163

Browse files
committed
chore(Algebra/Polynomial/Degree): split Definitions.lean (#19038)
This file was very big and contained many more things than just the definitions of `Polynomial.degree` et al. This PR splits off the following files: * `Degree/Operations.lean`: lemmas to compute the degree for various ring operations (this is probably the file to import if you're missing something after merging this PR) * `Degree/Domain.lean`: `Polynomial.instIsDomain` instance and its consequences * `Degree/Monomial.lean`: lemmas to compute the degree for monomials * `Degree/SmallDegree.lean`: lemmas where the degree is a specific small number such as 2 or 3 * `Degree/Support.lean`: lemmas about writing polynomials as sum over the support * `Degree/Units.lean`: lemmas about the degree of units
1 parent 15661bf commit 12f3163

File tree

15 files changed

+1289
-1046
lines changed

15 files changed

+1289
-1046
lines changed

Mathlib.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -761,8 +761,14 @@ import Mathlib.Algebra.Polynomial.Cardinal
761761
import Mathlib.Algebra.Polynomial.Coeff
762762
import Mathlib.Algebra.Polynomial.Degree.CardPowDegree
763763
import Mathlib.Algebra.Polynomial.Degree.Definitions
764+
import Mathlib.Algebra.Polynomial.Degree.Domain
764765
import Mathlib.Algebra.Polynomial.Degree.Lemmas
766+
import Mathlib.Algebra.Polynomial.Degree.Monomial
767+
import Mathlib.Algebra.Polynomial.Degree.Operations
768+
import Mathlib.Algebra.Polynomial.Degree.SmallDegree
769+
import Mathlib.Algebra.Polynomial.Degree.Support
765770
import Mathlib.Algebra.Polynomial.Degree.TrailingDegree
771+
import Mathlib.Algebra.Polynomial.Degree.Units
766772
import Mathlib.Algebra.Polynomial.DenomsClearable
767773
import Mathlib.Algebra.Polynomial.Derivation
768774
import Mathlib.Algebra.Polynomial.Derivative

Mathlib/Algebra/Polynomial/CancelLeads.lean

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,6 @@ Copyright (c) 2020 Aaron Anderson. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Aaron Anderson
55
-/
6-
import Mathlib.Algebra.Polynomial.Degree.Definitions
76
import Mathlib.Algebra.Polynomial.Degree.Lemmas
87
import Mathlib.Tactic.ComputeDegree
98

Mathlib/Algebra/Polynomial/Degree/Definitions.lean

Lines changed: 12 additions & 1040 deletions
Large diffs are not rendered by default.
Lines changed: 95 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,95 @@
1+
/-
2+
Copyright (c) 2018 Chris Hughes. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker
5+
-/
6+
import Mathlib.Algebra.Polynomial.Degree.Operations
7+
8+
/-!
9+
# Univariate polynomials form a domain
10+
11+
## Main results
12+
13+
* `Polynomial.instNoZeroDivisors`: `R[X]` has no zero divisors if `R` does not
14+
* `Polynomial.instDomain`: `R[X]` is a domain if `R` is
15+
-/
16+
17+
noncomputable section
18+
19+
open Finsupp Finset
20+
21+
open Polynomial
22+
23+
namespace Polynomial
24+
25+
universe u v
26+
27+
variable {R : Type u} {S : Type v} {a b c d : R} {n m : ℕ}
28+
29+
section Semiring
30+
31+
variable [Semiring R] [NoZeroDivisors R] {p q : R[X]}
32+
33+
instance : NoZeroDivisors R[X] where
34+
eq_zero_or_eq_zero_of_mul_eq_zero h := by
35+
rw [← leadingCoeff_eq_zero, ← leadingCoeff_eq_zero]
36+
refine eq_zero_or_eq_zero_of_mul_eq_zero ?_
37+
rw [← leadingCoeff_zero, ← leadingCoeff_mul, h]
38+
39+
lemma natDegree_mul (hp : p ≠ 0) (hq : q ≠ 0) : (p*q).natDegree = p.natDegree + q.natDegree := by
40+
rw [← Nat.cast_inj (R := WithBot ℕ), ← degree_eq_natDegree (mul_ne_zero hp hq),
41+
Nat.cast_add, ← degree_eq_natDegree hp, ← degree_eq_natDegree hq, degree_mul]
42+
43+
@[simp]
44+
lemma natDegree_pow (p : R[X]) (n : ℕ) : natDegree (p ^ n) = n * natDegree p := by
45+
classical
46+
obtain rfl | hp := eq_or_ne p 0
47+
· obtain rfl | hn := eq_or_ne n 0 <;> simp [*]
48+
exact natDegree_pow' <| by
49+
rw [← leadingCoeff_pow, Ne, leadingCoeff_eq_zero]; exact pow_ne_zero _ hp
50+
51+
lemma natDegree_le_of_dvd (h1 : p ∣ q) (h2 : q ≠ 0) : p.natDegree ≤ q.natDegree := by
52+
obtain ⟨q, rfl⟩ := h1
53+
rw [mul_ne_zero_iff] at h2
54+
rw [natDegree_mul h2.1 h2.2]; exact Nat.le_add_right _ _
55+
56+
lemma degree_le_of_dvd (h1 : p ∣ q) (h2 : q ≠ 0) : degree p ≤ degree q := by
57+
rcases h1 with ⟨q, rfl⟩; rw [mul_ne_zero_iff] at h2
58+
exact degree_le_mul_left p h2.2
59+
60+
lemma eq_zero_of_dvd_of_degree_lt (h₁ : p ∣ q) (h₂ : degree q < degree p) : q = 0 := by
61+
by_contra hc
62+
exact (lt_iff_not_ge _ _).mp h₂ (degree_le_of_dvd h₁ hc)
63+
64+
lemma eq_zero_of_dvd_of_natDegree_lt (h₁ : p ∣ q) (h₂ : natDegree q < natDegree p) :
65+
q = 0 := by
66+
by_contra hc
67+
exact (lt_iff_not_ge _ _).mp h₂ (natDegree_le_of_dvd h₁ hc)
68+
69+
lemma not_dvd_of_degree_lt (h0 : q ≠ 0) (hl : q.degree < p.degree) : ¬p ∣ q := by
70+
by_contra hcontra
71+
exact h0 (eq_zero_of_dvd_of_degree_lt hcontra hl)
72+
73+
lemma not_dvd_of_natDegree_lt (h0 : q ≠ 0) (hl : q.natDegree < p.natDegree) :
74+
¬p ∣ q := by
75+
by_contra hcontra
76+
exact h0 (eq_zero_of_dvd_of_natDegree_lt hcontra hl)
77+
78+
/-- This lemma is useful for working with the `intDegree` of a rational function. -/
79+
lemma natDegree_sub_eq_of_prod_eq {p₁ p₂ q₁ q₂ : R[X]} (hp₁ : p₁ ≠ 0) (hq₁ : q₁ ≠ 0)
80+
(hp₂ : p₂ ≠ 0) (hq₂ : q₂ ≠ 0) (h_eq : p₁ * q₂ = p₂ * q₁) :
81+
(p₁.natDegree : ℤ) - q₁.natDegree = (p₂.natDegree : ℤ) - q₂.natDegree := by
82+
rw [sub_eq_sub_iff_add_eq_add]
83+
norm_cast
84+
rw [← natDegree_mul hp₁ hq₂, ← natDegree_mul hp₂ hq₁, h_eq]
85+
86+
end Semiring
87+
88+
section Ring
89+
90+
variable [Ring R] [Nontrivial R] [IsDomain R] {p q : R[X]}
91+
92+
instance : IsDomain R[X] := NoZeroDivisors.to_isDomain _
93+
94+
end Ring
95+
end Polynomial
Lines changed: 46 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,46 @@
1+
/-
2+
Copyright (c) 2018 Chris Hughes. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker
5+
-/
6+
import Mathlib.Algebra.Polynomial.Degree.Definitions
7+
import Mathlib.Algebra.Polynomial.Monomial
8+
import Mathlib.Data.Nat.SuccPred
9+
10+
/-!
11+
# Degree of univariate monomials
12+
-/
13+
14+
noncomputable section
15+
16+
open Finsupp Finset Polynomial
17+
18+
namespace Polynomial
19+
20+
universe u v
21+
22+
variable {R : Type u} {S : Type v} {a b c d : R} {n m : ℕ}
23+
24+
section Semiring
25+
26+
variable [Semiring R] {p q : R[X]} {ι : Type*}
27+
28+
lemma natDegree_le_pred (hf : p.natDegree ≤ n) (hn : p.coeff n = 0) : p.natDegree ≤ n - 1 := by
29+
obtain _ | n := n
30+
· exact hf
31+
· refine (Nat.le_succ_iff_eq_or_le.1 hf).resolve_left fun h ↦ ?_
32+
rw [← Nat.succ_eq_add_one, ← h, coeff_natDegree, leadingCoeff_eq_zero] at hn
33+
aesop
34+
35+
theorem monomial_natDegree_leadingCoeff_eq_self (h : #p.support ≤ 1) :
36+
monomial p.natDegree p.leadingCoeff = p := by
37+
classical
38+
rcases card_support_le_one_iff_monomial.1 h with ⟨n, a, rfl⟩
39+
by_cases ha : a = 0 <;> simp [ha]
40+
41+
theorem C_mul_X_pow_eq_self (h : #p.support ≤ 1) : C p.leadingCoeff * X ^ p.natDegree = p := by
42+
rw [C_mul_X_pow_eq_monomial, monomial_natDegree_leadingCoeff_eq_self h]
43+
44+
end Semiring
45+
46+
end Polynomial

0 commit comments

Comments
 (0)