Skip to content

Commit

Permalink
feat(data/polynomial): %ₘ as a linear map (#9532)
Browse files Browse the repository at this point in the history
This PR bundles `(%ₘ) = polynomial.mod_by_monic` into a linear map. In a follow-up PR, I'll show this map descends to a linear map `adjoin_root q → polynomial R`, thereby (computably) assigning coefficients to each element in `adjoin_root q`.
  • Loading branch information
Vierkantor committed Oct 8, 2021
1 parent 99c3e22 commit 745cbfc
Show file tree
Hide file tree
Showing 2 changed files with 48 additions and 0 deletions.
34 changes: 34 additions & 0 deletions src/data/polynomial/ring_division.lean
Expand Up @@ -43,6 +43,40 @@ lemma aeval_mod_by_monic_eq_self_of_root [algebra R S]
aeval x (p %ₘ q) = aeval x p :=
eval₂_mod_by_monic_eq_self_of_root hq hx

lemma mod_by_monic_eq_of_dvd_sub [nontrivial R] (hq : q.monic) {p₁ p₂ : polynomial R}
(h : q ∣ (p₁ - p₂)) :
p₁ %ₘ q = p₂ %ₘ q :=
begin
obtain ⟨f, sub_eq⟩ := h,
refine (div_mod_by_monic_unique (p₂ /ₘ q + f) _ hq
⟨_, degree_mod_by_monic_lt _ hq⟩).2,
rw [sub_eq_iff_eq_add.mp sub_eq, mul_add, ← add_assoc, mod_by_monic_add_div _ hq, add_comm]
end

lemma add_mod_by_monic [nontrivial R] (hq : q.monic)
(p₁ p₂ : polynomial R) : (p₁ + p₂) %ₘ q = p₁ %ₘ q + p₂ %ₘ q :=
(div_mod_by_monic_unique (p₁ /ₘ q + p₂ /ₘ q) _ hq
by rw [mul_add, add_left_comm, add_assoc, mod_by_monic_add_div _ hq, ← add_assoc,
add_comm (q * _), mod_by_monic_add_div _ hq],
(degree_add_le _ _).trans_lt (max_lt (degree_mod_by_monic_lt _ hq)
(degree_mod_by_monic_lt _ hq))⟩).2

lemma smul_mod_by_monic [nontrivial R] (hq : q.monic)
(c : R) (p : polynomial R) : (c • p) %ₘ q = c • (p %ₘ q) :=
(div_mod_by_monic_unique (c • (p /ₘ q)) (c • (p %ₘ q)) hq
by rw [mul_smul_comm, ← smul_add, mod_by_monic_add_div p hq],
(degree_smul_le _ _).trans_lt (degree_mod_by_monic_lt _ hq)⟩).2

/--
`polynomial.mod_by_monic_hom (hq : monic (q : polynomial R))` is `_ %ₘ q` as a `R`-linear map.
-/
@[simps]
def mod_by_monic_hom [nontrivial R] (hq : q.monic) :
polynomial R →ₗ[R] polynomial R :=
{ to_fun := λ p, p %ₘ q,
map_add' := add_mod_by_monic hq,
map_smul' := smul_mod_by_monic hq }

end comm_ring

section no_zero_divisors
Expand Down
14 changes: 14 additions & 0 deletions src/ring_theory/polynomial/basic.lean
Expand Up @@ -303,6 +303,20 @@ begin
exact subtype.mem (coeff p n : T)
end

section mod_by_monic

variables {q : polynomial R}

lemma mem_ker_mod_by_monic [nontrivial R] (hq : q.monic) {p : polynomial R} :
p ∈ (mod_by_monic_hom hq).ker ↔ q ∣ p :=
linear_map.mem_ker.trans (dvd_iff_mod_by_monic_eq_zero hq)

@[simp] lemma ker_mod_by_monic_hom [nontrivial R] (hq : q.monic) :
(polynomial.mod_by_monic_hom hq).ker = (ideal.span {q}).restrict_scalars R :=
submodule.ext (λ f, (mem_ker_mod_by_monic hq).trans ideal.mem_span_singleton.symm)

end mod_by_monic

end polynomial

variables {R : Type u} {σ : Type v} {M : Type w} [comm_ring R] [add_comm_group M] [module R M]
Expand Down

0 comments on commit 745cbfc

Please sign in to comment.