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

Commit f61cc72

Browse files
YaelDillieskhwilson
authored andcommitted
feat(data/finset/pointwise): Singleton arithmetic (#15435)
Lemmas about pointwise operations with a singleton. Also make `s` explicit in `finset.card_inv`.
1 parent dcd4122 commit f61cc72

File tree

2 files changed

+73
-11
lines changed

2 files changed

+73
-11
lines changed

src/data/finset/n_ary.lean

Lines changed: 35 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -140,6 +140,41 @@ begin
140140
mem_insert_self _ _, ha⟩, h.trans $ image₂_subset (subset_insert _ _) $ subset_insert _ _⟩⟩,
141141
end
142142

143+
variables (s t)
144+
145+
lemma card_image₂_singleton_left (hf : injective (f a)) : (image₂ f {a} t).card = t.card :=
146+
by rw [image₂_singleton_left, card_image_of_injective _ hf]
147+
148+
lemma card_image₂_singleton_right (hf : injective (λ a, f a b)) : (image₂ f s {b}).card = s.card :=
149+
by rw [image₂_singleton_right, card_image_of_injective _ hf]
150+
151+
lemma image₂_singleton_inter [decidable_eq β] (t₁ t₂ : finset β) (hf : injective (f a)) :
152+
image₂ f {a} (t₁ ∩ t₂) = image₂ f {a} t₁ ∩ image₂ f {a} t₂ :=
153+
by simp_rw [image₂_singleton_left, image_inter _ _ hf]
154+
155+
lemma image₂_inter_singleton [decidable_eq α] (s₁ s₂ : finset α) (hf : injective (λ a, f a b)) :
156+
image₂ f (s₁ ∩ s₂) {b} = image₂ f s₁ {b} ∩ image₂ f s₂ {b} :=
157+
by simp_rw [image₂_singleton_right, image_inter _ _ hf]
158+
159+
lemma card_le_card_image₂_left {s : finset α} (hs : s.nonempty) (hf : ∀ a, injective (f a)) :
160+
t.card ≤ (image₂ f s t).card :=
161+
begin
162+
obtain ⟨a, ha⟩ := hs,
163+
rw ←card_image₂_singleton_left _ (hf a),
164+
exact card_le_of_subset (image₂_subset_right $ singleton_subset_iff.2 ha),
165+
end
166+
167+
lemma card_le_card_image₂_right {t : finset β} (ht : t.nonempty)
168+
(hf : ∀ b, injective (λ a, f a b)) :
169+
s.card ≤ (image₂ f s t).card :=
170+
begin
171+
obtain ⟨b, hb⟩ := ht,
172+
rw ←card_image₂_singleton_right _ (hf b),
173+
exact card_le_of_subset (image₂_subset_left $ singleton_subset_iff.2 hb),
174+
end
175+
176+
variables {s t}
177+
143178
lemma bUnion_image_left : s.bUnion (λ a, t.image $ f a) = image₂ f s t :=
144179
coe_injective $ by { push_cast, exact set.Union_image_left _ }
145180

src/data/finset/pointwise.lean

Lines changed: 38 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -124,10 +124,10 @@ end has_inv
124124
open_locale pointwise
125125

126126
section has_involutive_inv
127-
variables [decidable_eq α] [has_involutive_inv α] {s t : finset α}
127+
variables [decidable_eq α] [has_involutive_inv α] (s : finset α)
128128

129129
@[simp, norm_cast, to_additive]
130-
lemma coe_inv (s : finset α) : ↑(s⁻¹) = (s : set α)⁻¹ := coe_image.trans set.image_inv
130+
lemma coe_inv : ↑(s⁻¹) = (s : set α)⁻¹ := coe_image.trans set.image_inv
131131

132132
@[simp, to_additive] lemma card_inv : s⁻¹.card = s.card := card_image_of_injective _ inv_injective
133133

@@ -178,9 +178,8 @@ image₂_nonempty_iff
178178
@[to_additive] lemma nonempty.of_mul_left : (s * t).nonempty → s.nonempty := nonempty.of_image₂_left
179179
@[to_additive] lemma nonempty.of_mul_right : (s * t).nonempty → t.nonempty :=
180180
nonempty.of_image₂_right
181-
@[simp, to_additive] lemma mul_singleton (a : α) : s * {a} = s.image (* a) := image₂_singleton_right
182-
@[simp, to_additive] lemma singleton_mul (a : α) : {a} * s = s.image ((*) a) :=
183-
image₂_singleton_left
181+
@[to_additive] lemma mul_singleton (a : α) : s * {a} = s.image (* a) := image₂_singleton_right
182+
@[to_additive] lemma singleton_mul (a : α) : {a} * s = s.image ((*) a) := image₂_singleton_left
184183
@[simp, to_additive] lemma singleton_mul_singleton (a b : α) : ({a} : finset α) * {b} = {a * b} :=
185184
image₂_singleton
186185

@@ -643,9 +642,9 @@ image₂_nonempty_iff
643642
nonempty.of_image₂_left
644643
@[to_additive] lemma nonempty.of_smul_right : (s • t).nonempty → t.nonempty :=
645644
nonempty.of_image₂_right
646-
@[simp, to_additive] lemma smul_singleton (b : β) : s • ({b} : finset β) = s.image (• b) :=
645+
@[to_additive] lemma smul_singleton (b : β) : s • ({b} : finset β) = s.image (• b) :=
647646
image₂_singleton_right
648-
@[simp, to_additive] lemma singleton_smul (a : α) : ({a} : finset α) • t = t.image ((•) a) :=
647+
@[to_additive] lemma singleton_smul (a : α) : ({a} : finset α) • t = t.image ((•) a) :=
649648
image₂_singleton_left
650649
@[simp, to_additive] lemma singleton_smul_singleton (a : α) (b : β) :
651650
({a} : finset α) • ({b} : finset β) = {a • b} :=
@@ -709,8 +708,7 @@ lemma nonempty.of_vsub_left : (s -ᵥ t : finset α).nonempty → s.nonempty :=
709708
lemma nonempty.of_vsub_right : (s -ᵥ t : finset α).nonempty → t.nonempty := nonempty.of_image₂_right
710709
@[simp] lemma vsub_singleton (b : β) : s -ᵥ ({b} : finset β) = s.image (-ᵥ b) :=
711710
image₂_singleton_right
712-
@[simp] lemma singleton_vsub (a : β) : ({a} : finset β) -ᵥ t = t.image ((-ᵥ) a) :=
713-
image₂_singleton_left
711+
lemma singleton_vsub (a : β) : ({a} : finset β) -ᵥ t = t.image ((-ᵥ) a) := image₂_singleton_left
714712
@[simp] lemma singleton_vsub_singleton (a b : β) : ({a} : finset β) -ᵥ {b} = {a -ᵥ b} :=
715713
image₂_singleton
716714

@@ -888,11 +886,40 @@ coe_injective.no_zero_smul_divisors _ coe_zero coe_smul_finset
888886

889887
end instances
890888

891-
@[to_additive] lemma pairwise_disjoint_smul_iff [decidable_eq α] [left_cancel_semigroup α]
892-
{s : set α} {t : finset α} :
889+
section left_cancel_semigroup
890+
variables [left_cancel_semigroup α] [decidable_eq α] (s t : finset α) (a : α)
891+
892+
@[to_additive] lemma pairwise_disjoint_smul_iff {s : set α} {t : finset α} :
893893
s.pairwise_disjoint (• t) ↔ ((s : set α) ×ˢ (t : set α) : set (α × α)).inj_on (λ p, p.1 * p.2) :=
894894
by simp_rw [←pairwise_disjoint_coe, coe_smul_finset, set.pairwise_disjoint_smul_iff]
895895

896+
@[simp, to_additive] lemma card_singleton_mul : ({a} * t).card = t.card :=
897+
card_image₂_singleton_left _ $ mul_right_injective _
898+
899+
@[to_additive] lemma singleton_mul_inter : {a} * (s ∩ t) = ({a} * s) ∩ ({a} * t) :=
900+
image₂_singleton_inter _ _ $ mul_right_injective _
901+
902+
@[to_additive] lemma card_le_card_mul_left {s : finset α} (hs : s.nonempty) :
903+
t.card ≤ (s * t).card :=
904+
card_le_card_image₂_left _ hs mul_right_injective
905+
906+
end left_cancel_semigroup
907+
908+
section
909+
variables [right_cancel_semigroup α] [decidable_eq α] (s t : finset α) (a : α)
910+
911+
@[simp, to_additive] lemma card_mul_singleton : (s * {a}).card = s.card :=
912+
card_image₂_singleton_right _ $ mul_left_injective _
913+
914+
@[to_additive] lemma inter_mul_singleton : (s ∩ t) * {a} = (s * {a}) ∩ (t * {a}) :=
915+
image₂_inter_singleton _ _ $ mul_left_injective _
916+
917+
@[to_additive] lemma card_le_card_mul_right {t : finset α} (ht : t.nonempty) :
918+
s.card ≤ (s * t).card :=
919+
card_le_card_image₂_right _ ht mul_left_injective
920+
921+
end
922+
896923
open_locale pointwise
897924

898925
section group

0 commit comments

Comments
 (0)