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

Commit e22789e

Browse files
committed
feat(algebra/big_operators/finprod): add a few lemmas (#8261)
* add `finprod_eq_single` and `finsum_eq_single`; * add `finprod_induction` and `finsum_induction`; * add `single_le_finprod` and `single_le_finsum`; * add `one_le_finprod'` and `finsum_nonneg`.
1 parent d5c6f61 commit e22789e

File tree

2 files changed

+42
-17
lines changed

2 files changed

+42
-17
lines changed

src/algebra/big_operators/finprod.lean

Lines changed: 39 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -137,15 +137,18 @@ by { rw ← finprod_one, congr }
137137
@[simp, to_additive] lemma finprod_false (f : false → M) : ∏ᶠ i, f i = 1 :=
138138
finprod_of_is_empty _
139139

140-
@[to_additive] lemma finprod_unique [unique α] (f : α → M) : ∏ᶠ i, f i = f (default α) :=
140+
@[to_additive] lemma finprod_eq_single (f : α → M) (a : α) (ha : ∀ x ≠ a, f x = 1) :
141+
∏ᶠ x, f x = f a :=
141142
begin
142-
have : mul_support (f ∘ plift.down) ⊆ (finset.univ : finset (plift α)),
143-
from λ x _, finset.mem_univ x,
144-
rw [finprod_eq_prod_plift_of_mul_support_subset this, univ_unique,
145-
finset.prod_singleton],
146-
exact congr_arg f (plift.down_up _)
143+
have : mul_support (f ∘ plift.down) ⊆ ({plift.up a} : finset (plift α)),
144+
{ intro x, contrapose,
145+
simpa [plift.eq_up_iff_down_eq] using ha x.down },
146+
rw [finprod_eq_prod_plift_of_mul_support_subset this, finset.prod_singleton],
147147
end
148148

149+
@[to_additive] lemma finprod_unique [unique α] (f : α → M) : ∏ᶠ i, f i = f (default α) :=
150+
finprod_eq_single f (default α) $ λ x hx, (hx $ unique.eq_default _).elim
151+
149152
@[simp, to_additive] lemma finprod_true (f : true → M) : ∏ᶠ i, f i = f trivial :=
150153
@finprod_unique M true _ ⟨⟨trivial⟩, λ _, rfl⟩ f
151154

@@ -172,15 +175,30 @@ by { subst q, exact finprod_congr hfg }
172175

173176
attribute [congr] finsum_congr_Prop
174177

175-
lemma finprod_nonneg {R : Type*} [ordered_comm_semiring R] {f : α → R} (hf : ∀ x, 0 ≤ f x) :
176-
0 ≤ ∏ᶠ x, f x :=
178+
/-- To prove a property of a finite product, it suffices to prove that the property is
179+
multiplicative and holds on multipliers. -/
180+
@[to_additive] lemma finprod_induction {f : α → M} (p : M → Prop) (hp₀ : p 1)
181+
(hp₁ : ∀ x y, p x → p y → p (x * y)) (hp₂ : ∀ i, p (f i)) :
182+
p (∏ᶠ i, f i) :=
177183
begin
178184
rw finprod,
179185
split_ifs,
180-
{ exact finset.prod_nonneg (λ x _, hf _) },
181-
{ exact zero_le_one }
186+
exacts [finset.prod_induction _ _ hp₁ hp₀ (λ i hi, hp₂ _), hp₀]
182187
end
183188

189+
/-- To prove a property of a finite sum, it suffices to prove that the property is
190+
additive and holds on summands. -/
191+
add_decl_doc finsum_induction
192+
193+
lemma finprod_nonneg {R : Type*} [ordered_comm_semiring R] {f : α → R} (hf : ∀ x, 0 ≤ f x) :
194+
0 ≤ ∏ᶠ x, f x :=
195+
finprod_induction (λ x, 0 ≤ x) zero_le_one (λ x y, mul_nonneg) hf
196+
197+
@[to_additive finsum_nonneg]
198+
lemma one_le_finprod' {M : Type*} [ordered_comm_monoid M] {f : α → M} (hf : ∀ i, 1 ≤ f i) :
199+
1 ≤ ∏ᶠ i, f i :=
200+
finprod_induction _ le_rfl (λ _ _, one_le_mul) hf
201+
184202
@[to_additive] lemma monoid_hom.map_finprod_plift (f : M →* N) (g : α → M)
185203
(h : finite (mul_support $ g ∘ plift.down)) :
186204
f (∏ᶠ x, g x) = ∏ᶠ x, f (g x) :=
@@ -651,19 +669,23 @@ multiplicative and holds on multipliers. -/
651669
@[to_additive] lemma finprod_mem_induction (p : M → Prop) (hp₀ : p 1)
652670
(hp₁ : ∀ x y, p x → p y → p (x * y)) (hp₂ : ∀ x ∈ s, p $ f x) :
653671
p (∏ᶠ i ∈ s, f i) :=
654-
begin
655-
by_cases hs : (s ∩ mul_support f).finite,
656-
{ rw [finprod_mem_eq_prod _ hs],
657-
refine finset.prod_induction _ p hp₁ hp₀ (λ x hx, hp₂ x _),
658-
rw hs.mem_to_finset at hx, exact hx.1 },
659-
{ exact (finprod_mem_eq_one_of_infinite hs).symm ▸ hp₀ }
660-
end
672+
finprod_induction _ hp₀ hp₁ $ λ x, finprod_induction _ hp₀ hp₁ $ hp₂ x
661673

662674
lemma finprod_cond_nonneg {R : Type*} [ordered_comm_semiring R] {p : α → Prop} {f : α → R}
663675
(hf : ∀ x, p x → 0 ≤ f x) :
664676
0 ≤ ∏ᶠ x (h : p x), f x :=
665677
finprod_nonneg $ λ x, finprod_nonneg $ hf x
666678

679+
@[to_additive]
680+
lemma single_le_finprod {M : Type*} [ordered_comm_monoid M] (i : α) {f : α → M}
681+
(hf : finite (mul_support f)) (h : ∀ j, 1 ≤ f j) :
682+
f i ≤ ∏ᶠ j, f j :=
683+
by classical;
684+
calc f i ≤ ∏ j in insert i hf.to_finset, f j :
685+
finset.single_le_prod' (λ j hj, h j) (finset.mem_insert_self _ _)
686+
... = ∏ᶠ j, f j :
687+
(finprod_eq_prod_of_mul_support_to_finset_subset _ hf (finset.subset_insert _ _)).symm
688+
667689
lemma finprod_eq_zero {M₀ : Type*} [comm_monoid_with_zero M₀] (f : α → M₀) (x : α)
668690
(hx : f x = 0) (hf : finite (mul_support f)) :
669691
∏ᶠ x, f x = 0 :=

src/data/equiv/basic.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2001,6 +2001,9 @@ end swap
20012001

20022002
end equiv
20032003

2004+
lemma plift.eq_up_iff_down_eq {x : plift α} {y : α} : x = plift.up y ↔ x.down = y :=
2005+
equiv.plift.eq_symm_apply
2006+
20042007
lemma function.injective.map_swap {α β : Type*} [decidable_eq α] [decidable_eq β]
20052008
{f : α → β} (hf : function.injective f) (x y z : α) :
20062009
f (equiv.swap x y z) = equiv.swap (f x) (f y) (f z) :=

0 commit comments

Comments
 (0)