@@ -319,6 +319,21 @@ protected theorem induction_on {α : Type*} {p : finset α → Prop} [decidable_
319
319
(s : finset α) (h₁ : p ∅) (h₂ : ∀ ⦃a : α⦄ {s : finset α}, a ∉ s → p s → p (insert a s)) : p s :=
320
320
finset.induction h₁ h₂ s
321
321
322
+ /-- Inserting an element to a finite set is equivalent to the option type. -/
323
+ def subtype_insert_equiv_option {t : finset α} {x : α} (h : x ∉ t) :
324
+ {i // i ∈ insert x t} ≃ option {i // i ∈ t} :=
325
+ begin
326
+ refine
327
+ { to_fun := λ y, if h : y.1 = x then none else some ⟨y, (finset.mem_insert.mp y.2 ).resolve_left h⟩,
328
+ inv_fun := λ y, y.elim ⟨x, finset.mem_insert_self _ _⟩ $ λ z, ⟨z.1 , finset.mem_insert_of_mem z.2 ⟩,
329
+ .. },
330
+ { intro y, by_cases h : y.1 = x, simp only [subtype.ext, h, option.elim, dif_pos],
331
+ simp only [h, option.elim, dif_neg, not_false_iff, subtype.coe_eta] },
332
+ { rintro (_|y), simp only [option.elim, dif_pos],
333
+ have : y.val ≠ x, { rintro ⟨⟩, exact h y.2 },
334
+ simp only [this , option.elim, subtype.eta, dif_neg, not_false_iff, subtype.coe_mk] },
335
+ end
336
+
322
337
/-! ### union -/
323
338
324
339
/-- `s ∪ t` is the set such that `a ∈ s ∪ t` iff `a ∈ s` or `a ∈ t`. -/
@@ -1542,6 +1557,13 @@ lemma card_union_le [decidable_eq α] (s t : finset α) :
1542
1557
(s ∪ t).card ≤ s.card + t.card :=
1543
1558
card_union_add_card_inter s t ▸ le_add_right _ _
1544
1559
1560
+ lemma card_union_eq [decidable_eq α] {s t : finset α} (h : disjoint s t) :
1561
+ (s ∪ t).card = s.card + t.card :=
1562
+ begin
1563
+ rw [← card_union_add_card_inter],
1564
+ convert (add_zero _).symm, rw [card_eq_zero], rwa [disjoint_iff] at h
1565
+ end
1566
+
1545
1567
lemma surj_on_of_inj_on_of_card_le {s : finset α} {t : finset β}
1546
1568
(f : Π a ∈ s, β) (hf : ∀ a ha, f a ha ∈ t)
1547
1569
(hinj : ∀ a₁ a₂ ha₁ ha₂, f a₁ ha₁ = f a₂ ha₂ → a₁ = a₂)
@@ -2064,10 +2086,15 @@ by letI := classical.dec_eq β; from
2064
2086
⟨ λh b hb, lt_of_le_of_lt (le_sup hb) h,
2065
2087
finset.induction_on s (by simp [ha]) (by simp {contextual := tt}) ⟩
2066
2088
2067
- lemma comp_sup_eq_sup_comp [is_total α (≤)] {γ : Type } [semilattice_sup_bot γ]
2068
- (g : α → γ) (mono_g : monotone g) (bot : g ⊥ = ⊥) : g (s.sup f) = s.sup (g ∘ f) :=
2089
+ lemma comp_sup_eq_sup_comp [semilattice_sup_bot γ] {s : finset β}
2090
+ {f : β → α} (g : α → γ) (g_sup : ∀ x y, g (x ⊔ y) = g x ⊔ g y) (bot : g ⊥ = ⊥) :
2091
+ g (s.sup f) = s.sup (g ∘ f) :=
2069
2092
by letI := classical.dec_eq β; from
2070
- finset.induction_on s (by simp [bot]) (by simp [mono_g.map_sup] {contextual := tt})
2093
+ finset.induction_on s (by simp [bot]) (by simp [g_sup] {contextual := tt})
2094
+
2095
+ lemma comp_sup_eq_sup_comp_of_is_total [is_total α (≤)] {γ : Type } [semilattice_sup_bot γ]
2096
+ (g : α → γ) (mono_g : monotone g) (bot : g ⊥ = ⊥) : g (s.sup f) = s.sup (g ∘ f) :=
2097
+ comp_sup_eq_sup_comp g mono_g.map_sup bot
2071
2098
2072
2099
theorem subset_range_sup_succ (s : finset ℕ) : s ⊆ range (s.sup id).succ :=
2073
2100
λ n hn, mem_range.2 $ nat.lt_succ_of_le $ le_sup hn
@@ -2126,9 +2153,14 @@ le_inf $ assume b hb, inf_le (h hb)
2126
2153
lemma lt_inf_iff [h : is_total α (≤)] {a : α} (ha : a < ⊤) : a < s.inf f ↔ (∀b ∈ s, a < f b) :=
2127
2154
@sup_lt_iff (order_dual α) _ _ _ _ (@is_total.swap α _ h) _ ha
2128
2155
2129
- lemma comp_inf_eq_inf_comp [h : is_total α (≤)] {γ : Type } [semilattice_inf_top γ]
2156
+ lemma comp_inf_eq_inf_comp [semilattice_inf_top γ] {s : finset β}
2157
+ {f : β → α} (g : α → γ) (g_inf : ∀ x y, g (x ⊓ y) = g x ⊓ g y) (top : g ⊤ = ⊤) :
2158
+ g (s.inf f) = s.inf (g ∘ f) :=
2159
+ @comp_sup_eq_sup_comp (order_dual α) _ (order_dual γ) _ _ _ _ _ g_inf top
2160
+
2161
+ lemma comp_inf_eq_inf_comp_of_is_total [h : is_total α (≤)] {γ : Type } [semilattice_inf_top γ]
2130
2162
(g : α → γ) (mono_g : monotone g) (top : g ⊤ = ⊤) : g (s.inf f) = s.inf (g ∘ f) :=
2131
- @comp_sup_eq_sup_comp (order_dual α) _ _ _ _ (@is_total.swap α _ h) _ _ _ mono_g.order_dual top
2163
+ comp_inf_eq_inf_comp g mono_g.map_inf top
2132
2164
2133
2165
end inf
2134
2166
@@ -2874,14 +2906,6 @@ end
2874
2906
2875
2907
@[simp] lemma Ico_ℤ.card (l u : ℤ) : (Ico_ℤ l u).card = (u - l).to_nat := by simp [Ico_ℤ]
2876
2908
2877
- lemma supr_coe [has_Sup β] (f : α → β) (s : finset α) :
2878
- (⨆ x ∈ (↑s : set α), f x) = ⨆ x ∈ s, f x :=
2879
- rfl
2880
-
2881
- lemma infi_coe [has_Inf β] (f : α → β) (s : finset α) :
2882
- (⨅ x ∈ (↑s : set α), f x) = ⨅ x ∈ s, f x :=
2883
- rfl
2884
-
2885
2909
end finset
2886
2910
2887
2911
namespace multiset
@@ -2967,34 +2991,92 @@ end finset
2967
2991
2968
2992
namespace finset
2969
2993
2970
- /-! ### bUnion -/
2994
+ /-! ### Interaction with big lattice/set operations -/
2995
+
2996
+ section lattice
2997
+
2998
+ lemma supr_coe [has_Sup β] (f : α → β) (s : finset α) :
2999
+ (⨆ x ∈ (↑s : set α), f x) = ⨆ x ∈ s, f x :=
3000
+ rfl
3001
+
3002
+ lemma infi_coe [has_Inf β] (f : α → β) (s : finset α) :
3003
+ (⨅ x ∈ (↑s : set α), f x) = ⨅ x ∈ s, f x :=
3004
+ rfl
3005
+
3006
+ variables [complete_lattice β]
3007
+
3008
+ theorem supr_singleton (a : α) (s : α → β) : (⨆ x ∈ ({a} : finset α), s x) = s a :=
3009
+ by simp
3010
+
3011
+ theorem infi_singleton (a : α) (s : α → β) : (⨅ x ∈ ({a} : finset α), s x) = s a :=
3012
+ by simp
3013
+
3014
+ variables [decidable_eq α]
3015
+
3016
+ theorem supr_union {f : α → β} {s t : finset α} :
3017
+ (⨆ x ∈ s ∪ t, f x) = (⨆x∈s, f x) ⊔ (⨆x∈t, f x) :=
3018
+ by simp [supr_or, supr_sup_eq]
3019
+
3020
+ theorem infi_union {f : α → β} {s t : finset α} :
3021
+ (⨅ x ∈ s ∪ t, f x) = (⨅ x ∈ s, f x) ⊓ (⨅ x ∈ t, f x) :=
3022
+ by simp [infi_or, infi_inf_eq]
3023
+
3024
+ lemma supr_insert (a : α) (s : finset α) (t : α → β) :
3025
+ (⨆ x ∈ insert a s, t x) = t a ⊔ (⨆ x ∈ s, t x) :=
3026
+ by { rw insert_eq, simp only [supr_union, finset.supr_singleton] }
3027
+
3028
+ lemma infi_insert (a : α) (s : finset α) (t : α → β) :
3029
+ (⨅ x ∈ insert a s, t x) = t a ⊓ (⨅ x ∈ s, t x) :=
3030
+ by { rw insert_eq, simp only [infi_union, finset.infi_singleton] }
3031
+
3032
+ lemma supr_finset_image {f : γ → α} {g : α → β} {s : finset γ} :
3033
+ (⨆ x ∈ s.image f, g x) = (⨆ y ∈ s, g (f y)) :=
3034
+ by rw [← supr_coe, coe_image, supr_image, supr_coe]
3035
+
3036
+ lemma infi_finset_image {f : γ → α} {g : α → β} {s : finset γ} :
3037
+ (⨅ x ∈ s.image f, g x) = (⨅ y ∈ s, g (f y)) :=
3038
+ by rw [← infi_coe, coe_image, infi_image, infi_coe]
3039
+
3040
+ end lattice
2971
3041
2972
3042
@[simp] theorem bUnion_coe (s : finset α) (t : α → set β) :
2973
3043
(⋃ x ∈ (↑s : set α), t x) = ⋃ x ∈ s, t x :=
2974
3044
rfl
2975
3045
3046
+ @[simp] theorem bInter_coe (s : finset α) (t : α → set β) :
3047
+ (⋂ x ∈ (↑s : set α), t x) = ⋂ x ∈ s, t x :=
3048
+ rfl
3049
+
2976
3050
@[simp] theorem bUnion_singleton (a : α) (s : α → set β) : (⋃ x ∈ ({a} : finset α), s x) = s a :=
2977
- by rw [← bUnion_coe, coe_singleton, set.bUnion_singleton]
3051
+ supr_singleton a s
2978
3052
2979
- @[simp] lemma bUnion_preimage_singleton (f : α → β) (s : finset β) :
2980
- (⋃ y ∈ s, f ⁻¹' {y}) = f ⁻¹' ↑s :=
2981
- set.bUnion_preimage_singleton f ↑s
3053
+ @[simp] theorem bInter_singleton (a : α) (s : α → set β) : (⋂ x ∈ ({a} : finset α), s x) = s a :=
3054
+ infi_singleton a s
2982
3055
2983
3056
variables [decidable_eq α]
2984
3057
2985
- theorem supr_union {α} [complete_lattice α] {β} [decidable_eq β] {f : β → α} {s t : finset β} :
2986
- (⨆ x ∈ s ∪ t, f x) = (⨆x∈s, f x) ⊔ (⨆x∈t, f x) :=
2987
- calc (⨆ x ∈ s ∪ t, f x) = (⨆ x, (⨆h : x∈s, f x) ⊔ (⨆h : x∈t, f x)) :
2988
- congr_arg _ $ funext $ λ x, by { convert supr_or, rw finset.mem_union, rw finset.mem_union, refl,
2989
- refl }
2990
- ... = (⨆x∈s, f x) ⊔ (⨆x∈t, f x) : supr_sup_eq
2991
-
2992
3058
lemma bUnion_union (s t : finset α) (u : α → set β) :
2993
3059
(⋃ x ∈ s ∪ t, u x) = (⋃ x ∈ s, u x) ∪ (⋃ x ∈ t, u x) :=
2994
3060
supr_union
2995
3061
3062
+ lemma bInter_inter (s t : finset α) (u : α → set β) :
3063
+ (⋂ x ∈ s ∪ t, u x) = (⋂ x ∈ s, u x) ∩ (⋂ x ∈ t, u x) :=
3064
+ infi_union
3065
+
2996
3066
@[simp] lemma bUnion_insert (a : α) (s : finset α) (t : α → set β) :
2997
3067
(⋃ x ∈ insert a s, t x) = t a ∪ (⋃ x ∈ s, t x) :=
2998
- begin rw insert_eq, simp only [bUnion_union, finset.bUnion_singleton] end
3068
+ supr_insert a s t
3069
+
3070
+ @[simp] lemma bInter_insert (a : α) (s : finset α) (t : α → set β) :
3071
+ (⋂ x ∈ insert a s, t x) = t a ∩ (⋂ x ∈ s, t x) :=
3072
+ infi_insert a s t
3073
+
3074
+ @[simp] lemma bUnion_finset_image {f : γ → α} {g : α → set β} {s : finset γ} :
3075
+ (⋃x ∈ s.image f, g x) = (⋃y ∈ s, g (f y)) :=
3076
+ supr_finset_image
3077
+
3078
+ @[simp] lemma bInter_finset_image {f : γ → α} {g : α → set β} {s : finset γ} :
3079
+ (⋂ x ∈ s.image f, g x) = (⋂ y ∈ s, g (f y)) :=
3080
+ infi_finset_image
2999
3081
3000
3082
end finset
0 commit comments