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

Commit d9992ad

Browse files
feat(data/set/basic): Add natural missing lemmas to set.subsingleton and slightly refactor (#15886)
- `subsingleton.image` is moved and documentation is added so that it is near related lemmas. - `subsingleton_of_preimage` is added to go with `subsingleton.preimage`, `subsingleton_of_image`., and `subsingleton.image`. - We add `subsingleton_anti` analogously to `nontrivial_mono`. - Some small style tweaks are made. - `subsingleton.mono` is renamed to `subsingleton.anti`. Co-authored-by: Wrenna Robson <34025592+linesthatinterlace@users.noreply.github.com>
1 parent 1d7f479 commit d9992ad

File tree

6 files changed

+25
-22
lines changed

6 files changed

+25
-22
lines changed

src/data/set/basic.lean

Lines changed: 19 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -1683,18 +1683,14 @@ end image
16831683

16841684
/-! ### Subsingleton -/
16851685

1686-
/-- A set `s` is a `subsingleton`, if it has at most one element. -/
1686+
/-- A set `s` is a `subsingleton` if it has at most one element. -/
16871687
protected def subsingleton (s : set α) : Prop :=
16881688
∀ ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s), x = y
16891689

1690-
lemma subsingleton.mono (ht : t.subsingleton) (hst : s ⊆ t) : s.subsingleton :=
1690+
lemma subsingleton.anti (ht : t.subsingleton) (hst : s ⊆ t) : s.subsingleton :=
16911691
λ x hx y hy, ht (hst hx) (hst hy)
16921692

1693-
lemma subsingleton.image (hs : s.subsingleton) (f : α → β) : (f '' s).subsingleton :=
1694-
λ _ ⟨x, hx, Hx⟩ _ ⟨y, hy, Hy⟩, Hx ▸ Hy ▸ congr_arg f (hs hx hy)
1695-
1696-
lemma subsingleton.eq_singleton_of_mem (hs : s.subsingleton) {x:α} (hx : x ∈ s) :
1697-
s = {x} :=
1693+
lemma subsingleton.eq_singleton_of_mem (hs : s.subsingleton) {x:α} (hx : x ∈ s) : s = {x} :=
16981694
ext $ λ y, ⟨λ hy, (hs hx hy) ▸ mem_singleton _, λ hy, (eq_of_mem_singleton hy).symm ▸ hx⟩
16991695

17001696
@[simp] lemma subsingleton_empty : (∅ : set α).subsingleton := λ x, false.elim
@@ -1703,7 +1699,7 @@ ext $ λ y, ⟨λ hy, (hs hx hy) ▸ mem_singleton _, λ hy, (eq_of_mem_singleto
17031699
λ x hx y hy, (eq_of_mem_singleton hx).symm ▸ (eq_of_mem_singleton hy).symm ▸ rfl
17041700

17051701
lemma subsingleton_of_subset_singleton (h : s ⊆ {a}) : s.subsingleton :=
1706-
subsingleton_singleton.mono h
1702+
subsingleton_singleton.anti h
17071703

17081704
lemma subsingleton_of_forall_eq (a : α) (h : ∀ b ∈ s, b = a) : s.subsingleton :=
17091705
λ b hb c hc, (h _ hb).trans (h _ hc).symm
@@ -1729,7 +1725,7 @@ lemma subsingleton_of_univ_subsingleton (h : (univ : set α).subsingleton) : sub
17291725
⟨subsingleton_of_univ_subsingleton, λ h, @subsingleton_univ _ h⟩
17301726

17311727
lemma subsingleton_of_subsingleton [subsingleton α] {s : set α} : set.subsingleton s :=
1732-
subsingleton.mono subsingleton_univ (subset_univ s)
1728+
subsingleton_univ.anti (subset_univ s)
17331729

17341730
lemma subsingleton_is_top (α : Type*) [partial_order α] : set.subsingleton {x : α | is_top x} :=
17351731
λ x hx y hy, hx.is_max.eq_of_le (hy x)
@@ -1746,8 +1742,7 @@ begin
17461742
{ exact h.2.eq_empty_or_singleton.resolve_left h.1.ne_empty },
17471743
end
17481744

1749-
/-- `s`, coerced to a type, is a subsingleton type if and only if `s`
1750-
is a subsingleton set. -/
1745+
/-- `s`, coerced to a type, is a subsingleton type if and only if `s` is a subsingleton set. -/
17511746
@[simp, norm_cast] lemma subsingleton_coe (s : set α) : subsingleton s ↔ s.subsingleton :=
17521747
begin
17531748
split,
@@ -1761,16 +1756,24 @@ For the corresponding result for `subtype`, see `subtype.subsingleton`. -/
17611756
instance subsingleton_coe_of_subsingleton [subsingleton α] {s : set α} : subsingleton s :=
17621757
by { rw [s.subsingleton_coe], exact subsingleton_of_subsingleton }
17631758

1759+
/-- The image of a subsingleton is a subsingleton. -/
1760+
lemma subsingleton.image (hs : s.subsingleton) (f : α → β) : (f '' s).subsingleton :=
1761+
λ _ ⟨x, hx, Hx⟩ _ ⟨y, hy, Hy⟩, Hx ▸ Hy ▸ congr_arg f (hs hx hy)
1762+
17641763
/-- The preimage of a subsingleton under an injective map is a subsingleton. -/
17651764
theorem subsingleton.preimage {s : set β} (hs : s.subsingleton) {f : α → β}
1766-
(hf : function.injective f) :
1767-
(f ⁻¹' s).subsingleton :=
1768-
λ a ha b hb, hf $ hs ha hb
1765+
(hf : function.injective f) : (f ⁻¹' s).subsingleton := λ a ha b hb, hf $ hs ha hb
17691766

1770-
/-- `s` is a subsingleton, if its image of an injective function is. -/
1767+
/-- If the image of a set under an injective map is a subsingleton, the set is a subsingleton. -/
17711768
theorem subsingleton_of_image {α β : Type*} {f : α → β} (hf : function.injective f)
17721769
(s : set α) (hs : (f '' s).subsingleton) : s.subsingleton :=
1773-
(hs.preimage hf).mono $ subset_preimage_image _ _
1770+
(hs.preimage hf).anti $ subset_preimage_image _ _
1771+
1772+
/-- If the preimage of a set under an surjective map is a subsingleton,
1773+
the set is a subsingleton. -/
1774+
theorem subsingleton_of_preimage {α β : Type*} {f : α → β} (hf : function.surjective f)
1775+
(s : set β) (hs : (f ⁻¹' s).subsingleton) : s.subsingleton :=
1776+
λ fx hx fy hy, by { rcases ⟨hf fx, hf fy⟩ with ⟨⟨x, rfl⟩, ⟨y, rfl⟩⟩, exact congr_arg f (hs hx hy) }
17741777

17751778
/-! ### Nontrivial -/
17761779

src/data/set/pointwise.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1236,7 +1236,7 @@ lemma zero_smul_set_subset (s : set β) : (0 : α) • s ⊆ 0 :=
12361236
image_subset_iff.2 $ λ x _, zero_smul α x
12371237

12381238
lemma subsingleton_zero_smul_set (s : set β) : ((0 : α) • s).subsingleton :=
1239-
subsingleton_singleton.mono $ zero_smul_set_subset s
1239+
subsingleton_singleton.anti $ zero_smul_set_subset s
12401240

12411241
lemma zero_mem_smul_set {t : set β} {a : α} (h : (0 : β) ∈ t) : (0 : β) ∈ a • t :=
12421242
0, h, smul_zero' _ _⟩

src/measure_theory/measure/hausdorff.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -883,7 +883,7 @@ begin
883883
{ simp only [hs, measure_empty, zero_le], },
884884
have : f ⁻¹' s = {x},
885885
{ haveI : subsingleton X := hf.subsingleton,
886-
have : (f ⁻¹' s).subsingleton, from subsingleton_univ.mono (subset_univ _),
886+
have : (f ⁻¹' s).subsingleton, from subsingleton_univ.anti (subset_univ _),
887887
exact (subsingleton_iff_singleton hx).1 this },
888888
rw this,
889889
rcases eq_or_lt_of_le hd with rfl|h'd,

src/topology/connected.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -1246,7 +1246,7 @@ theorem is_totally_disconnected_empty : is_totally_disconnected (∅ : set α) :
12461246
λ _ ht _ _ x_in _ _, (ht x_in).elim
12471247

12481248
theorem is_totally_disconnected_singleton {x} : is_totally_disconnected ({x} : set α) :=
1249-
λ _ ht _, subsingleton.mono subsingleton_singleton ht
1249+
λ _ ht _, subsingleton_singleton.anti ht
12501250

12511251
/-- A space is totally disconnected if all of its connected components are singletons. -/
12521252
class totally_disconnected_space (α : Type u) [topological_space α] : Prop :=
@@ -1322,7 +1322,7 @@ begin
13221322
intros s s_sub hs,
13231323
rcases eq_empty_or_nonempty s with rfl | ⟨x, x_in⟩,
13241324
{ exact subsingleton_empty },
1325-
{ exact (h x).mono (hs.subset_connected_component x_in) }
1325+
{ exact (h x).anti (hs.subset_connected_component x_in) }
13261326
end
13271327

13281328
/-- A space is totally disconnected iff its connected components are singletons. -/

src/topology/metric_space/basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -2564,7 +2564,7 @@ begin
25642564
end
25652565

25662566
lemma subsingleton_sphere (x : γ) {r : ℝ} (hr : r ≤ 0) : (sphere x r).subsingleton :=
2567-
(subsingleton_closed_ball x hr).mono sphere_subset_closed_ball
2567+
(subsingleton_closed_ball x hr).anti sphere_subset_closed_ball
25682568

25692569
/-- A map between metric spaces is a uniform embedding if and only if the distance between `f x`
25702570
and `f y` is controlled in terms of the distance between `x` and `y` and conversely. -/

src/topology/separation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -516,7 +516,7 @@ hs.induction_on (by simp) $ λ x, by simp
516516

517517
@[simp] lemma subsingleton_closure [t1_space α] {s : set α} :
518518
(closure s).subsingleton ↔ s.subsingleton :=
519-
⟨λ h, h.mono subset_closure, λ h, h.closure⟩
519+
⟨λ h, h.anti subset_closure, λ h, h.closure⟩
520520

521521
lemma is_closed_map_const {α β} [topological_space α] [topological_space β] [t1_space β] {y : β} :
522522
is_closed_map (function.const α y) :=

0 commit comments

Comments
 (0)