@@ -1683,18 +1683,14 @@ end image
1683
1683
1684
1684
/-! ### Subsingleton -/
1685
1685
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. -/
1687
1687
protected def subsingleton (s : set α) : Prop :=
1688
1688
∀ ⦃x⦄ (hx : x ∈ s) ⦃y⦄ (hy : y ∈ s), x = y
1689
1689
1690
- lemma subsingleton.mono (ht : t.subsingleton) (hst : s ⊆ t) : s.subsingleton :=
1690
+ lemma subsingleton.anti (ht : t.subsingleton) (hst : s ⊆ t) : s.subsingleton :=
1691
1691
λ x hx y hy, ht (hst hx) (hst hy)
1692
1692
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} :=
1698
1694
ext $ λ y, ⟨λ hy, (hs hx hy) ▸ mem_singleton _, λ hy, (eq_of_mem_singleton hy).symm ▸ hx⟩
1699
1695
1700
1696
@[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
1703
1699
λ x hx y hy, (eq_of_mem_singleton hx).symm ▸ (eq_of_mem_singleton hy).symm ▸ rfl
1704
1700
1705
1701
lemma subsingleton_of_subset_singleton (h : s ⊆ {a}) : s.subsingleton :=
1706
- subsingleton_singleton.mono h
1702
+ subsingleton_singleton.anti h
1707
1703
1708
1704
lemma subsingleton_of_forall_eq (a : α) (h : ∀ b ∈ s, b = a) : s.subsingleton :=
1709
1705
λ b hb c hc, (h _ hb).trans (h _ hc).symm
@@ -1729,7 +1725,7 @@ lemma subsingleton_of_univ_subsingleton (h : (univ : set α).subsingleton) : sub
1729
1725
⟨subsingleton_of_univ_subsingleton, λ h, @subsingleton_univ _ h⟩
1730
1726
1731
1727
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)
1733
1729
1734
1730
lemma subsingleton_is_top (α : Type *) [partial_order α] : set.subsingleton {x : α | is_top x} :=
1735
1731
λ x hx y hy, hx.is_max.eq_of_le (hy x)
@@ -1746,8 +1742,7 @@ begin
1746
1742
{ exact h.2 .eq_empty_or_singleton.resolve_left h.1 .ne_empty },
1747
1743
end
1748
1744
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. -/
1751
1746
@[simp, norm_cast] lemma subsingleton_coe (s : set α) : subsingleton s ↔ s.subsingleton :=
1752
1747
begin
1753
1748
split,
@@ -1761,16 +1756,24 @@ For the corresponding result for `subtype`, see `subtype.subsingleton`. -/
1761
1756
instance subsingleton_coe_of_subsingleton [subsingleton α] {s : set α} : subsingleton s :=
1762
1757
by { rw [s.subsingleton_coe], exact subsingleton_of_subsingleton }
1763
1758
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
+
1764
1763
/-- The preimage of a subsingleton under an injective map is a subsingleton. -/
1765
1764
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
1769
1766
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 . -/
1771
1768
theorem subsingleton_of_image {α β : Type *} {f : α → β} (hf : function.injective f)
1772
1769
(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) }
1774
1777
1775
1778
/-! ### Nontrivial -/
1776
1779
0 commit comments