@@ -151,10 +151,10 @@ coeff_monomial_same 0 1
151
151
lemma monomial_zero_one : monomial R (0 : σ →₀ ℕ) 1 = 1 := rfl
152
152
153
153
instance : has_mul (mv_power_series σ R) :=
154
- ⟨λ φ ψ n, ∑ p in ( finsupp.antidiagonal n).support , coeff R p.1 φ * coeff R p.2 ψ⟩
154
+ ⟨λ φ ψ n, ∑ p in finsupp.antidiagonal n, coeff R p.1 φ * coeff R p.2 ψ⟩
155
155
156
156
lemma coeff_mul : coeff R n (φ * ψ) =
157
- ∑ p in ( finsupp.antidiagonal n).support , coeff R p.1 φ * coeff R p.2 ψ := rfl
157
+ ∑ p in finsupp.antidiagonal n, coeff R p.1 φ * coeff R p.2 ψ := rfl
158
158
159
159
protected lemma zero_mul : (0 : mv_power_series σ R) * φ = 0 :=
160
160
ext $ λ n, by simp [coeff_mul]
@@ -165,21 +165,21 @@ ext $ λ n, by simp [coeff_mul]
165
165
lemma coeff_monomial_mul (a : R) :
166
166
coeff R m (monomial R n a * φ) = if n ≤ m then a * coeff R (m - n) φ else 0 :=
167
167
begin
168
- have : ∀ p ∈ ( antidiagonal m).support ,
168
+ have : ∀ p ∈ antidiagonal m,
169
169
coeff R (p : (σ →₀ ℕ) × (σ →₀ ℕ)).1 (monomial R n a) * coeff R p.2 φ ≠ 0 → p.1 = n :=
170
170
λ p _ hp, eq_of_coeff_monomial_ne_zero (left_ne_zero_of_mul hp),
171
- rw [coeff_mul, ← finset.sum_filter_of_ne this , antidiagonal_support_filter_fst_eq ,
171
+ rw [coeff_mul, ← finset.sum_filter_of_ne this , antidiagonal_filter_fst_eq ,
172
172
finset.sum_ite_index],
173
173
simp only [finset.sum_singleton, coeff_monomial_same, finset.sum_empty]
174
174
end
175
175
176
176
lemma coeff_mul_monomial (a : R) :
177
177
coeff R m (φ * monomial R n a) = if n ≤ m then coeff R (m - n) φ * a else 0 :=
178
178
begin
179
- have : ∀ p ∈ ( antidiagonal m).support ,
179
+ have : ∀ p ∈ antidiagonal m,
180
180
coeff R (p : (σ →₀ ℕ) × (σ →₀ ℕ)).1 φ * coeff R p.2 (monomial R n a) ≠ 0 → p.2 = n :=
181
181
λ p _ hp, eq_of_coeff_monomial_ne_zero (right_ne_zero_of_mul hp),
182
- rw [coeff_mul, ← finset.sum_filter_of_ne this , antidiagonal_support_filter_snd_eq ,
182
+ rw [coeff_mul, ← finset.sum_filter_of_ne this , antidiagonal_filter_snd_eq ,
183
183
finset.sum_ite_index],
184
184
simp only [finset.sum_singleton, coeff_monomial_same, finset.sum_empty]
185
185
end
@@ -218,7 +218,7 @@ begin
218
218
ext1 n,
219
219
simp only [coeff_mul, finset.sum_mul, finset.mul_sum, finset.sum_sigma'],
220
220
refine finset.sum_bij (λ p _, ⟨(p.2 .1 , p.2 .2 + p.1 .2 ), (p.2 .2 , p.1 .2 )⟩) _ _ _ _;
221
- simp only [mem_antidiagonal_support , finset.mem_sigma, heq_iff_eq, prod.mk.inj_iff, and_imp,
221
+ simp only [mem_antidiagonal , finset.mem_sigma, heq_iff_eq, prod.mk.inj_iff, and_imp,
222
222
exists_prop],
223
223
{ rintros ⟨⟨i,j⟩, ⟨k,l⟩⟩, dsimp only, rintro rfl rfl,
224
224
simp [add_assoc] },
@@ -246,7 +246,7 @@ end semiring
246
246
247
247
instance [comm_semiring R] : comm_semiring (mv_power_series σ R) :=
248
248
{ mul_comm := λ φ ψ, ext $ λ n, by simpa only [coeff_mul, mul_comm]
249
- using sum_antidiagonal_support_swap n (λ a b, coeff R a φ * coeff R b ψ),
249
+ using sum_antidiagonal_swap n (λ a b, coeff R a φ * coeff R b ψ),
250
250
.. mv_power_series.semiring }
251
251
252
252
instance [ring R] : ring (mv_power_series σ R) :=
@@ -527,21 +527,21 @@ begin
527
527
{ rintros ⟨φ, rfl⟩ m h,
528
528
rw [coeff_mul, finset.sum_eq_zero],
529
529
rintros ⟨i,j⟩ hij, rw [coeff_X_pow, if_neg, zero_mul],
530
- contrapose! h, subst i, rw finsupp.mem_antidiagonal_support at hij,
530
+ contrapose! h, subst i, rw finsupp.mem_antidiagonal at hij,
531
531
rw [← hij, finsupp.add_apply, finsupp.single_eq_same], exact nat.le_add_right n _ },
532
532
{ intro h, refine ⟨λ m, coeff R (m + (single s n)) φ, _⟩,
533
533
ext m, by_cases H : m - single s n + single s n = m,
534
534
{ rw [coeff_mul, finset.sum_eq_single (single s n, m - single s n)],
535
535
{ rw [coeff_X_pow, if_pos rfl, one_mul],
536
536
simpa using congr_arg (λ (m : σ →₀ ℕ), coeff R m φ) H.symm },
537
- { rintros ⟨i,j⟩ hij hne, rw finsupp.mem_antidiagonal_support at hij,
537
+ { rintros ⟨i,j⟩ hij hne, rw finsupp.mem_antidiagonal at hij,
538
538
rw coeff_X_pow, split_ifs with hi,
539
539
{ exfalso, apply hne, rw [← hij, ← hi, prod.mk.inj_iff], refine ⟨rfl, _⟩,
540
540
ext t, simp only [nat.add_sub_cancel_left, finsupp.add_apply, finsupp.nat_sub_apply] },
541
541
{ exact zero_mul _ } },
542
- { intro hni, exfalso, apply hni, rwa [finsupp.mem_antidiagonal_support , add_comm] } },
542
+ { intro hni, exfalso, apply hni, rwa [finsupp.mem_antidiagonal , add_comm] } },
543
543
{ rw [h, coeff_mul, finset.sum_eq_zero],
544
- { rintros ⟨i,j⟩ hij, rw finsupp.mem_antidiagonal_support at hij,
544
+ { rintros ⟨i,j⟩ hij, rw finsupp.mem_antidiagonal at hij,
545
545
rw coeff_X_pow, split_ifs with hi,
546
546
{ exfalso, apply H, rw [← hij, hi], ext,
547
547
rw [coe_add, coe_add, pi.add_apply, pi.add_apply, nat_add_sub_cancel_left, add_comm], },
@@ -576,15 +576,15 @@ well-founded recursion on the coeffients of the inverse.
576
576
an inverse of the constant coefficient `inv_of_unit`.-/
577
577
protected noncomputable def inv.aux (a : R) (φ : mv_power_series σ R) : mv_power_series σ R
578
578
| n := if n = 0 then a else
579
- - a * ∑ x in n.antidiagonal.support ,
579
+ - a * ∑ x in n.antidiagonal,
580
580
if h : x.2 < n then coeff R x.1 φ * inv.aux x.2 else 0
581
581
using_well_founded
582
582
{ rel_tac := λ _ _, `[exact ⟨_, finsupp.lt_wf σ⟩],
583
583
dec_tac := tactic.assumption }
584
584
585
585
lemma coeff_inv_aux (n : σ →₀ ℕ) (a : R) (φ : mv_power_series σ R) :
586
586
coeff R n (inv.aux a φ) = if n = 0 then a else
587
- - a * ∑ x in n.antidiagonal.support ,
587
+ - a * ∑ x in n.antidiagonal,
588
588
if x.2 < n then coeff R x.1 φ * coeff R x.2 (inv.aux a φ) else 0 :=
589
589
show inv.aux a φ n = _, by { rw inv.aux, refl }
590
590
@@ -594,7 +594,7 @@ inv.aux (↑u⁻¹) φ
594
594
595
595
lemma coeff_inv_of_unit (n : σ →₀ ℕ) (φ : mv_power_series σ R) (u : units R) :
596
596
coeff R n (inv_of_unit φ u) = if n = 0 then ↑u⁻¹ else
597
- - ↑u⁻¹ * ∑ x in n.antidiagonal.support ,
597
+ - ↑u⁻¹ * ∑ x in n.antidiagonal,
598
598
if x.2 < n then coeff R x.1 φ * coeff R x.2 (inv_of_unit φ u) else 0 :=
599
599
coeff_inv_aux n (↑u⁻¹) φ
600
600
@@ -607,16 +607,16 @@ lemma mul_inv_of_unit (φ : mv_power_series σ R) (u : units R) (h : constant_co
607
607
ext $ λ n, if H : n = 0 then by { rw H, simp [coeff_mul, support_single_ne_zero, h], }
608
608
else
609
609
begin
610
- have : ((0 : σ →₀ ℕ), n) ∈ n.antidiagonal.support ,
611
- { rw [finsupp.mem_antidiagonal_support , zero_add] },
610
+ have : ((0 : σ →₀ ℕ), n) ∈ n.antidiagonal,
611
+ { rw [finsupp.mem_antidiagonal , zero_add] },
612
612
rw [coeff_one, if_neg H, coeff_mul,
613
613
← finset.insert_erase this , finset.sum_insert (finset.not_mem_erase _ _),
614
614
coeff_zero_eq_constant_coeff_apply, h, coeff_inv_of_unit, if_neg H,
615
615
neg_mul_eq_neg_mul_symm, mul_neg_eq_neg_mul_symm, units.mul_inv_cancel_left,
616
616
← finset.insert_erase this , finset.sum_insert (finset.not_mem_erase _ _),
617
617
finset.insert_erase this , if_neg (not_lt_of_ge $ le_refl _), zero_add, add_comm,
618
618
← sub_eq_add_neg, sub_eq_zero, finset.sum_congr rfl],
619
- rintros ⟨i,j⟩ hij, rw [finset.mem_erase, finsupp.mem_antidiagonal_support ] at hij,
619
+ rintros ⟨i,j⟩ hij, rw [finset.mem_erase, finsupp.mem_antidiagonal ] at hij,
620
620
cases hij with h₁ h₂,
621
621
subst n, rw if_pos,
622
622
suffices : (0 : _) + j < i + j, {simpa},
@@ -681,7 +681,7 @@ instance : has_inv (mv_power_series σ k) := ⟨mv_power_series.inv⟩
681
681
682
682
lemma coeff_inv (n : σ →₀ ℕ) (φ : mv_power_series σ k) :
683
683
coeff k n (φ⁻¹) = if n = 0 then (constant_coeff σ k φ)⁻¹ else
684
- - (constant_coeff σ k φ)⁻¹ * ∑ x in n.antidiagonal.support ,
684
+ - (constant_coeff σ k φ)⁻¹ * ∑ x in n.antidiagonal,
685
685
if x.2 < n then coeff k x.1 φ * coeff k x.2 (φ⁻¹) else 0 :=
686
686
coeff_inv_aux n _ φ
687
687
@@ -942,13 +942,13 @@ begin
942
942
symmetry,
943
943
apply finset.sum_bij (λ (p : ℕ × ℕ) h, (single () p.1 , single () p.2 )),
944
944
{ rintros ⟨i,j⟩ hij, rw finset.nat.mem_antidiagonal at hij,
945
- rw [finsupp.mem_antidiagonal_support , ← finsupp.single_add, hij], },
945
+ rw [finsupp.mem_antidiagonal , ← finsupp.single_add, hij], },
946
946
{ rintros ⟨i,j⟩ hij, refl },
947
947
{ rintros ⟨i,j⟩ ⟨k,l⟩ hij hkl,
948
948
simpa only [prod.mk.inj_iff, finsupp.unique_single_eq_iff] using id },
949
949
{ rintros ⟨f,g⟩ hfg,
950
950
refine ⟨(f (), g ()), _, _⟩,
951
- { rw finsupp.mem_antidiagonal_support at hfg,
951
+ { rw finsupp.mem_antidiagonal at hfg,
952
952
rw [finset.nat.mem_antidiagonal, ← finsupp.add_apply, hfg, finsupp.single_eq_same] },
953
953
{ rw prod.mk.inj_iff, dsimp,
954
954
exact ⟨finsupp.unique_single f, finsupp.unique_single g⟩ } }
@@ -1166,7 +1166,7 @@ begin
1166
1166
symmetry,
1167
1167
apply finset.sum_bij (λ (p : ℕ × ℕ) h, (single () p.1 , single () p.2 )),
1168
1168
{ rintros ⟨i,j⟩ hij, rw finset.nat.mem_antidiagonal at hij,
1169
- rw [finsupp.mem_antidiagonal_support , ← finsupp.single_add, hij], },
1169
+ rw [finsupp.mem_antidiagonal , ← finsupp.single_add, hij], },
1170
1170
{ rintros ⟨i,j⟩ hij,
1171
1171
by_cases H : j < n,
1172
1172
{ rw [if_pos H, if_pos], {refl},
@@ -1180,7 +1180,7 @@ begin
1180
1180
simpa only [prod.mk.inj_iff, finsupp.unique_single_eq_iff] using id },
1181
1181
{ rintros ⟨f,g⟩ hfg,
1182
1182
refine ⟨(f (), g ()), _, _⟩,
1183
- { rw finsupp.mem_antidiagonal_support at hfg,
1183
+ { rw finsupp.mem_antidiagonal at hfg,
1184
1184
rw [finset.nat.mem_antidiagonal, ← finsupp.add_apply, hfg, finsupp.single_eq_same] },
1185
1185
{ rw prod.mk.inj_iff, dsimp,
1186
1186
exact ⟨finsupp.unique_single f, finsupp.unique_single g⟩ } }
0 commit comments