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] - refactor: prefer s ∩ . when passing to a subset of s #10433

Closed
wants to merge 3 commits into from
Closed
Show file tree
Hide file tree
Changes from 1 commit
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Jump to
Jump to file
Failed to load files.
Diff view
Diff view
2 changes: 1 addition & 1 deletion Mathlib/Analysis/InnerProductSpace/Projection.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1444,7 +1444,7 @@ theorem maximal_orthonormal_iff_orthogonalComplement_eq_bot (hv : Orthonormal
obtain ⟨l, hl, rfl⟩ :
∃ l ∈ Finsupp.supported 𝕜 𝕜 ((↑) ⁻¹' v : Set u), (Finsupp.total (↥u) E 𝕜 (↑)) l = y := by
rw [← Finsupp.mem_span_image_iff_total]
simp [huv, inter_eq_self_of_subset_left, hy]
simp [huv, inter_eq_self_of_subset_right, hy]
exact hu.inner_finsupp_eq_zero hxv' hl
#align maximal_orthonormal_iff_orthogonal_complement_eq_bot maximal_orthonormal_iff_orthogonalComplement_eq_bot

Expand Down
27 changes: 15 additions & 12 deletions Mathlib/Data/Set/Image.lean
Original file line number Diff line number Diff line change
Expand Up @@ -766,14 +766,17 @@ theorem insert_image_compl_eq_range (f : α → β) (x : α) : insert (f x) (f '
rw [← image_insert_eq, insert_eq, union_compl_self, image_univ]
#align set.insert_image_compl_eq_range Set.insert_image_compl_eq_range

theorem image_preimage_eq_inter_range {f : α → β} {t : Set β} : f '' (f ⁻¹' t) = t ∩ range f :=
theorem image_preimage_eq_range_inter {f : α → β} {t : Set β} : f '' (f ⁻¹' t) = range f ∩ t :=
ext fun x =>
⟨fun ⟨x, hx, HEq⟩ => HEq ▸ ⟨hx, mem_range_self _⟩, fun ⟨hx, ⟨y, h_eq⟩⟩ =>
⟨fun ⟨x, hx, HEq⟩ => HEq ▸ ⟨mem_range_self _, hx⟩, fun ⟨⟨y, h_eq⟩, hx⟩ =>
h_eq ▸ mem_image_of_mem f <| show y ∈ f ⁻¹' t by rw [preimage, mem_setOf, h_eq]; exact hx⟩

theorem image_preimage_eq_inter_range {f : α → β} {t : Set β} : f '' (f ⁻¹' t) = t ∩ range f := by
rw [image_preimage_eq_range_inter, inter_comm]
#align set.image_preimage_eq_inter_range Set.image_preimage_eq_inter_range

theorem image_preimage_eq_of_subset {f : α → β} {s : Set β} (hs : s ⊆ range f) :
f '' (f ⁻¹' s) = s := by rw [image_preimage_eq_inter_range, inter_eq_self_of_subset_left hs]
f '' (f ⁻¹' s) = s := by rw [image_preimage_eq_range_inter, inter_eq_self_of_subset_right hs]
#align set.image_preimage_eq_of_subset Set.image_preimage_eq_of_subset

theorem image_preimage_eq_iff {f : α → β} {s : Set β} : f '' (f ⁻¹' s) = s ↔ s ⊆ range f :=
Expand Down Expand Up @@ -838,7 +841,7 @@ theorem preimage_range_inter {f : α → β} {s : Set β} : f ⁻¹' (range f
#align set.preimage_range_inter Set.preimage_range_inter

theorem preimage_image_preimage {f : α → β} {s : Set β} : f ⁻¹' (f '' (f ⁻¹' s)) = f ⁻¹' s := by
rw [image_preimage_eq_inter_range, preimage_inter_range]
rw [image_preimage_eq_range_inter, preimage_range_inter]
#align set.preimage_image_preimage Set.preimage_image_preimage

@[simp, mfld_simps]
Expand Down Expand Up @@ -1398,28 +1401,28 @@ theorem coe_image_univ (s : Set α) : ((↑) : s → α) '' Set.univ = s :=
#align subtype.coe_image_univ Subtype.coe_image_univ

@[simp]
theorem image_preimage_coe (s t : Set α) : ((↑) : s → α) '' (((↑) : s → α) ⁻¹' t) = ts :=
image_preimage_eq_inter_range.trans <| congr_arg _ range_coe
theorem image_preimage_coe (s t : Set α) : ((↑) : s → α) '' (((↑) : s → α) ⁻¹' t) = st :=
image_preimage_eq_range_inter.trans <| congr_arg (· ∩ t) range_coe
#align subtype.image_preimage_coe Subtype.image_preimage_coe

theorem image_preimage_val (s t : Set α) : (Subtype.val : s → α) '' (Subtype.val ⁻¹' t) = ts :=
theorem image_preimage_val (s t : Set α) : (Subtype.val : s → α) '' (Subtype.val ⁻¹' t) = st :=
image_preimage_coe s t
#align subtype.image_preimage_val Subtype.image_preimage_val

theorem preimage_coe_eq_preimage_coe_iff {s t u : Set α} :
((↑) : s → α) ⁻¹' t = ((↑) : s → α) ⁻¹' u ↔ ts = us := by
((↑) : s → α) ⁻¹' t = ((↑) : s → α) ⁻¹' u ↔ st = su := by
rw [← image_preimage_coe, ← image_preimage_coe, coe_injective.image_injective.eq_iff]
#align subtype.preimage_coe_eq_preimage_coe_iff Subtype.preimage_coe_eq_preimage_coe_iff

-- Porting note:
-- @[simp] `simp` can prove this
theorem preimage_coe_inter_self (s t : Set α) :
eric-wieser marked this conversation as resolved.
Show resolved Hide resolved
((↑) : s → α) ⁻¹' (ts) = ((↑) : s → α) ⁻¹' t := by
rw [preimage_coe_eq_preimage_coe_iff, inter_assoc, inter_self]
((↑) : s → α) ⁻¹' (st) = ((↑) : s → α) ⁻¹' t := by
rw [preimage_coe_eq_preimage_coe_iff, inter_assoc, inter_self]
#align subtype.preimage_coe_inter_self Subtype.preimage_coe_inter_self

theorem preimage_val_eq_preimage_val_iff (s t u : Set α) :
(Subtype.val : s → α) ⁻¹' t = Subtype.val ⁻¹' u ↔ ts = us :=
(Subtype.val : s → α) ⁻¹' t = Subtype.val ⁻¹' u ↔ st = su :=
preimage_coe_eq_preimage_coe_iff
#align subtype.preimage_val_eq_preimage_val_iff Subtype.preimage_val_eq_preimage_val_iff

Expand All @@ -1434,7 +1437,7 @@ theorem forall_set_subtype {t : Set α} (p : Set α → Prop) :

theorem preimage_coe_nonempty {s t : Set α} :
(((↑) : s → α) ⁻¹' t).Nonempty ↔ (s ∩ t).Nonempty := by
rw [inter_comm, ← image_preimage_coe, image_nonempty]
rw [← image_preimage_coe, image_nonempty]
#align subtype.preimage_coe_nonempty Subtype.preimage_coe_nonempty

theorem preimage_coe_eq_empty {s t : Set α} : ((↑) : s → α) ⁻¹' t = ∅ ↔ s ∩ t = ∅ := by
Expand Down
28 changes: 14 additions & 14 deletions Mathlib/Data/Set/Intervals/Image.lean
Original file line number Diff line number Diff line change
Expand Up @@ -289,27 +289,27 @@ theorem image_subtype_val_Ioi_subset (a : {x // p x}) :

@[simp]
lemma image_subtype_val_Ici_Iic {a : α} (b : Ici a) : Subtype.val '' Iic b = Icc a b :=
(Subtype.image_preimage_val (Ici a) (Iic b.1)).trans Iic_inter_Ici
(Subtype.image_preimage_val (Ici a) (Iic b.1)).trans Ici_inter_Iic

@[simp]
lemma image_subtype_val_Ici_Iio {a : α} (b : Ici a) : Subtype.val '' Iio b = Ico a b :=
(Subtype.image_preimage_val (Ici a) (Iio b.1)).trans Iio_inter_Ici
(Subtype.image_preimage_val (Ici a) (Iio b.1)).trans Ici_inter_Iio

@[simp]
lemma image_subtype_val_Ici_Ici {a : α} (b : Ici a) : Subtype.val '' Ici b = Ici b.1 :=
(Subtype.image_preimage_val (Ici a) (Ici b.1)).trans <| inter_eq_left.2 <| Ici_subset_Ici.2 b.2
(Subtype.image_preimage_val (Ici a) (Ici b.1)).trans <| inter_eq_right.2 <| Ici_subset_Ici.2 b.2

@[simp]
lemma image_subtype_val_Ici_Ioi {a : α} (b : Ici a) : Subtype.val '' Ioi b = Ioi b.1 :=
(Subtype.image_preimage_val (Ici a) (Ioi b.1)).trans <| inter_eq_left.2 <| Ioi_subset_Ici b.2
(Subtype.image_preimage_val (Ici a) (Ioi b.1)).trans <| inter_eq_right.2 <| Ioi_subset_Ici b.2

@[simp]
lemma image_subtype_val_Iic_Ici {a : α} (b : Iic a) : Subtype.val '' Ici b = Icc b.1 a :=
Subtype.image_preimage_val (Iic a) (Ici b.1)
(Subtype.image_preimage_val _ _).trans <| inter_comm _ _

@[simp]
lemma image_subtype_val_Iic_Ioi {a : α} (b : Iic a) : Subtype.val '' Ioi b = Ioc b.1 a :=
Subtype.image_preimage_val (Iic a) (Ioi b.1)
(Subtype.image_preimage_val _ _).trans <| inter_comm _ _

@[simp]
lemma image_subtype_val_Iic_Iic {a : α} (b : Iic a) : Subtype.val '' Iic b = Iic b.1 :=
Expand All @@ -321,31 +321,31 @@ lemma image_subtype_val_Iic_Iio {a : α} (b : Iic a) : Subtype.val '' Iio b = Ii

@[simp]
lemma image_subtype_val_Ioi_Ici {a : α} (b : Ioi a) : Subtype.val '' Ici b = Ici b.1 :=
(Subtype.image_preimage_val (Ioi a) (Ici b.1)).trans <| inter_eq_left.2 <| Ici_subset_Ioi.2 b.2
(Subtype.image_preimage_val (Ioi a) (Ici b.1)).trans <| inter_eq_right.2 <| Ici_subset_Ioi.2 b.2

@[simp]
lemma image_subtype_val_Ioi_Iic {a : α} (b : Ioi a) : Subtype.val '' Iic b = Ioc a b :=
(Subtype.image_preimage_val (Ioi a) (Iic b.1)).trans Iic_inter_Ioi
(Subtype.image_preimage_val (Ioi a) (Iic b.1)).trans Ioi_inter_Iic

@[simp]
lemma image_subtype_val_Ioi_Ioi {a : α} (b : Ioi a) : Subtype.val '' Ioi b = Ioi b.1 :=
(Subtype.image_preimage_val (Ioi a) (Ioi b.1)).trans <| inter_eq_left.2 <| Ioi_subset_Ioi b.2.le
(Subtype.image_preimage_val (Ioi a) (Ioi b.1)).trans <| inter_eq_right.2 <| Ioi_subset_Ioi b.2.le

@[simp]
lemma image_subtype_val_Ioi_Iio {a : α} (b : Ioi a) : Subtype.val '' Iio b = Ioo a b :=
(Subtype.image_preimage_val (Ioi a) (Iio b.1)).trans Iio_inter_Ioi
(Subtype.image_preimage_val (Ioi a) (Iio b.1)).trans Ioi_inter_Iio

@[simp]
lemma image_subtype_val_Iio_Ici {a : α} (b : Iio a) : Subtype.val '' Ici b = Ico b.1 a :=
Subtype.image_preimage_val (Iio a) (Ici b.1)
(Subtype.image_preimage_val _ _).trans <| inter_comm _ _

@[simp]
lemma image_subtype_val_Iio_Iic {a : α} (b : Iio a) : Subtype.val '' Iic b = Iic b.1 :=
image_subtype_val_Ioi_Ici (α := αᵒᵈ) _

@[simp]
lemma image_subtype_val_Iio_Ioi {a : α} (b : Iio a) : Subtype.val '' Ioi b = Ioo b.1 a :=
Subtype.image_preimage_val (Iio a) (Ioi b.1)
(Subtype.image_preimage_val _ _).trans <| inter_comm _ _

@[simp]
lemma image_subtype_val_Iio_Iio {a : α} (b : Iio a) : Subtype.val '' Iio b = Iio b.1 :=
Expand All @@ -355,13 +355,13 @@ private lemma image_subtype_val_Ixx_Ixi {p q r : α → α → Prop} {a b : α}
(h : ∀ {x}, r c x → p a x) :
Subtype.val '' {y : {x // p a x ∧ q x b} | r c.1 y.1} = {y : α | r c.1 y ∧ q y b} :=
(Subtype.image_preimage_val {x | p a x ∧ q x b} {y | r c.1 y}).trans <| by
ext; simp (config := { contextual := true }) [h]
ext; simp (config := { contextual := true }) [@and_comm (r _ _), h]

private lemma image_subtype_val_Ixx_Iix {p q r : α → α → Prop} {a b : α} (c : {x // p a x ∧ q x b})
(h : ∀ {x}, r x c → q x b) :
Subtype.val '' {y : {x // p a x ∧ q x b} | r y.1 c.1} = {y : α | p a y ∧ r y c.1} :=
(Subtype.image_preimage_val {x | p a x ∧ q x b} {y | r y c.1}).trans <| by
ext; simp (config := { contextual := true }) [@and_comm (p _ _), h]
ext; simp (config := { contextual := true}) [h]

@[simp]
lemma image_subtype_val_Icc_Ici {a b : α} (c : Icc a b) : Subtype.val '' Ici c = Icc c.1 b :=
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Geometry/Manifold/IntegralCurve.lean
Original file line number Diff line number Diff line change
Expand Up @@ -444,7 +444,7 @@ theorem isIntegralCurveOn_Ioo_eqOn_of_contMDiff (ht₀ : t₀ ∈ Ioo a b)
⟨t₀, ⟨ht₀, ⟨h, ht₀⟩⟩⟩
· -- is this really the most convenient way to pass to subtype topology?
-- TODO: shorten this when better API around subtype topology exists
rw [hs, ← Subtype.image_preimage_val, ← Subtype.image_preimage_val,
rw [hs, inter_comm, ← Subtype.image_preimage_val, inter_comm, ← Subtype.image_preimage_val,
image_subset_image_iff Subtype.val_injective, preimage_setOf_eq]
intros t ht
rw [mem_preimage, ← closure_subtype] at ht
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/GroupTheory/Subgroup/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1679,7 +1679,7 @@ theorem mem_subgroupOf {H K : Subgroup G} {h : K} : h ∈ H.subgroupOf K ↔ (h

@[to_additive (attr := simp)]
theorem subgroupOf_map_subtype (H K : Subgroup G) : (H.subgroupOf K).map K.subtype = H ⊓ K :=
Copy link
Member

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Arguably this should flip too

Copy link
Contributor Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I added a TODO

SetLike.ext' <| Subtype.image_preimage_coe _ _
SetLike.ext' <| by refine Subtype.image_preimage_coe _ _ |>.trans ?_; apply Set.inter_comm
#align subgroup.subgroup_of_map_subtype Subgroup.subgroupOf_map_subtype
#align add_subgroup.add_subgroup_of_map_subtype AddSubgroup.addSubgroupOf_map_subtype

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/MeasureTheory/Constructions/Polish.lean
Original file line number Diff line number Diff line change
Expand Up @@ -589,7 +589,7 @@ if and only if the intesection with `Set.range f` is measurable. -/
theorem measurableSet_preimage_iff_inter_range {f : X → Y} [SecondCountableTopology (range f)]
(hf : Measurable f) (hr : MeasurableSet (range f)) {s : Set Y} :
MeasurableSet (f ⁻¹' s) ↔ MeasurableSet (s ∩ range f) := by
rw [hf.measurableSet_preimage_iff_preimage_val,
rw [hf.measurableSet_preimage_iff_preimage_val, inter_comm,
← (MeasurableEmbedding.subtype_coe hr).measurableSet_image, Subtype.image_preimage_coe]
#align measurable.measurable_set_preimage_iff_inter_range Measurable.measurableSet_preimage_iff_inter_range

Expand Down
2 changes: 1 addition & 1 deletion Mathlib/MeasureTheory/MeasurableSpace/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -612,7 +612,7 @@ theorem MeasurableSet.subtype_image {s : Set α} {t : Set s} (hs : MeasurableSet
MeasurableSet t → MeasurableSet (((↑) : s → α) '' t) := by
rintro ⟨u, hu, rfl⟩
rw [Subtype.image_preimage_coe]
exact hu.inter hs
exact hs.inter hu
#align measurable_set.subtype_image MeasurableSet.subtype_image

@[measurability]
Expand Down
1 change: 1 addition & 0 deletions Mathlib/MeasureTheory/Measure/AEMeasurable.lean
Original file line number Diff line number Diff line change
Expand Up @@ -412,6 +412,7 @@ lemma MeasureTheory.NullMeasurable.aemeasurable {f : α → β}
refine measurable_generateFrom fun s hs ↦ .of_subtype_image ?_
rw [preimage_comp, Subtype.image_preimage_coe]
convert (hTm s hs).diff hvm using 1
rw [inter_comm]
refine Set.ext fun x ↦ and_congr_left fun hxv ↦ ⟨fun hx ↦ ?_, fun hx ↦ hTf s hs hx⟩
exact by_contra fun hx' ↦ hxv <| mem_biUnion hs ⟨hUf s hs hx, hx'⟩

Expand Down
6 changes: 4 additions & 2 deletions Mathlib/MeasureTheory/Measure/OuterMeasure.lean
Original file line number Diff line number Diff line change
Expand Up @@ -551,9 +551,10 @@ def restrict (s : Set α) : OuterMeasure α →ₗ[ℝ≥0∞] OuterMeasure α :
(map (↑)).comp (comap ((↑) : s → α))
#align measure_theory.outer_measure.restrict MeasureTheory.OuterMeasure.restrict

-- TODO (kmill): change `m (t ∩ s)` to `m (s ∩ t)`
@[simp]
theorem restrict_apply (s t : Set α) (m : OuterMeasure α) : restrict s m t = m (t ∩ s) := by
simp [restrict]
simp [restrict, inter_comm t]
#align measure_theory.outer_measure.restrict_apply MeasureTheory.OuterMeasure.restrict_apply

@[mono]
Expand Down Expand Up @@ -796,11 +797,12 @@ theorem map_ofFunction {β} {f : α → β} (hf : Injective f) :
simp [hf.preimage_image]
#align measure_theory.outer_measure.map_of_function MeasureTheory.OuterMeasure.map_ofFunction

-- TODO (kmill): change `m (t ∩ s)` to `m (s ∩ t)`
theorem restrict_ofFunction (s : Set α) (hm : Monotone m) :
restrict s (OuterMeasure.ofFunction m m_empty) =
OuterMeasure.ofFunction (fun t => m (t ∩ s)) (by simp; simp [m_empty]) := by
rw [restrict]
simp only [LinearMap.comp_apply]
simp only [inter_comm _ s, LinearMap.comp_apply]
rw [comap_ofFunction _ (Or.inl hm)]
simp only [map_ofFunction Subtype.coe_injective, Subtype.image_preimage_coe]
#align measure_theory.outer_measure.restrict_of_function MeasureTheory.OuterMeasure.restrict_ofFunction
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/MeasureTheory/Measure/Restrict.lean
Original file line number Diff line number Diff line change
Expand Up @@ -810,7 +810,7 @@ theorem MeasurableSet.nullMeasurableSet_subtype_coe {t : Set s} (hs : NullMeasur
{ t : Set s | ∃ s' : Set α, MeasurableSet s' ∧ (↑) ⁻¹' s' = t } _ _ _ _ ht
· rintro t' ⟨s', hs', rfl⟩
rw [Subtype.image_preimage_coe]
exact hs'.nullMeasurableSet.inter hs
exact hs.inter (hs'.nullMeasurableSet)
· simp only [image_empty, nullMeasurableSet_empty]
· intro t'
simp only [← range_diff_image Subtype.coe_injective, Subtype.range_coe_subtype, setOf_mem_eq]
Expand Down
4 changes: 3 additions & 1 deletion Mathlib/RingTheory/GradedAlgebra/HomogeneousIdeal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -182,7 +182,9 @@ theorem Ideal.homogeneous_span (s : Set A) (h : ∀ x ∈ s, Homogeneous 𝒜 x)
is the largest homogeneous ideal of `A` contained in `I`. -/
def Ideal.homogeneousCore : HomogeneousIdeal 𝒜 :=
⟨Ideal.homogeneousCore' 𝒜 I,
Ideal.homogeneous_span _ _ fun _ h => (Subtype.image_preimage_coe _ _ ▸ h).2⟩
Ideal.homogeneous_span _ _ fun _ h => by
have := Subtype.image_preimage_coe (setOf (Homogeneous 𝒜)) (I : Set A)
exact (cast congr(_ ∈ $this) h).1⟩
#align ideal.homogeneous_core Ideal.homogeneousCore

theorem Ideal.homogeneousCore_mono : Monotone (Ideal.homogeneousCore 𝒜) :=
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Topology/Connected/Basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -640,14 +640,14 @@ theorem IsPreconnected.subset_connectedComponentIn {x : α} {F : Set α} (hs : I
(hxs : x ∈ s) (hsF : s ⊆ F) : s ⊆ connectedComponentIn F x := by
have : IsPreconnected (((↑) : F → α) ⁻¹' s) := by
refine' inducing_subtype_val.isPreconnected_image.mp _
rwa [Subtype.image_preimage_coe, inter_eq_left.mpr hsF]
rwa [Subtype.image_preimage_coe, inter_eq_right.mpr hsF]
have h2xs : (⟨x, hsF hxs⟩ : F) ∈ (↑) ⁻¹' s := by
rw [mem_preimage]
exact hxs
have := this.subset_connectedComponent h2xs
rw [connectedComponentIn_eq_image (hsF hxs)]
refine' Subset.trans _ (image_subset _ this)
rw [Subtype.image_preimage_coe, inter_eq_left.mpr hsF]
rw [Subtype.image_preimage_coe, inter_eq_right.mpr hsF]
#align is_preconnected.subset_connected_component_in IsPreconnected.subset_connectedComponentIn

theorem IsConnected.subset_connectedComponent {x : α} {s : Set α} (H1 : IsConnected s)
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Topology/Constructions.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1706,12 +1706,12 @@ variable [TopologicalSpace X] {s : Set X} {t : Set s}
theorem IsOpen.trans (ht : IsOpen t) (hs : IsOpen s) : IsOpen (t : Set X) := by
rcases isOpen_induced_iff.mp ht with ⟨s', hs', rfl⟩
rw [Subtype.image_preimage_coe]
exact hs'.inter hs
exact hs.inter hs'

theorem IsClosed.trans (ht : IsClosed t) (hs : IsClosed s) : IsClosed (t : Set X) := by
rcases isClosed_induced_iff.mp ht with ⟨s', hs', rfl⟩
rw [Subtype.image_preimage_coe]
convert hs'.inter hs
exact hs.inter hs'

end Monad

Expand Down
6 changes: 3 additions & 3 deletions Mathlib/Topology/ContinuousOn.lean
Original file line number Diff line number Diff line change
Expand Up @@ -669,7 +669,7 @@ theorem continuousOn_iff' {f : α → β} {s : Set α} :
simp only [Subtype.preimage_coe_eq_preimage_coe_iff]
constructor <;>
· rintro ⟨u, ou, useq⟩
exact ⟨u, ou, useq.symm
exact ⟨u, ou, by simpa only [Set.inter_comm, eq_comm] using useq
rw [continuousOn_iff_continuous_restrict, continuous_def]; simp only [this]
#align continuous_on_iff' continuousOn_iff'

Expand All @@ -694,7 +694,7 @@ theorem continuousOn_iff_isClosed {f : α → β} {s : Set α} :
have : ∀ t, IsClosed (s.restrict f ⁻¹' t) ↔ ∃ u : Set α, IsClosed u ∧ f ⁻¹' t ∩ s = u ∩ s := by
intro t
rw [isClosed_induced_iff, Set.restrict_eq, Set.preimage_comp]
simp only [Subtype.preimage_coe_eq_preimage_coe_iff, eq_comm]
simp only [Subtype.preimage_coe_eq_preimage_coe_iff, eq_comm, Set.inter_comm s]
rw [continuousOn_iff_continuous_restrict, continuous_iff_isClosed]; simp only [this]
#align continuous_on_iff_is_closed continuousOn_iff_isClosed

Expand Down Expand Up @@ -1325,7 +1325,7 @@ theorem continuousOn_piecewise_ite {s s' t : Set α} {f f' : α → β} [∀ x,

theorem frontier_inter_open_inter {s t : Set α} (ht : IsOpen t) :
frontier (s ∩ t) ∩ t = frontier s ∩ t := by
simp only [← Subtype.preimage_coe_eq_preimage_coe_iff,
simp only [Set.inter_comm _ t, ← Subtype.preimage_coe_eq_preimage_coe_iff,
ht.isOpenMap_subtype_val.preimage_frontier_eq_frontier_preimage continuous_subtype_val,
Subtype.preimage_coe_inter_self]
#align frontier_inter_open_inter frontier_inter_open_inter
Expand Down
4 changes: 2 additions & 2 deletions Mathlib/Topology/Separation.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2376,12 +2376,12 @@ theorem loc_compact_Haus_tot_disc_of_zero_dim [TotallyDisconnectedSpace H] :
refine' ⟨f0, _⟩
· have : Set.range ((↑) : u → H) = interior s := by
rw [this, Set.range_comp, Subtype.range_coe, Subtype.image_preimage_coe]
apply Set.inter_eq_self_of_subset_left interior_subset
apply Set.inter_eq_self_of_subset_right interior_subset
rw [this]
apply isOpen_interior
have f2 : IsOpen v := VisClopen.2.preimage continuous_subtype_val
have f3 : ((↑) : s → H) '' V = ((↑) : u → H) '' v := by
rw [this, image_comp, Subtype.image_preimage_coe, inter_eq_self_of_subset_left V_sub]
rw [this, image_comp, Subtype.image_preimage_coe, inter_eq_self_of_subset_right V_sub]
rw [f3]
apply f1.isOpenMap v f2
refine' ⟨(↑) '' V, VisClopen', by simp [Vx], Subset.trans _ sU⟩
Expand Down
2 changes: 1 addition & 1 deletion Mathlib/Topology/Sober.lean
Original file line number Diff line number Diff line change
Expand Up @@ -236,7 +236,7 @@ theorem quasiSober_of_open_cover (S : Set (Set α)) (hS : ∀ s : S, IsOpen (s :
rw [← image_singleton, ← closure_image_closure continuous_subtype_val, H.genericPoint_spec.def]
refine' (subset_closure_inter_of_isPreirreducible_of_isOpen h.2 (hS ⟨U, hU⟩) ⟨x, hx, hU'⟩).trans
(closure_mono _)
rw [← Subtype.image_preimage_coe]
rw [inter_comm t, ← Subtype.image_preimage_coe]
exact Set.image_subset _ subset_closure
#align quasi_sober_of_open_cover quasiSober_of_open_cover

Expand Down