@@ -6,7 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Yury Kudryashov
6
6
import data.set.intervals.pi
7
7
import data.set.pointwise.interval
8
8
import order.filter.interval
9
- import topology.algebra.group.basic
9
+ import topology.support
10
10
import topology.algebra.order.left_right
11
11
12
12
/-!
@@ -94,9 +94,6 @@ instance [topological_space α] [h : first_countable_topology α] : first_counta
94
94
instance [topological_space α] [h : second_countable_topology α] : second_countable_topology αᵒᵈ :=
95
95
h
96
96
97
- @[to_additive]
98
- instance [topological_space α] [has_mul α] [h : has_continuous_mul α] : has_continuous_mul αᵒᵈ := h
99
-
100
97
lemma dense.order_dual [topological_space α] {s : set α} (hs : dense s) :
101
98
dense (order_dual.of_dual ⁻¹' s) := hs
102
99
@@ -1608,6 +1605,7 @@ end order_topology
1608
1605
end linear_order
1609
1606
1610
1607
section linear_ordered_add_comm_group
1608
+
1611
1609
variables [topological_space α] [linear_ordered_add_comm_group α] [order_topology α]
1612
1610
variables {l : filter β} {f g : β → α}
1613
1611
@@ -1643,50 +1641,35 @@ lemma eventually_abs_sub_lt (a : α) {ε : α} (hε : 0 < ε) : ∀ᶠ x in 𝓝
1643
1641
(nhds_eq_infi_abs_sub a).symm ▸ mem_infi_of_mem ε
1644
1642
(mem_infi_of_mem hε $ by simp only [abs_sub_comm, mem_principal_self])
1645
1643
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 :=
1683
1648
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)
1688
1653
end
1689
1654
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
+
1690
1673
lemma nhds_basis_Ioo_pos [no_min_order α] [no_max_order α] (a : α) :
1691
1674
(𝓝 a).has_basis (λ ε : α, (0 : α) < ε) (λ ε, Ioo (a-ε) (a+ε)) :=
1692
1675
⟨begin
@@ -1729,54 +1712,6 @@ lemma nhds_basis_Ioo_pos_of_pos [no_min_order α] [no_max_order α]
1729
1712
(sub_le_sub_left (min_le_left i a) a) (add_le_add_left (min_le_left i a) a)) hit⟩,
1730
1713
λ h, let ⟨i, hi, hit⟩ := h in ⟨i, hi.1 , hit⟩ ⟩ ⟩
1731
1714
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
-
1780
1715
end linear_ordered_add_comm_group
1781
1716
1782
1717
lemma preimage_neg [add_group α] : preimage (has_neg.neg : α → α) = image (has_neg.neg : α → α) :=
@@ -2693,18 +2628,15 @@ section nhds_with_pos
2693
2628
2694
2629
section linear_ordered_add_comm_group
2695
2630
2696
- variables [linear_ordered_add_comm_group α] [topological_space α] [order_topology α]
2631
+ variables [linear_order α] [has_zero α] [topological_space α] [order_topology α]
2697
2632
2698
2633
lemma eventually_nhds_within_pos_mem_Ioo {ε : α} (h : 0 < ε) :
2699
2634
∀ᶠ 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)
2704
2636
2705
2637
lemma eventually_nhds_within_pos_mem_Ioc {ε : α} (h : 0 < ε) :
2706
2638
∀ᶠ 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)
2708
2640
2709
2641
end linear_ordered_add_comm_group
2710
2642
0 commit comments