diff --git a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean index b2745e2b69ffe..21e4f3cdbc54f 100644 --- a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean +++ b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean @@ -42,6 +42,10 @@ def nonZeroDivisorsLeft : Submonoid M₀ where x ∈ nonZeroDivisorsLeft M₀ ↔ ∀ y, y * x = 0 → y = 0 := Iff.rfl +lemma nmem_nonZeroDivisorsLeft_iff {r : M₀} : + r ∉ nonZeroDivisorsLeft M₀ ↔ {s | s * r = 0 ∧ s ≠ 0}.Nonempty := by + simpa [mem_nonZeroDivisorsLeft_iff] using Set.nonempty_def.symm + /-- The collection of elements of a `MonoidWithZero` that are not right zero divisors form a `Submonoid`. -/ def nonZeroDivisorsRight : Submonoid M₀ where @@ -53,6 +57,10 @@ def nonZeroDivisorsRight : Submonoid M₀ where x ∈ nonZeroDivisorsRight M₀ ↔ ∀ y, x * y = 0 → y = 0 := Iff.rfl +lemma nmem_nonZeroDivisorsRight_iff {r : M₀} : + r ∉ nonZeroDivisorsRight M₀ ↔ {s | r * s = 0 ∧ s ≠ 0}.Nonempty := by + simpa [mem_nonZeroDivisorsRight_iff] using Set.nonempty_def.symm + lemma nonZeroDivisorsLeft_eq_right (M₀ : Type*) [CommMonoidWithZero M₀] : nonZeroDivisorsLeft M₀ = nonZeroDivisorsRight M₀ := by ext x; simp [mul_comm x] @@ -107,6 +115,9 @@ variable {M M' M₁ R R' F : Type*} [MonoidWithZero M] [MonoidWithZero M'] [Comm theorem mem_nonZeroDivisors_iff {r : M} : r ∈ M⁰ ↔ ∀ x, x * r = 0 → x = 0 := Iff.rfl #align mem_non_zero_divisors_iff mem_nonZeroDivisors_iff +lemma nmem_nonZeroDivisors_iff {r : M} : r ∉ M⁰ ↔ {s | s * r = 0 ∧ s ≠ 0}.Nonempty := by + simpa [mem_nonZeroDivisors_iff] using Set.nonempty_def.symm + theorem mul_right_mem_nonZeroDivisors_eq_zero_iff {x r : M} (hr : r ∈ M⁰) : x * r = 0 ↔ x = 0 := ⟨hr _, by simp (config := { contextual := true })⟩ #align mul_right_mem_non_zero_divisors_eq_zero_iff mul_right_mem_nonZeroDivisors_eq_zero_iff diff --git a/Mathlib/Algebra/Polynomial/RingDivision.lean b/Mathlib/Algebra/Polynomial/RingDivision.lean index 52e46f0bad238..b14952410234f 100644 --- a/Mathlib/Algebra/Polynomial/RingDivision.lean +++ b/Mathlib/Algebra/Polynomial/RingDivision.lean @@ -364,6 +364,51 @@ theorem natDegree_pos_of_monic_of_not_isUnit {a : R[X]} (hu : ¬ IsUnit a) (ha : 0 < natDegree a := natDegree_pos_iff_degree_pos.mpr <| degree_pos_of_monic_of_not_isUnit hu ha +theorem eq_zero_of_mul_eq_zero_of_smul (P : R[X]) (h : ∀ r : R, r • P = 0 → r = 0) : + ∀ (Q : R[X]), P * Q = 0 → Q = 0 := by + intro Q hQ + suffices ∀ i, P.coeff i • Q = 0 by + rw [← leadingCoeff_eq_zero] + apply h + simpa [ext_iff, mul_comm Q.leadingCoeff] using fun i ↦ congr_arg (·.coeff Q.natDegree) (this i) + apply Nat.strong_decreasing_induction + · use P.natDegree + intro i hi + rw [coeff_eq_zero_of_natDegree_lt hi, zero_smul] + intro l IH + obtain _|hl := (natDegree_smul_le (P.coeff l) Q).lt_or_eq + · apply eq_zero_of_mul_eq_zero_of_smul _ h (P.coeff l • Q) + rw [smul_eq_C_mul, mul_left_comm, hQ, mul_zero] + suffices P.coeff l * Q.leadingCoeff = 0 by + rwa [← leadingCoeff_eq_zero, ← coeff_natDegree, coeff_smul, hl, coeff_natDegree, smul_eq_mul] + let m := Q.natDegree + suffices (P * Q).coeff (l + m) = P.coeff l * Q.leadingCoeff by rw [← this, hQ, coeff_zero] + rw [coeff_mul] + apply Finset.sum_eq_single (l, m) _ (by simp) + simp only [Finset.mem_antidiagonal, ne_eq, Prod.forall, Prod.mk.injEq, not_and] + intro i j hij H + obtain hi|rfl|hi := lt_trichotomy i l + · have hj : m < j := by omega + rw [coeff_eq_zero_of_natDegree_lt hj, mul_zero] + · omega + · rw [← coeff_C_mul, ← smul_eq_C_mul, IH _ hi, coeff_zero] +termination_by Q => Q.natDegree + +open nonZeroDivisors in +/-- *McCoy theorem*: a polynomial `P : R[X]` is a zerodivisor if and only if there is `a : R` +such that `a ≠ 0` and `a • P = 0`. -/ +theorem nmem_nonZeroDivisors_iff {P : R[X]} : P ∉ R[X]⁰ ↔ ∃ a : R, a ≠ 0 ∧ a • P = 0 := by + refine ⟨fun hP ↦ ?_, fun ⟨a, ha, h⟩ h1 ↦ ha <| C_eq_zero.1 <| (h1 _) <| smul_eq_C_mul a ▸ h⟩ + by_contra! h + obtain ⟨Q, hQ⟩ := _root_.nmem_nonZeroDivisors_iff.1 hP + refine hQ.2 (eq_zero_of_mul_eq_zero_of_smul P (fun a ha ↦ ?_) Q (mul_comm P _ ▸ hQ.1)) + contrapose! ha + exact h a ha + +open nonZeroDivisors in +protected lemma mem_nonZeroDivisors_iff {P : R[X]} : P ∈ R[X]⁰ ↔ ∀ a : R, a • P = 0 → a = 0 := by + simpa [not_imp_not] using (nmem_nonZeroDivisors_iff (P := P)).not + end CommSemiring section CommRing diff --git a/Mathlib/Data/Nat/Interval.lean b/Mathlib/Data/Nat/Interval.lean index f362f15ab4d2a..59911d126eab2 100644 --- a/Mathlib/Data/Nat/Interval.lean +++ b/Mathlib/Data/Nat/Interval.lean @@ -361,6 +361,19 @@ theorem Nat.decreasing_induction_of_not_bddAbove (hP : ¬BddAbove { x | P x }) ( decreasingInduction h hl.le hm #align nat.decreasing_induction_of_not_bdd_above Nat.decreasing_induction_of_not_bddAbove +@[elab_as_elim] +lemma Nat.strong_decreasing_induction (base : ∃ n, ∀ m > n, P m) (step : ∀ n, (∀ m > n, P m) → P n) + (n : ℕ) : P n := by + apply Nat.decreasing_induction_of_not_bddAbove (P := fun n ↦ ∀ m ≥ n, P m) _ _ n n le_rfl + · intro n ih m hm + rcases hm.eq_or_lt with rfl | hm + · exact step n ih + · exact ih m hm + · rintro ⟨b, hb⟩ + rcases base with ⟨n, hn⟩ + specialize @hb (n + b + 1) (fun m hm ↦ hn _ _) + all_goals omega + theorem Nat.decreasing_induction_of_infinite (hP : { x | P x }.Infinite) (n : ℕ) : P n := Nat.decreasing_induction_of_not_bddAbove h (mt BddAbove.finite hP) n #align nat.decreasing_induction_of_infinite Nat.decreasing_induction_of_infinite