Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 84dc0bd

Browse files
committed
chore(topology/order/basic): split (#18363)
Move material that requires topological groups to a new file to reduce transitive imports.
1 parent 2705404 commit 84dc0bd

File tree

11 files changed

+123
-104
lines changed

11 files changed

+123
-104
lines changed

src/analysis/box_integral/partition/basic.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -3,6 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yury Kudryashov
55
-/
6+
import algebra.big_operators.option
67
import analysis.box_integral.box.basic
78

89
/-!

src/analysis/convex/strict.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yaël Dillies
55
-/
66
import analysis.convex.basic
7-
import topology.order.basic
7+
import topology.algebra.order.group
88

99
/-!
1010
# Strictly convex sets

src/topology/algebra/monoid.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -49,6 +49,8 @@ section has_continuous_mul
4949

5050
variables [topological_space M] [has_mul M] [has_continuous_mul M]
5151

52+
@[to_additive] instance : has_continuous_mul Mᵒᵈ := ‹has_continuous_mul M›
53+
5254
@[to_additive]
5355
lemma continuous_mul : continuous (λp:M×M, p.1 * p.2) :=
5456
has_continuous_mul.continuous_mul

src/topology/algebra/order/field.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@ Authors: Benjamin Davidson, Devon Tuma, Eric Rodriguez, Oliver Nash
55
-/
66

77
import tactic.positivity
8-
import topology.order.basic
8+
import topology.algebra.order.group
99
import topology.algebra.field
1010

1111
/-!

src/topology/algebra/order/floor.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Anatole Dedecker
55
-/
66
import algebra.order.floor
7-
import topology.order.basic
7+
import topology.algebra.order.group
88

99
/-!
1010
# Topological facts about `int.floor`, `int.ceil` and `int.fract`
@@ -156,7 +156,7 @@ tendsto_nhds_within_of_tendsto_nhds_of_eventually_within _
156156

157157
local notation `I` := (Icc 0 1 : set α)
158158

159-
variables [order_topology α] [topological_add_group α] [topological_space β] [topological_space γ]
159+
variables [order_topology α] [topological_space β] [topological_space γ]
160160

161161
/-- Do not use this, use `continuous_on.comp_fract` instead. -/
162162
lemma continuous_on.comp_fract' {f : β → α → γ}
Lines changed: 80 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,80 @@
1+
/-
2+
Copyright (c) 2020 Yury Kudryashov. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Yury Kudryashov
5+
-/
6+
import topology.order.basic
7+
import topology.algebra.group.basic
8+
9+
/-!
10+
# Topology on a linear ordered additive commutative group
11+
12+
In this file we prove that a linear ordered additive commutative group with order topology is a
13+
topological group. We also prove continuity of `abs : G → G` and provide convenience lemmas like
14+
`continuous_at.abs`.
15+
-/
16+
17+
open set filter
18+
open_locale topology filter
19+
20+
variables {α G : Type*} [topological_space G] [linear_ordered_add_comm_group G] [order_topology G]
21+
variables {l : filter α} {f g : α → G}
22+
23+
@[priority 100] -- see Note [lower instance priority]
24+
instance linear_ordered_add_comm_group.topological_add_group : topological_add_group G :=
25+
{ continuous_add :=
26+
begin
27+
refine continuous_iff_continuous_at.2 _,
28+
rintro ⟨a, b⟩,
29+
refine linear_ordered_add_comm_group.tendsto_nhds.2 (λ ε ε0, _),
30+
rcases dense_or_discrete 0 ε with (⟨δ, δ0, δε⟩|⟨h₁, h₂⟩),
31+
{ -- If there exists `δ ∈ (0, ε)`, then we choose `δ`-nhd of `a` and `(ε-δ)`-nhd of `b`
32+
filter_upwards [(eventually_abs_sub_lt a δ0).prod_nhds
33+
(eventually_abs_sub_lt b (sub_pos.2 δε))],
34+
rintros ⟨x, y⟩ ⟨hx : |x - a| < δ, hy : |y - b| < ε - δ⟩,
35+
rw [add_sub_add_comm],
36+
calc |x - a + (y - b)| ≤ |x - a| + |y - b| : abs_add _ _
37+
... < δ + (ε - δ) : add_lt_add hx hy
38+
... = ε : add_sub_cancel'_right _ _ },
39+
{ -- Otherwise `ε`-nhd of each point `a` is `{a}`
40+
have hε : ∀ {x y}, |x - y| < ε → x = y,
41+
{ intros x y h,
42+
simpa [sub_eq_zero] using h₂ _ h },
43+
filter_upwards [(eventually_abs_sub_lt a ε0).prod_nhds (eventually_abs_sub_lt b ε0)],
44+
rintros ⟨x, y⟩ ⟨hx : |x - a| < ε, hy : |y - b| < ε⟩,
45+
simpa [hε hx, hε hy] }
46+
end,
47+
continuous_neg := continuous_iff_continuous_at.2 $ λ a,
48+
linear_ordered_add_comm_group.tendsto_nhds.2 $ λ ε ε0,
49+
(eventually_abs_sub_lt a ε0).mono $ λ x hx, by rwa [neg_sub_neg, abs_sub_comm] }
50+
51+
@[continuity]
52+
lemma continuous_abs : continuous (abs : G → G) := continuous_id.max continuous_neg
53+
54+
protected lemma filter.tendsto.abs {a : G} (h : tendsto f l (𝓝 a)) :
55+
tendsto (λ x, |f x|) l (𝓝 (|a|)) :=
56+
(continuous_abs.tendsto _).comp h
57+
58+
lemma tendsto_zero_iff_abs_tendsto_zero (f : α → G) :
59+
tendsto f l (𝓝 0) ↔ tendsto (abs ∘ f) l (𝓝 0) :=
60+
begin
61+
refine ⟨λ h, (abs_zero : |(0 : G)| = 0) ▸ h.abs, λ h, _⟩,
62+
have : tendsto (λ a, -|f a|) l (𝓝 0) := (neg_zero : -(0 : G) = 0) ▸ h.neg,
63+
exact tendsto_of_tendsto_of_tendsto_of_le_of_le this h
64+
(λ x, neg_abs_le_self $ f x) (λ x, le_abs_self $ f x),
65+
end
66+
67+
variables [topological_space α] {a : α} {s : set α}
68+
69+
protected lemma continuous.abs (h : continuous f) : continuous (λ x, |f x|) := continuous_abs.comp h
70+
71+
protected lemma continuous_at.abs (h : continuous_at f a) : continuous_at (λ x, |f x|) a := h.abs
72+
73+
protected lemma continuous_within_at.abs (h : continuous_within_at f s a) :
74+
continuous_within_at (λ x, |f x|) s a := h.abs
75+
76+
protected lemma continuous_on.abs (h : continuous_on f s) : continuous_on (λ x, |f x|) s :=
77+
λ x hx, (h x hx).abs
78+
79+
lemma tendsto_abs_nhds_within_zero : tendsto (abs : G → G) (𝓝[≠] 0) (𝓝[>] 0) :=
80+
(continuous_abs.tendsto' (0 : G) 0 abs_zero).inf $ tendsto_principal_principal.2 $ λ x, abs_pos.2

src/topology/algebra/order/liminf_limsup.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov
55
-/
66
import algebra.big_operators.intervals
7+
import algebra.big_operators.order
8+
import algebra.indicator_function
79
import order.liminf_limsup
810
import order.filter.archimedean
911
import topology.order.basic

src/topology/algebra/order/monotone_continuity.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Yury G. Kudryashov, Heather Macbeth
55
-/
66
import topology.order.basic
7-
import topology.algebra.order.left_right
7+
import topology.homeomorph
88

99
/-!
1010
# Continuity of monotone functions

src/topology/continuous_function/ordered.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -5,6 +5,7 @@ Authors: Scott Morrison, Shing Tak Lam
55
-/
66

77
import topology.algebra.order.proj_Icc
8+
import topology.algebra.order.group
89
import topology.continuous_function.basic
910

1011
/-!

src/topology/order/basic.lean

Lines changed: 31 additions & 99 deletions
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov
66
import data.set.intervals.pi
77
import data.set.pointwise.interval
88
import order.filter.interval
9-
import topology.algebra.group.basic
9+
import topology.support
1010
import topology.algebra.order.left_right
1111

1212
/-!
@@ -94,9 +94,6 @@ instance [topological_space α] [h : first_countable_topology α] : first_counta
9494
instance [topological_space α] [h : second_countable_topology α] : second_countable_topology αᵒᵈ :=
9595
h
9696

97-
@[to_additive]
98-
instance [topological_space α] [has_mul α] [h : has_continuous_mul α] : has_continuous_mul αᵒᵈ := h
99-
10097
lemma dense.order_dual [topological_space α] {s : set α} (hs : dense s) :
10198
dense (order_dual.of_dual ⁻¹' s) := hs
10299

@@ -1608,6 +1605,7 @@ end order_topology
16081605
end linear_order
16091606

16101607
section linear_ordered_add_comm_group
1608+
16111609
variables [topological_space α] [linear_ordered_add_comm_group α] [order_topology α]
16121610
variables {l : filter β} {f g : β → α}
16131611

@@ -1643,50 +1641,35 @@ lemma eventually_abs_sub_lt (a : α) {ε : α} (hε : 0 < ε) : ∀ᶠ x in 𝓝
16431641
(nhds_eq_infi_abs_sub a).symm ▸ mem_infi_of_mem ε
16441642
(mem_infi_of_mem hε $ by simp only [abs_sub_comm, mem_principal_self])
16451643

1646-
@[priority 100] -- see Note [lower instance priority]
1647-
instance linear_ordered_add_comm_group.topological_add_group : topological_add_group α :=
1648-
{ continuous_add :=
1649-
begin
1650-
refine continuous_iff_continuous_at.2 _,
1651-
rintro ⟨a, b⟩,
1652-
refine linear_ordered_add_comm_group.tendsto_nhds.2 (λ ε ε0, _),
1653-
rcases dense_or_discrete 0 ε with (⟨δ, δ0, δε⟩|⟨h₁, h₂⟩),
1654-
{ -- If there exists `δ ∈ (0, ε)`, then we choose `δ`-nhd of `a` and `(ε-δ)`-nhd of `b`
1655-
filter_upwards [(eventually_abs_sub_lt a δ0).prod_nhds
1656-
(eventually_abs_sub_lt b (sub_pos.2 δε))],
1657-
rintros ⟨x, y⟩ ⟨hx : |x - a| < δ, hy : |y - b| < ε - δ⟩,
1658-
rw [add_sub_add_comm],
1659-
calc |x - a + (y - b)| ≤ |x - a| + |y - b| : abs_add _ _
1660-
... < δ + (ε - δ) : add_lt_add hx hy
1661-
... = ε : add_sub_cancel'_right _ _ },
1662-
{ -- Otherwise `ε`-nhd of each point `a` is `{a}`
1663-
have hε : ∀ {x y}, |x - y| < ε → x = y,
1664-
{ intros x y h,
1665-
simpa [sub_eq_zero] using h₂ _ h },
1666-
filter_upwards [(eventually_abs_sub_lt a ε0).prod_nhds (eventually_abs_sub_lt b ε0)],
1667-
rintros ⟨x, y⟩ ⟨hx : |x - a| < ε, hy : |y - b| < ε⟩,
1668-
simpa [hε hx, hε hy] }
1669-
end,
1670-
continuous_neg := continuous_iff_continuous_at.2 $ λ a,
1671-
linear_ordered_add_comm_group.tendsto_nhds.2 $ λ ε ε0,
1672-
(eventually_abs_sub_lt a ε0).mono $ λ x hx, by rwa [neg_sub_neg, abs_sub_comm] }
1673-
1674-
@[continuity]
1675-
lemma continuous_abs : continuous (abs : α → α) := continuous_id.max continuous_neg
1676-
1677-
lemma filter.tendsto.abs {f : β → α} {a : α} {l : filter β} (h : tendsto f l (𝓝 a)) :
1678-
tendsto (λ x, |f x|) l (𝓝 (|a|)) :=
1679-
(continuous_abs.tendsto _).comp h
1680-
1681-
lemma tendsto_zero_iff_abs_tendsto_zero (f : β → α) {l : filter β} :
1682-
tendsto f l (𝓝 0) ↔ tendsto (abs ∘ f) l (𝓝 0) :=
1644+
/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `C`
1645+
and `g` tends to `at_top` then `f + g` tends to `at_top`. -/
1646+
lemma filter.tendsto.add_at_top {C : α} (hf : tendsto f l (𝓝 C)) (hg : tendsto g l at_top) :
1647+
tendsto (λ x, f x + g x) l at_top :=
16831648
begin
1684-
refine ⟨λ h, (abs_zero : |(0 : α)| = 0) ▸ h.abs, λ h, _⟩,
1685-
have : tendsto (λ a, -|f a|) l (𝓝 0) := (neg_zero : -(0 : α) = 0) ▸ h.neg,
1686-
exact tendsto_of_tendsto_of_tendsto_of_le_of_le this h
1687-
(λ x, neg_abs_le_self $ f x) (λ x, le_abs_self $ f x),
1649+
nontriviality α,
1650+
obtain ⟨C', hC'⟩ : ∃ C', C' < C := exists_lt C,
1651+
refine tendsto_at_top_add_left_of_le' _ C' _ hg,
1652+
exact (hf.eventually (lt_mem_nhds hC')).mono (λ x, le_of_lt)
16881653
end
16891654

1655+
/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `C`
1656+
and `g` tends to `at_bot` then `f + g` tends to `at_bot`. -/
1657+
lemma filter.tendsto.add_at_bot {C : α} (hf : tendsto f l (𝓝 C)) (hg : tendsto g l at_bot) :
1658+
tendsto (λ x, f x + g x) l at_bot :=
1659+
@filter.tendsto.add_at_top αᵒᵈ _ _ _ _ _ _ _ _ hf hg
1660+
1661+
/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to
1662+
`at_top` and `g` tends to `C` then `f + g` tends to `at_top`. -/
1663+
lemma filter.tendsto.at_top_add {C : α} (hf : tendsto f l at_top) (hg : tendsto g l (𝓝 C)) :
1664+
tendsto (λ x, f x + g x) l at_top :=
1665+
by { conv in (_ + _) { rw add_comm }, exact hg.add_at_top hf }
1666+
1667+
/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to
1668+
`at_bot` and `g` tends to `C` then `f + g` tends to `at_bot`. -/
1669+
lemma filter.tendsto.at_bot_add {C : α} (hf : tendsto f l at_bot) (hg : tendsto g l (𝓝 C)) :
1670+
tendsto (λ x, f x + g x) l at_bot :=
1671+
by { conv in (_ + _) { rw add_comm }, exact hg.add_at_bot hf }
1672+
16901673
lemma nhds_basis_Ioo_pos [no_min_order α] [no_max_order α] (a : α) :
16911674
(𝓝 a).has_basis (λ ε : α, (0 : α) < ε) (λ ε, Ioo (a-ε) (a+ε)) :=
16921675
begin
@@ -1729,54 +1712,6 @@ lemma nhds_basis_Ioo_pos_of_pos [no_min_order α] [no_max_order α]
17291712
(sub_le_sub_left (min_le_left i a) a) (add_le_add_left (min_le_left i a) a)) hit⟩,
17301713
λ h, let ⟨i, hi, hit⟩ := h in ⟨i, hi.1, hit⟩ ⟩ ⟩
17311714

1732-
section
1733-
1734-
variables [topological_space β] {b : β} {a : α} {s : set β}
1735-
1736-
lemma continuous.abs (h : continuous f) : continuous (λ x, |f x|) := continuous_abs.comp h
1737-
1738-
lemma continuous_at.abs (h : continuous_at f b) : continuous_at (λ x, |f x|) b := h.abs
1739-
1740-
lemma continuous_within_at.abs (h : continuous_within_at f s b) :
1741-
continuous_within_at (λ x, |f x|) s b := h.abs
1742-
1743-
lemma continuous_on.abs (h : continuous_on f s) : continuous_on (λ x, |f x|) s :=
1744-
λ x hx, (h x hx).abs
1745-
1746-
lemma tendsto_abs_nhds_within_zero : tendsto (abs : α → α) (𝓝[≠] 0) (𝓝[>] 0) :=
1747-
(continuous_abs.tendsto' (0 : α) 0 abs_zero).inf $ tendsto_principal_principal.2 $ λ x, abs_pos.2
1748-
1749-
end
1750-
1751-
/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `C`
1752-
and `g` tends to `at_top` then `f + g` tends to `at_top`. -/
1753-
lemma filter.tendsto.add_at_top {C : α} (hf : tendsto f l (𝓝 C)) (hg : tendsto g l at_top) :
1754-
tendsto (λ x, f x + g x) l at_top :=
1755-
begin
1756-
nontriviality α,
1757-
obtain ⟨C', hC'⟩ : ∃ C', C' < C := exists_lt C,
1758-
refine tendsto_at_top_add_left_of_le' _ C' _ hg,
1759-
exact (hf.eventually (lt_mem_nhds hC')).mono (λ x, le_of_lt)
1760-
end
1761-
1762-
/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `C`
1763-
and `g` tends to `at_bot` then `f + g` tends to `at_bot`. -/
1764-
lemma filter.tendsto.add_at_bot {C : α} (hf : tendsto f l (𝓝 C)) (hg : tendsto g l at_bot) :
1765-
tendsto (λ x, f x + g x) l at_bot :=
1766-
@filter.tendsto.add_at_top αᵒᵈ _ _ _ _ _ _ _ _ hf hg
1767-
1768-
/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to
1769-
`at_top` and `g` tends to `C` then `f + g` tends to `at_top`. -/
1770-
lemma filter.tendsto.at_top_add {C : α} (hf : tendsto f l at_top) (hg : tendsto g l (𝓝 C)) :
1771-
tendsto (λ x, f x + g x) l at_top :=
1772-
by { conv in (_ + _) { rw add_comm }, exact hg.add_at_top hf }
1773-
1774-
/-- In a linearly ordered additive commutative group with the order topology, if `f` tends to
1775-
`at_bot` and `g` tends to `C` then `f + g` tends to `at_bot`. -/
1776-
lemma filter.tendsto.at_bot_add {C : α} (hf : tendsto f l at_bot) (hg : tendsto g l (𝓝 C)) :
1777-
tendsto (λ x, f x + g x) l at_bot :=
1778-
by { conv in (_ + _) { rw add_comm }, exact hg.add_at_bot hf }
1779-
17801715
end linear_ordered_add_comm_group
17811716

17821717
lemma preimage_neg [add_group α] : preimage (has_neg.neg : α → α) = image (has_neg.neg : α → α) :=
@@ -2693,18 +2628,15 @@ section nhds_with_pos
26932628

26942629
section linear_ordered_add_comm_group
26952630

2696-
variables [linear_ordered_add_comm_group α] [topological_space α] [order_topology α]
2631+
variables [linear_order α] [has_zero α] [topological_space α] [order_topology α]
26972632

26982633
lemma eventually_nhds_within_pos_mem_Ioo {ε : α} (h : 0 < ε) :
26992634
∀ᶠ x in 𝓝[>] 0, x ∈ Ioo 0 ε :=
2700-
begin
2701-
rw [eventually_iff, mem_nhds_within],
2702-
exact ⟨Ioo (-ε) ε, is_open_Ioo, by simp [h], λ x hx, ⟨hx.2, hx.1.2⟩⟩,
2703-
end
2635+
Ioo_mem_nhds_within_Ioi (left_mem_Ico.2 h)
27042636

27052637
lemma eventually_nhds_within_pos_mem_Ioc {ε : α} (h : 0 < ε) :
27062638
∀ᶠ x in 𝓝[>] 0, x ∈ Ioc 0 ε :=
2707-
(eventually_nhds_within_pos_mem_Ioo h).mono Ioo_subset_Ioc_self
2639+
Ioc_mem_nhds_within_Ioi (left_mem_Ico.2 h)
27082640

27092641
end linear_ordered_add_comm_group
27102642

0 commit comments

Comments
 (0)