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

Commit 56f021a

Browse files
committed
feat(data/polynomial/{erase_lead + denoms_clearable}): strengthen a lemma (#11257)
This PR is a step in the direction of simplifying #11139 . It strengthens lemma `induction_with_nat_degree_le` by showing that restriction can be strengthened on one of the assumptions. ~~It proves a lemma computing the `nat_degree` under functions on polynomials that include the shift of a variable.~~ EDIT: the lemma was moved to the later PR #11265. It fixes the unique use of lemma `induction_with_nat_degree_le`. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/144837-PR-reviews/topic/.2311139.20taylor.20sum.20and.20nat_degree_taylor)
1 parent b181a12 commit 56f021a

File tree

2 files changed

+21
-12
lines changed

2 files changed

+21
-12
lines changed

src/data/polynomial/denoms_clearable.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -60,10 +60,10 @@ lemma denoms_clearable.add {N : ℕ} {f g : polynomial R} :
6060

6161
lemma denoms_clearable_of_nat_degree_le (N : ℕ) (a : R) (bu : bi * i b = 1) :
6262
∀ (f : polynomial R), f.nat_degree ≤ N → denoms_clearable a b N f i :=
63-
induction_with_nat_degree_le N
63+
induction_with_nat_degree_le _ N
6464
(denoms_clearable_zero N a bu)
6565
(λ N_1 r r0, denoms_clearable_C_mul_X_pow a bu r)
66-
(λ f g fN gN df dg, df.add dg)
66+
(λ f g fg gN df dg, df.add dg)
6767

6868
/-- If `i : R → K` is a ring homomorphism, `f` is a polynomial with coefficients in `R`,
6969
`a, b` are elements of `R`, with `i b` invertible, then there is a `D ∈ R` such that

src/data/polynomial/erase_lead.lean

Lines changed: 19 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -160,10 +160,11 @@ end erase_lead
160160
required to be at least as big as the `nat_degree` of the polynomial. This is useful to prove
161161
results where you want to change each term in a polynomial to something else depending on the
162162
`nat_degree` of the polynomial itself and not on the specific `nat_degree` of each term. -/
163-
lemma induction_with_nat_degree_le {R : Type*} [semiring R] {P : polynomial R → Prop} (N : ℕ)
163+
lemma induction_with_nat_degree_le (P : polynomial R → Prop) (N : ℕ)
164164
(P_0 : P 0)
165165
(P_C_mul_pow : ∀ n : ℕ, ∀ r : R, r ≠ 0 → n ≤ N → P (C r * X ^ n))
166-
(P_C_add : ∀ f g : polynomial R, f.nat_degree ≤ N → g.nat_degree ≤ N → P f → P g → P (f + g)) :
166+
(P_C_add : ∀ f g : polynomial R, f.nat_degree < g.nat_degree →
167+
g.nat_degree ≤ N → P f → P g → P (f + g)) :
167168
∀ f : polynomial R, f.nat_degree ≤ N → P f :=
168169
begin
169170
intros f df,
@@ -172,18 +173,26 @@ begin
172173
induction c with c hc,
173174
{ assume f df f0,
174175
convert P_0,
175-
simpa only [support_eq_empty, card_eq_zero] using f0},
176-
177-
-- exact λ f df f0, by rwa (finsupp.support_eq_empty.mp (card_eq_zero.mp f0)) },
176+
simpa only [support_eq_empty, card_eq_zero] using f0 },
178177
{ intros f df f0,
179-
rw ← erase_lead_add_C_mul_X_pow f,
180-
refine P_C_add f.erase_lead _ (erase_lead_nat_degree_le.trans df) _ _ _,
178+
rw [← erase_lead_add_C_mul_X_pow f],
179+
cases c,
180+
{ convert P_C_mul_pow f.nat_degree f.leading_coeff _ df,
181+
{ convert zero_add _,
182+
rw [← card_support_eq_zero, erase_lead_card_support f0] },
183+
{ rw [leading_coeff_ne_zero, ne.def, ← card_support_eq_zero, f0],
184+
exact zero_ne_one.symm } },
185+
refine P_C_add f.erase_lead _ _ _ _ _,
186+
{ refine (erase_lead_nat_degree_lt _).trans_le (le_of_eq _),
187+
{ exact (nat.succ_le_succ (nat.succ_le_succ (nat.zero_le _))).trans f0.ge },
188+
{ rw [nat_degree_C_mul_X_pow _ _ (leading_coeff_ne_zero.mpr _)],
189+
rintro rfl,
190+
simpa using f0 } },
181191
{ exact (nat_degree_C_mul_X_pow_le f.leading_coeff f.nat_degree).trans df },
182192
{ exact hc _ (erase_lead_nat_degree_le.trans df) (erase_lead_card_support f0) },
183193
{ refine P_C_mul_pow _ _ _ df,
184-
rw [ne.def, leading_coeff_eq_zero],
185-
rintro rfl,
186-
exact not_le.mpr c.succ_pos f0.ge } }
194+
rw [ne.def, leading_coeff_eq_zero, ← card_support_eq_zero, f0],
195+
exact nat.succ_ne_zero _ } }
187196
end
188197

189198
end polynomial

0 commit comments

Comments
 (0)