Skip to content

Commit

Permalink
chore(topology/uniform_space/basic): rename symmetric_rel_inter to …
Browse files Browse the repository at this point in the history
…`symmetric_rel.inter` (#15441)

Also add `symmetric_rel.eq`
  • Loading branch information
urkud committed Jul 19, 2022
1 parent 65a44e6 commit 9137502
Showing 1 changed file with 5 additions and 6 deletions.
11 changes: 5 additions & 6 deletions src/topology/uniform_space/basic.lean
Expand Up @@ -195,12 +195,11 @@ lemma symmetric_rel.mk_mem_comm {V : set (α × α)} (hV : symmetric_rel V) {x y
(x, y) ∈ V ↔ (y, x) ∈ V :=
set.ext_iff.1 hV (y, x)

lemma symmetric_rel_inter {U V : set (α × α)} (hU : symmetric_rel U) (hV : symmetric_rel V) :
lemma symmetric_rel.eq {U : set (α × α)} (hU : symmetric_rel U) : prod.swap ⁻¹' U = U := hU

lemma symmetric_rel.inter {U V : set (α × α)} (hU : symmetric_rel U) (hV : symmetric_rel V) :
symmetric_rel (U ∩ V) :=
begin
unfold symmetric_rel at *,
rw [preimage_inter, hU, hV],
end
by rw [symmetric_rel, preimage_inter, hU.eq, hV.eq]

/-- This core description of a uniform space is outside of the type class hierarchy. It is useful
for constructions of uniform spaces, when the topology is derived from the uniform space. -/
Expand Down Expand Up @@ -658,7 +657,7 @@ begin
rw nhds_prod_eq,
apply (has_basis_nhds x).prod' (has_basis_nhds y),
rintro U V ⟨U_in, U_symm⟩ ⟨V_in, V_symm⟩,
exact ⟨U ∩ V, ⟨(𝓤 α).inter_sets U_in V_in, symmetric_rel_inter U_symm V_symm⟩,
exact ⟨U ∩ V, ⟨(𝓤 α).inter_sets U_in V_in, U_symm.inter V_symm⟩,
ball_inter_left x U V, ball_inter_right y U V⟩,
end

Expand Down

0 comments on commit 9137502

Please sign in to comment.