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

Commit f66a5dd

Browse files
committed
chore(data/set/basic): add a few lemmas and a @[simp] attribute (#12176)
* rename `set.exists_eq_singleton_iff_nonempty_unique_mem` to `set.exists_eq_singleton_iff_nonempty_subsingleton`, use `set.subsingleton` in the statement; * add `@[simp]` to `set.subset_compl_singleton_iff`; * add `set.diff_diff_right_self`.
1 parent 0eb5e2d commit f66a5dd

File tree

4 files changed

+17
-15
lines changed

4 files changed

+17
-15
lines changed

src/data/set/basic.lean

Lines changed: 13 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -773,16 +773,6 @@ subset.antisymm_iff.trans $ and.comm.trans $ and_congr_left' singleton_subset_if
773773
lemma eq_singleton_iff_nonempty_unique_mem : s = {a} ↔ s.nonempty ∧ ∀ x ∈ s, x = a :=
774774
eq_singleton_iff_unique_mem.trans $ and_congr_left $ λ H, ⟨λ h', ⟨_, h'⟩, λ ⟨x, h⟩, H x h ▸ h⟩
775775

776-
lemma exists_eq_singleton_iff_nonempty_unique_mem :
777-
(∃ a : α, s = {a}) ↔ (s.nonempty ∧ ∀ a b ∈ s, a = b) :=
778-
begin
779-
refine ⟨_, λ h, _⟩,
780-
{ rintros ⟨a, rfl⟩,
781-
refine ⟨set.singleton_nonempty a, λ b hb c hc, hb.trans hc.symm⟩ },
782-
{ obtain ⟨a, ha⟩ := h.1,
783-
refine ⟨a, set.eq_singleton_iff_unique_mem.mpr ⟨ha, λ b hb, (h.2 b hb a ha)⟩⟩ },
784-
end
785-
786776
-- while `simp` is capable of proving this, it is not capable of turning the LHS into the RHS.
787777
@[simp] lemma default_coe_singleton (x : α) : (default : ({x} : set α)) = ⟨x, rfl⟩ := rfl
788778

@@ -921,7 +911,7 @@ forall_congr $ λ a, imp_not_comm
921911
theorem subset_compl_iff_disjoint {s t : set α} : s ⊆ tᶜ ↔ s ∩ t = ∅ :=
922912
iff.trans (forall_congr $ λ a, and_imp.symm) subset_empty_iff
923913

924-
lemma subset_compl_singleton_iff {a : α} {s : set α} : s ⊆ {a}ᶜ ↔ a ∉ s :=
914+
@[simp] lemma subset_compl_singleton_iff {a : α} {s : set α} : s ⊆ {a}ᶜ ↔ a ∉ s :=
925915
subset_compl_comm.trans singleton_subset_iff
926916

927917
theorem inter_subset (a b c : set α) : a ∩ b ⊆ c ↔ a ⊆ bᶜ ∪ c :=
@@ -1109,6 +1099,8 @@ by simp [insert_eq, union_diff_self, -union_singleton, -singleton_union]
11091099

11101100
@[simp] lemma diff_self {s : set α} : s \ s = ∅ := sdiff_self
11111101

1102+
lemma diff_diff_right_self (s t : set α) : s \ (s \ t) = s ∩ t := sdiff_sdiff_right_self
1103+
11121104
lemma diff_diff_cancel_left {s t : set α} (h : s ⊆ t) : t \ (t \ s) = s :=
11131105
sdiff_sdiff_eq_self h
11141106

@@ -1637,6 +1629,16 @@ lemma subsingleton_is_top (α : Type*) [partial_order α] : set.subsingleton {x
16371629
lemma subsingleton_is_bot (α : Type*) [partial_order α] : set.subsingleton {x : α | is_bot x} :=
16381630
λ x hx y hy, hx.is_min.eq_of_ge (hy x)
16391631

1632+
lemma exists_eq_singleton_iff_nonempty_subsingleton :
1633+
(∃ a : α, s = {a}) ↔ s.nonempty ∧ s.subsingleton :=
1634+
begin
1635+
refine ⟨_, λ h, _⟩,
1636+
{ rintros ⟨a, rfl⟩,
1637+
exact ⟨singleton_nonempty a, subsingleton_singleton⟩ },
1638+
{ obtain ⟨a, ha⟩ := h.1,
1639+
exact ⟨a, eq_singleton_iff_unique_mem.mpr ⟨ha, λ b hb, h.2 hb ha⟩⟩ },
1640+
end
1641+
16401642
/-- `s`, coerced to a type, is a subsingleton type if and only if `s`
16411643
is a subsingleton set. -/
16421644
@[simp, norm_cast] lemma subsingleton_coe (s : set α) : subsingleton s ↔ s.subsingleton :=

src/group_theory/complement.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -100,7 +100,7 @@ end
100100

101101
@[to_additive] lemma is_complement_top_left : is_complement ⊤ S ↔ ∃ g : G, S = {g} :=
102102
begin
103-
refine ⟨λ h, set.exists_eq_singleton_iff_nonempty_unique_mem.mpr ⟨_, λ a ha b hb, _⟩, _⟩,
103+
refine ⟨λ h, set.exists_eq_singleton_iff_nonempty_subsingleton.mpr ⟨_, λ a ha b hb, _⟩, _⟩,
104104
{ obtain ⟨a, ha⟩ := h.2 1,
105105
exact ⟨a.2.1, a.2.2⟩ },
106106
{ have : (⟨⟨_, mem_top a⁻¹⟩, ⟨a, ha⟩⟩ : (⊤ : set G) × S) = ⟨⟨_, mem_top b⁻¹⟩, ⟨b, hb⟩⟩ :=
@@ -112,7 +112,7 @@ end
112112

113113
@[to_additive] lemma is_complement_top_right : is_complement S ⊤ ↔ ∃ g : G, S = {g} :=
114114
begin
115-
refine ⟨λ h, set.exists_eq_singleton_iff_nonempty_unique_mem.mpr ⟨_, λ a ha b hb, _⟩, _⟩,
115+
refine ⟨λ h, set.exists_eq_singleton_iff_nonempty_subsingleton.mpr ⟨_, λ a ha b hb, _⟩, _⟩,
116116
{ obtain ⟨a, ha⟩ := h.2 1,
117117
exact ⟨a.1.1, a.1.2⟩ },
118118
{ have : (⟨⟨a, ha⟩, ⟨_, mem_top a⁻¹⟩⟩ : S × (⊤ : set G)) = ⟨⟨b, hb⟩, ⟨_, mem_top b⁻¹⟩⟩ :=

src/topology/separation.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1169,7 +1169,7 @@ lemma is_irreducible_iff_singleton [t2_space α] (S : set α) :
11691169
begin
11701170
split,
11711171
{ intro h,
1172-
rw exists_eq_singleton_iff_nonempty_unique_mem,
1172+
rw exists_eq_singleton_iff_nonempty_subsingleton,
11731173
use h.1,
11741174
intros a ha b hb,
11751175
injection @@subsingleton.elim ((is_preirreducible_iff_subsingleton _).mp h.2) ⟨_, ha⟩ ⟨_, hb⟩ },

src/topology/subset_properties.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1405,7 +1405,7 @@ lemma is_preirreducible_of_subsingleton (s : set α) [hs : subsingleton s] : is_
14051405
begin
14061406
cases s.eq_empty_or_nonempty,
14071407
{ exact h.symm ▸ is_preirreducible_empty },
1408-
{ obtain ⟨x, e⟩ := exists_eq_singleton_iff_nonempty_unique_mem.mpr
1408+
{ obtain ⟨x, e⟩ := exists_eq_singleton_iff_nonempty_subsingleton.mpr
14091409
⟨h, λ _ ha _ hb, by injection @@subsingleton.elim hs ⟨_, ha⟩ ⟨_, hb⟩⟩,
14101410
exact e.symm ▸ is_irreducible_singleton.2 }
14111411
end

0 commit comments

Comments
 (0)