Skip to content

Commit

Permalink
chore(order/filter): turn tendsto_id' into an iff lemma (#14791)
Browse files Browse the repository at this point in the history
  • Loading branch information
urkud committed Jun 24, 2022
1 parent 6cefaf4 commit efe794c
Show file tree
Hide file tree
Showing 6 changed files with 11 additions and 15 deletions.
2 changes: 1 addition & 1 deletion src/dynamics/omega_limit.lean
Expand Up @@ -68,7 +68,7 @@ begin
end

lemma omega_limit_mono_left {f₁ f₂ : filter τ} (hf : f₁ ≤ f₂) : ω f₁ ϕ s ⊆ ω f₂ ϕ s :=
omega_limit_subset_of_tendsto ϕ s (tendsto_id' hf)
omega_limit_subset_of_tendsto ϕ s (tendsto_id'.2 hf)

lemma omega_limit_mono_right {s₁ s₂ : set α} (hs : s₁ ⊆ s₂) : ω f ϕ s₁ ⊆ ω f ϕ s₂ :=
Inter₂_mono $ λ u hu, closure_mono (image2_subset subset.rfl hs)
Expand Down
2 changes: 1 addition & 1 deletion src/measure_theory/measure/haar_lebesgue.lean
Expand Up @@ -445,7 +445,7 @@ begin
have A : tendsto (λ (r : ℝ), ennreal.of_real (r ^ (finrank ℝ E)) * μ (closed_ball (0 : E) 1))
(𝓝[<] 1) (𝓝 (ennreal.of_real (1 ^ (finrank ℝ E)) * μ (closed_ball (0 : E) 1))),
{ refine ennreal.tendsto.mul _ (by simp) tendsto_const_nhds (by simp),
exact ennreal.tendsto_of_real ((tendsto_id' nhds_within_le_nhds).pow _) },
exact ennreal.tendsto_of_real ((tendsto_id'.2 nhds_within_le_nhds).pow _) },
simp only [one_pow, one_mul, ennreal.of_real_one] at A,
refine le_of_tendsto A _,
refine mem_nhds_within_Iio_iff_exists_Ioo_subset.2 ⟨(0 : ℝ), by simp, λ r hr, _⟩,
Expand Down
9 changes: 3 additions & 6 deletions src/order/filter/basic.lean
Expand Up @@ -2320,16 +2320,13 @@ theorem tendsto.congr {f₁ f₂ : α → β} {l₁ : filter α} {l₂ : filter
(h : ∀ x, f₁ x = f₂ x) : tendsto f₁ l₁ l₂ → tendsto f₂ l₁ l₂ :=
(tendsto_congr h).1

lemma tendsto_id' {x y : filter α} : x ≤ y → tendsto id x y :=
by simp only [tendsto, map_id, forall_true_iff] {contextual := tt}
lemma tendsto_id' {x y : filter α} : tendsto id x y ↔ x ≤ y := iff.rfl

lemma tendsto_id {x : filter α} : tendsto id x x := tendsto_id' $ le_refl x
lemma tendsto_id {x : filter α} : tendsto id x x := le_refl x

lemma tendsto.comp {f : α → β} {g : β → γ} {x : filter α} {y : filter β} {z : filter γ}
(hg : tendsto g y z) (hf : tendsto f x y) : tendsto (g ∘ f) x z :=
calc map (g ∘ f) x = map g (map f x) : by rw [map_map]
... ≤ map g y : map_mono hf
... ≤ z : hg
λ s hs, hf (hg hs)

lemma tendsto.mono_left {f : α → β} {x y : filter α} {z : filter β}
(hx : tendsto f x z) (h : y ≤ x) : tendsto f y z :=
Expand Down
2 changes: 1 addition & 1 deletion src/topology/instances/ennreal.lean
Expand Up @@ -614,7 +614,7 @@ have Inf ((λb, ↑r - b) '' range b) = ↑r - (⨆i, b i),
from is_glb.Inf_eq $ is_lub_supr.is_glb_of_tendsto
(assume x _ y _, tsub_le_tsub (le_refl (r : ℝ≥0∞)))
(range_nonempty _)
(ennreal.tendsto_coe_sub.comp (tendsto_id' inf_le_left)),
(ennreal.tendsto_coe_sub.comp (tendsto_id'.2 inf_le_left)),
by rw [eq, ←this]; simp [Inf_image, infi_range, -mem_range]; exact le_rfl

lemma exists_countable_dense_no_zero_top :
Expand Down
9 changes: 4 additions & 5 deletions src/topology/order.lean
Expand Up @@ -725,12 +725,11 @@ continuous_iff_le_induced.2 $ bot_le
@[continuity] lemma continuous_top {t : tspace α} : cont t ⊤ f :=
continuous_iff_coinduced_le.2 $ le_top

lemma continuous_id_iff_le {t t' : tspace α} : cont t t' id ↔ t ≤ t' :=
@continuous_def _ _ t t' id

lemma continuous_id_of_le {t t' : tspace α} (h : t ≤ t') : cont t t' id :=
begin
rw continuous_def,
assume u hu,
exact h u hu
end
continuous_id_iff_le.2 h

/- 𝓝 in the induced topology -/

Expand Down
2 changes: 1 addition & 1 deletion src/topology/separation.lean
Expand Up @@ -587,7 +587,7 @@ lemma eq_of_tendsto_nhds [topological_space β] [t1_space β] {f : α → β} {a
(h : tendsto f (𝓝 a) (𝓝 b)) : f a = b :=
by_contra $ assume (hfa : f a ≠ b),
have fact₁ : {f a}ᶜ ∈ 𝓝 b := compl_singleton_mem_nhds hfa.symm,
have fact₂ : tendsto f (pure a) (𝓝 b) := h.comp (tendsto_id' $ pure_le_nhds a),
have fact₂ : tendsto f (pure a) (𝓝 b) := h.comp (tendsto_id'.2 $ pure_le_nhds a),
fact₂ fact₁ (eq.refl $ f a)

/-- To prove a function to a `t1_space` is continuous at some point `a`, it suffices to prove that
Expand Down

0 comments on commit efe794c

Please sign in to comment.