Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

[Merged by Bors] - feat(topology/discrete_quotient): add two lemmas #8464

Closed
wants to merge 8 commits into from
Closed
Show file tree
Hide file tree
Changes from 7 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
19 changes: 19 additions & 0 deletions src/topology/discrete_quotient.lean
Original file line number Diff line number Diff line change
Expand Up @@ -30,6 +30,8 @@ The type `discrete_quotient X` is endowed with an instance of a `semilattice_inf
The partial ordering `A ≤ B` mathematically means that `B.proj` factors through `A.proj`.
The top element `⊤` is the trivial quotient, meaning that every element of `X` is collapsed
to a point. Given `h : A ≤ B`, the map `A → B` is `discrete_quotient.of_le h`.
Whenever `X` is discrete, the type `discrete_quotient X` is also endowed with an instance of a
`semilattice_inf_bot`, where the bot element `⊥` is `X` itself.

Given `f : X → Y` and `h : continuous f`, we define a predicate `le_comap h A B` for
`A : discrete_quotient X` and `B : discrete_quotient Y`, asserting that `f` descends to `A → B`.
Expand Down Expand Up @@ -198,6 +200,23 @@ lemma of_le_proj_apply {A B : discrete_quotient X} (h : A ≤ B) (x : X) :

end of_le

/--
When X is discrete, there is a `semilattice_inf_bot` instance on `discrete_quotient X`
-/
instance [discrete_topology X] : semilattice_inf_bot (discrete_quotient X) :=
{ bot :=
{ rel := (=),
equiv := eq_equivalence,
clopen := λ x, is_clopen_discrete _ },
bot_le := by { rintro S a b (h : a = b), rw h, exact S.refl _ },
..(infer_instance : semilattice_inf _) }
faenuccio marked this conversation as resolved.
Show resolved Hide resolved
faenuccio marked this conversation as resolved.
Show resolved Hide resolved

lemma proj_bot_injective [discrete_topology X] :
function.injective (⊥ : discrete_quotient X).proj := λ a b h, quotient.exact' h

lemma proj_bot_bijective [discrete_topology X] :
function.bijective (⊥ : discrete_quotient X).proj := ⟨proj_bot_injective, proj_surjective _⟩

section map

variables {Y : Type*} [topological_space Y] {f : Y → X}
Expand Down
3 changes: 3 additions & 0 deletions src/topology/subset_properties.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1173,6 +1173,9 @@ begin
exact ⟨hx₁, by simpa [not_mem_of_mem_compl hx₂] using cover hx₁⟩ }
end

@[simp] lemma is_clopen_discrete [discrete_topology α] (x : set α) : is_clopen x :=
⟨is_open_discrete _, is_closed_discrete _⟩
faenuccio marked this conversation as resolved.
Show resolved Hide resolved

end clopen

section preirreducible
Expand Down