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

Commit 473da0e

Browse files
committed
feat(data/polynomial/{ basic + div + monic + degree/definitions }): lemmas about monic and forall_eq (#15817)
This PR reorganizes a couple of lemmas about monic. The main breaking change is the removal of `subsingleton_of_monic_zero'` in favour of `monic_zero_iff_subsingleton`. I left in `monic_zero_iff_subsingleton'` an iff version of `subsingleton_of_monic_zero'`, but I think that this can probably be removed, thanks to `monic_zero_iff_subsingleton` and `forall_eq_iff_forall_eq`. Also, if anyone wants to rename some/all subsingletons to forall_eq, I am happy to do the change!
1 parent 7e33e4f commit 473da0e

File tree

4 files changed

+25
-11
lines changed

4 files changed

+25
-11
lines changed

src/data/polynomial/basic.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -480,6 +480,15 @@ lemma monomial_eq_C_mul_X : ∀{n}, monomial n a = C a * X^n
480480
calc C a = 0 ↔ C a = C 0 : by rw C_0
481481
... ↔ a = 0 : C_inj
482482

483+
lemma subsingleton_iff_subsingleton :
484+
subsingleton R[X] ↔ subsingleton R :=
485+
⟨λ h, subsingleton_iff.mpr (λ a b, C_inj.mp (subsingleton_iff.mp h _ _)),
486+
by { introI, apply_instance } ⟩
487+
488+
lemma forall_eq_iff_forall_eq :
489+
(∀ f g : R[X], f = g) ↔ (∀ a b : R, a = b) :=
490+
by simpa only [← subsingleton_iff] using subsingleton_iff_subsingleton
491+
483492
theorem ext_iff {p q : R[X]} : p = q ↔ ∀ n, coeff p n = coeff q n :=
484493
by { rcases p, rcases q, simp [coeff, finsupp.ext_iff] }
485494

src/data/polynomial/degree/definitions.lean

Lines changed: 0 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -839,12 +839,6 @@ begin
839839
exact coeff_mul_degree_add_degree _ _ } }
840840
end
841841

842-
lemma subsingleton_of_monic_zero (h : monic (0 : R[X])) :
843-
(∀ p q : R[X], p = q) ∧ (∀ a b : R, a = b) :=
844-
by rw [monic.def, leading_coeff_zero] at h;
845-
exact ⟨λ p q, by rw [← mul_one p, ← mul_one q, ← C_1, ← h, C_0, mul_zero, mul_zero],
846-
λ a b, by rw [← mul_one a, ← mul_one b, ← h, mul_zero, mul_zero]⟩
847-
848842
lemma zero_le_degree_iff {p : R[X]} : 0 ≤ degree p ↔ p ≠ 0 :=
849843
by rw [ne.def, ← degree_eq_bot];
850844
cases degree p; exact dec_trivial

src/data/polynomial/div.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -150,12 +150,12 @@ begin
150150
end
151151

152152
@[simp] lemma mod_by_monic_zero (p : R[X]) : p %ₘ 0 = p :=
153-
if h : monic (0 : R[X]) then (subsingleton_of_monic_zero h).1 _ _ else
154-
by unfold mod_by_monic div_mod_by_monic_aux; rw dif_neg h
153+
if h : monic (0 : R[X]) then by { haveI := monic_zero_iff_subsingleton.mp h, simp }
154+
else by unfold mod_by_monic div_mod_by_monic_aux; rw dif_neg h
155155

156156
@[simp] lemma div_by_monic_zero (p : R[X]) : p /ₘ 0 = 0 :=
157-
if h : monic (0 : R[X]) then (subsingleton_of_monic_zero h).1 _ _ else
158-
by unfold div_by_monic div_mod_by_monic_aux; rw dif_neg h
157+
if h : monic (0 : R[X]) then by { haveI := monic_zero_iff_subsingleton.mp h, simp }
158+
else by unfold div_by_monic div_mod_by_monic_aux; rw dif_neg h
159159

160160
lemma div_by_monic_eq_of_not_monic (p : R[X]) (hq : ¬monic q) : p /ₘ q = 0 := dif_neg hq
161161

src/data/polynomial/monic.lean

Lines changed: 12 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,17 @@ variables {R : Type u} {S : Type v} {a b : R} {m n : ℕ} {ι : Type y}
2626
section semiring
2727
variables [semiring R] {p q r : R[X]}
2828

29+
lemma monic_zero_iff_subsingleton : monic (0 : R[X]) ↔ subsingleton R :=
30+
subsingleton_iff_zero_eq_one
31+
32+
lemma not_monic_zero_iff : ¬ monic (0 : R[X]) ↔ (0 : R) ≠ 1 :=
33+
(monic_zero_iff_subsingleton.trans subsingleton_iff_zero_eq_one.symm).not
34+
35+
lemma monic_zero_iff_subsingleton' :
36+
monic (0 : R[X]) ↔ (∀ f g : R[X], f = g) ∧ (∀ a b : R, a = b) :=
37+
polynomial.monic_zero_iff_subsingleton.trans ⟨by { introI, simp },
38+
λ h, subsingleton_iff.mpr h.2
39+
2940
lemma monic.as_sum (hp : p.monic) :
3041
p = X^(p.nat_degree) + (∑ i in range p.nat_degree, C (p.coeff i) * X^i) :=
3142
begin
@@ -392,7 +403,7 @@ section nonzero_semiring
392403
variables [semiring R] [nontrivial R] {p q : R[X]}
393404

394405
@[simp] lemma not_monic_zero : ¬monic (0 : R[X]) :=
395-
by simpa only [monic, leading_coeff_zero] using (zero_ne_one : (0 : R) ≠ 1)
406+
not_monic_zero_iff.mp zero_ne_one
396407

397408
end nonzero_semiring
398409

0 commit comments

Comments
 (0)