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

Commit e7e9f30

Browse files
committed
feat(set): preliminaries for Haar measure (#3189)
`comp_sup_eq_sup_comp` is renamed `comp_sup_eq_sup_comp_of_is_total` and there is a new version that doesn't assume that the order is linear. `set.image_injective` is renamed `function.injective.image_injective` (in the same way as the existing `function.surjective.preimage_injective`). `set.image_injective` is now an `iff`.
1 parent 8413b3f commit e7e9f30

File tree

6 files changed

+212
-47
lines changed

6 files changed

+212
-47
lines changed

src/data/finset.lean

Lines changed: 108 additions & 26 deletions
Original file line numberDiff line numberDiff line change
@@ -319,6 +319,21 @@ protected theorem induction_on {α : Type*} {p : finset α → Prop} [decidable_
319319
(s : finset α) (h₁ : p ∅) (h₂ : ∀ ⦃a : α⦄ {s : finset α}, a ∉ s → p s → p (insert a s)) : p s :=
320320
finset.induction h₁ h₂ s
321321

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+
322337
/-! ### union -/
323338

324339
/-- `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 α) :
15421557
(s ∪ t).card ≤ s.card + t.card :=
15431558
card_union_add_card_inter s t ▸ le_add_right _ _
15441559

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+
15451567
lemma surj_on_of_inj_on_of_card_le {s : finset α} {t : finset β}
15461568
(f : Π a ∈ s, β) (hf : ∀ a ha, f a ha ∈ t)
15471569
(hinj : ∀ a₁ a₂ ha₁ ha₂, f a₁ ha₁ = f a₂ ha₂ → a₁ = a₂)
@@ -2064,10 +2086,15 @@ by letI := classical.dec_eq β; from
20642086
⟨ λh b hb, lt_of_le_of_lt (le_sup hb) h,
20652087
finset.induction_on s (by simp [ha]) (by simp {contextual := tt}) ⟩
20662088

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) :=
20692092
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
20712098

20722099
theorem subset_range_sup_succ (s : finset ℕ) : s ⊆ range (s.sup id).succ :=
20732100
λ 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)
21262153
lemma lt_inf_iff [h : is_total α (≤)] {a : α} (ha : a < ⊤) : a < s.inf f ↔ (∀b ∈ s, a < f b) :=
21272154
@sup_lt_iff (order_dual α) _ _ _ _ (@is_total.swap α _ h) _ ha
21282155

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 γ]
21302162
(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
21322164

21332165
end inf
21342166

@@ -2874,14 +2906,6 @@ end
28742906

28752907
@[simp] lemma Ico_ℤ.card (l u : ℤ) : (Ico_ℤ l u).card = (u - l).to_nat := by simp [Ico_ℤ]
28762908

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-
28852909
end finset
28862910

28872911
namespace multiset
@@ -2967,34 +2991,92 @@ end finset
29672991

29682992
namespace finset
29692993

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
29713041

29723042
@[simp] theorem bUnion_coe (s : finset α) (t : α → set β) :
29733043
(⋃ x ∈ (↑s : set α), t x) = ⋃ x ∈ s, t x :=
29743044
rfl
29753045

3046+
@[simp] theorem bInter_coe (s : finset α) (t : α → set β) :
3047+
(⋂ x ∈ (↑s : set α), t x) = ⋂ x ∈ s, t x :=
3048+
rfl
3049+
29763050
@[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
29783052

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
29823055

29833056
variables [decidable_eq α]
29843057

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-
29923058
lemma bUnion_union (s t : finset α) (u : α → set β) :
29933059
(⋃ x ∈ s ∪ t, u x) = (⋃ x ∈ s, u x) ∪ (⋃ x ∈ t, u x) :=
29943060
supr_union
29953061

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+
29963066
@[simp] lemma bUnion_insert (a : α) (s : finset α) (t : α → set β) :
29973067
(⋃ 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
29993081

30003082
end finset

src/data/set/basic.lean

Lines changed: 77 additions & 18 deletions
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ import tactic.finish
88
import data.subtype
99
import logic.unique
1010
import data.prod
11+
import logic.function.basic
1112

1213
/-!
1314
# Basic properties of sets
@@ -75,6 +76,8 @@ singleton, complement, powerset
7576

7677
open function
7778

79+
universe variables u v w x
80+
7881
namespace set
7982

8083
/-- Coercion from a set to the corresponding subtype. -/
@@ -84,8 +87,6 @@ end set
8487

8588
section set_coe
8689

87-
universe u
88-
8990
variables {α : Type u}
9091

9192
theorem set.set_coe_eq_subtype (s : set α) :
@@ -119,8 +120,6 @@ lemma subtype.mem {α : Type*} {s : set α} (p : s) : (p : α) ∈ s := p.proper
119120

120121
namespace set
121122

122-
universes u v w x
123-
124123
variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {a : α} {s t : set α}
125124

126125
instance : inhabited (set α) := ⟨∅⟩
@@ -161,6 +160,7 @@ instance decidable_set_of (p : α → Prop) [H : decidable_pred p] : decidable_p
161160
theorem subset_def {s t : set α} : (s ⊆ t) = ∀ x, x ∈ s → x ∈ t := rfl
162161

163162
@[refl] theorem subset.refl (a : set α) : a ⊆ a := assume x, id
163+
theorem subset.rfl {s : set α} : s ⊆ s := subset.refl s
164164

165165
@[trans] theorem subset.trans {a b c : set α} (ab : a ⊆ b) (bc : b ⊆ c) : a ⊆ c :=
166166
assume x h, bc (ab h)
@@ -897,6 +897,10 @@ by rw [diff_subset_iff, diff_subset_iff, union_comm]
897897
lemma diff_inter {s t u : set α} : s \ (t ∩ u) = (s \ t) ∪ (s \ u) :=
898898
ext $ λ x, by simp [classical.not_and_distrib, and_or_distrib_left]
899899

900+
lemma diff_inter_diff {s t u : set α} : s \ t ∩ (s \ u) = s \ (t ∪ u) :=
901+
by { ext x, simp only [mem_inter_eq, mem_union_eq, mem_diff, not_or_distrib],
902+
exact ⟨λ ⟨⟨h1, h2⟩, _, h3⟩, ⟨h1, h2, h3⟩, λ ⟨h1, h2, h3⟩, ⟨⟨h1, h2⟩, h1, h3⟩⟩ }
903+
900904
lemma diff_compl : s \ -t = s ∩ t := by rw [diff_eq, compl_compl]
901905

902906
lemma diff_diff_right {s t u : set α} : s \ (t \ u) = (s \ t) ∪ (s ∩ u) :=
@@ -993,6 +997,8 @@ rfl
993997

994998
@[simp] theorem preimage_id {s : set α} : id ⁻¹' s = s := rfl
995999

1000+
@[simp] theorem preimage_id' {s : set α} : (λ x, x) ⁻¹' s = s := rfl
1001+
9961002
theorem preimage_const_of_mem {b : β} {s : set β} (h : b ∈ s) :
9971003
(λ (x : α), b) ⁻¹' s = univ :=
9981004
eq_univ_of_forall $ λ x, h
@@ -1007,6 +1013,10 @@ by { split_ifs with hb hb, exacts [preimage_const_of_mem hb, preimage_const_of_n
10071013

10081014
theorem preimage_comp {s : set γ} : (g ∘ f) ⁻¹' s = f ⁻¹' (g ⁻¹' s) := rfl
10091015

1016+
lemma preimage_preimage {g : β → γ} {f : α → β} {s : set γ} :
1017+
f ⁻¹' (g ⁻¹' s) = (λ x, g (f x)) ⁻¹' s :=
1018+
preimage_comp.symm
1019+
10101020
theorem eq_preimage_subtype_val_iff {p : α → Prop} {s : set (subtype p)} {t : set α} :
10111021
s = subtype.val ⁻¹' t ↔ (∀x (h : p x), (⟨x, h⟩ : subtype p) ∈ s ↔ x ∈ t) :=
10121022
⟨assume s_eq x h, by rw [s_eq]; simp,
@@ -1299,9 +1309,6 @@ begin
12991309
exact preimage_mono h
13001310
end
13011311

1302-
lemma image_injective {f : α → β} (hf : injective f) : injective (('') f) :=
1303-
assume s t, (image_eq_image hf).1
1304-
13051312
lemma prod_quotient_preimage_eq_image [s : setoid α] (g : quotient s → β) {h : α → β}
13061313
(Hh : h = g ∘ quotient.mk) (r : set (β × β)) :
13071314
{x : quotient s × quotient s | (g x.1, g x.2) ∈ r} =
@@ -1529,6 +1536,36 @@ theorem pairwise_on.mono' {s : set α} {r r' : α → α → Prop}
15291536
end set
15301537
open set
15311538

1539+
namespace function
1540+
1541+
variables {ι : Sort*} {α : Type*} {β : Type*} {f : α → β}
1542+
1543+
lemma surjective.preimage_injective (hf : surjective f) : injective (preimage f) :=
1544+
assume s t, (preimage_eq_preimage hf).1
1545+
1546+
lemma injective.preimage_surjective (hf : injective f) : surjective (preimage f) :=
1547+
by { intro s, use f '' s, rw preimage_image_eq _ hf }
1548+
1549+
lemma surjective.image_surjective (hf : surjective f) : surjective (image f) :=
1550+
by { intro s, use f ⁻¹' s, rw image_preimage_eq hf }
1551+
1552+
lemma injective.image_injective (hf : injective f) : injective (image f) :=
1553+
by { intros s t h, rw [←preimage_image_eq s hf, ←preimage_image_eq t hf, h] }
1554+
1555+
lemma surjective.range_eq {f : ι → α} (hf : surjective f) : range f = univ :=
1556+
range_iff_surjective.2 hf
1557+
1558+
lemma surjective.range_comp (g : α → β) {f : ι → α} (hf : surjective f) :
1559+
range (g ∘ f) = range g :=
1560+
by rw [range_comp, hf.range_eq, image_univ]
1561+
1562+
lemma injective.nonempty_apply_iff {f : set α → set β} (hf : injective f)
1563+
(h2 : f ∅ = ∅) {s : set α} : (f s).nonempty ↔ s.nonempty :=
1564+
by rw [← ne_empty_iff_nonempty, ← h2, ← ne_empty_iff_nonempty, hf.ne_iff]
1565+
1566+
end function
1567+
open function
1568+
15321569
/-! ### Image and preimage on subtypes -/
15331570

15341571
namespace subtype
@@ -1567,7 +1604,7 @@ theorem preimage_val_eq_preimage_val_iff (s t u : set α) :
15671604
begin
15681605
rw [←image_preimage_val, ←image_preimage_val],
15691606
split, { intro h, rw h },
1570-
intro h, exact set.image_injective (val_injective) h
1607+
intro h, exact val_injective.image_injective h
15711608
end
15721609

15731610
lemma exists_set_subtype {t : set α} (p : set α → Prop) :
@@ -1833,18 +1870,40 @@ s.eq_empty_or_nonempty.elim (λ h, h.symm ▸ h0) $ λ h, (eq_univ_of_nonempty h
18331870

18341871
end subsingleton
18351872

1836-
namespace function
1873+
namespace set
18371874

1838-
variables {ι : Sort*} {α : Type*} {β : Type*}
1875+
variables {α : Type u} {β : Type v} {f : α → β}
18391876

1840-
lemma surjective.preimage_injective {f : β → α} (hf : surjective f) : injective (preimage f) :=
1841-
assume s t, (preimage_eq_preimage hf).1
1877+
@[simp]
1878+
lemma preimage_injective : injective (preimage f) ↔ surjective f :=
1879+
begin
1880+
refine ⟨λ h y, _, surjective.preimage_injective⟩,
1881+
obtain ⟨x, hx⟩ : (f ⁻¹' {y}).nonempty,
1882+
{ rw [h.nonempty_apply_iff preimage_empty], apply singleton_nonempty },
1883+
exact ⟨x, hx⟩
1884+
end
18421885

1843-
lemma surjective.range_eq {f : ι → α} (hf : surjective f) : range f = univ :=
1844-
range_iff_surjective.2 hf
1886+
@[simp]
1887+
lemma preimage_surjective : surjective (preimage f) ↔ injective f :=
1888+
begin
1889+
refine ⟨λ h x x' hx, _, injective.preimage_surjective⟩,
1890+
cases h {x} with s hs, have := mem_singleton x,
1891+
rwa [← hs, mem_preimage, hx, ← mem_preimage, hs, mem_singleton_iff, eq_comm] at this
1892+
end
18451893

1846-
lemma surjective.range_comp (g : α → β) {f : ι → α} (hf : surjective f) :
1847-
range (g ∘ f) = range g :=
1848-
by rw [range_comp, hf.range_eq, image_univ]
1894+
@[simp] lemma image_surjective : surjective (image f) ↔ surjective f :=
1895+
begin
1896+
refine ⟨λ h y, _, surjective.image_surjective⟩,
1897+
cases h {y} with s hs,
1898+
have := mem_singleton y, rw [← hs] at this, rcases this with ⟨x, h1x, h2x⟩,
1899+
exact ⟨x, h2x⟩
1900+
end
18491901

1850-
end function
1902+
@[simp] lemma image_injective : injective (image f) ↔ injective f :=
1903+
begin
1904+
refine ⟨λ h x x' hx, _, injective.image_injective⟩,
1905+
rw [← singleton_eq_singleton_iff], apply h,
1906+
rw [image_singleton, image_singleton, hx]
1907+
end
1908+
1909+
end set

0 commit comments

Comments
 (0)