@@ -1527,8 +1527,8 @@ let ⟨u', hu'⟩ := exists_gt a in mem_nhds_within_Ici_iff_exists_Ico_subset' h
1527
1527
1528
1528
/-- A set is a neighborhood of `a` within `[a, +∞)` if and only if it contains an interval `[a, u]`
1529
1529
with `a < u`. -/
1530
- lemma mem_nhds_within_Ici_iff_exists_Icc_subset' [no_max_order α] [densely_ordered α]
1531
- {a : α} {s : set α} : s ∈ 𝓝[≥] a ↔ ∃u ∈ Ioi a, Icc a u ⊆ s :=
1530
+ lemma mem_nhds_within_Ici_iff_exists_Icc_subset [no_max_order α] [densely_ordered α]
1531
+ {a : α} {s : set α} : s ∈ 𝓝[≥] a ↔ ∃ u, a < u ∧ Icc a u ⊆ s :=
1532
1532
begin
1533
1533
rw mem_nhds_within_Ici_iff_exists_Ico_subset,
1534
1534
split,
@@ -1573,42 +1573,14 @@ let ⟨l', hl'⟩ := exists_lt a in mem_nhds_within_Iic_iff_exists_Ioc_subset' h
1573
1573
1574
1574
/-- A set is a neighborhood of `a` within `(-∞, a]` if and only if it contains an interval `[l, a]`
1575
1575
with `l < a`. -/
1576
- lemma mem_nhds_within_Iic_iff_exists_Icc_subset' [no_min_order α] [densely_ordered α]
1577
- {a : α} {s : set α} : s ∈ 𝓝[≤] a ↔ ∃l ∈ Iio a, Icc l a ⊆ s :=
1576
+ lemma mem_nhds_within_Iic_iff_exists_Icc_subset [no_min_order α] [densely_ordered α]
1577
+ {a : α} {s : set α} : s ∈ 𝓝[≤] a ↔ ∃ l, l < a ∧ Icc l a ⊆ s :=
1578
1578
begin
1579
- convert @mem_nhds_within_Ici_iff_exists_Icc_subset' αᵒᵈ _ _ _ _ _ _ _,
1579
+ convert @mem_nhds_within_Ici_iff_exists_Icc_subset αᵒᵈ _ _ _ _ _ _ _,
1580
1580
simp_rw (show ∀ u : αᵒᵈ, @Icc αᵒᵈ _ a u = @Icc α _ u a, from λ u, dual_Icc),
1581
1581
refl,
1582
1582
end
1583
1583
1584
- /-- A set is a neighborhood of `a` within `[a, +∞)` if and only if it contains an interval `[a, u]`
1585
- with `a < u`. -/
1586
- lemma mem_nhds_within_Ici_iff_exists_Icc_subset [no_max_order α] [densely_ordered α]
1587
- {a : α} {s : set α} : s ∈ 𝓝[≥] a ↔ ∃u, a < u ∧ Icc a u ⊆ s :=
1588
- begin
1589
- rw mem_nhds_within_Ici_iff_exists_Ico_subset,
1590
- split,
1591
- { rintros ⟨u, au, as⟩,
1592
- rcases exists_between au with ⟨v, hv⟩,
1593
- exact ⟨v, hv.1 , λx hx, as ⟨hx.1 , lt_of_le_of_lt hx.2 hv.2 ⟩⟩ },
1594
- { rintros ⟨u, au, as⟩,
1595
- exact ⟨u, au, subset.trans Ico_subset_Icc_self as⟩ }
1596
- end
1597
-
1598
- /-- A set is a neighborhood of `a` within `(-∞, a]` if and only if it contains an interval `[l, a]`
1599
- with `l < a`. -/
1600
- lemma mem_nhds_within_Iic_iff_exists_Icc_subset [no_min_order α] [densely_ordered α]
1601
- {a : α} {s : set α} : s ∈ 𝓝[≤] a ↔ ∃l, l < a ∧ Icc l a ⊆ s :=
1602
- begin
1603
- rw mem_nhds_within_Iic_iff_exists_Ioc_subset,
1604
- split,
1605
- { rintros ⟨l, la, as⟩,
1606
- rcases exists_between la with ⟨v, hv⟩,
1607
- refine ⟨v, hv.2 , λx hx, as ⟨lt_of_lt_of_le hv.1 hx.1 , hx.2 ⟩⟩, },
1608
- { rintros ⟨l, la, as⟩,
1609
- exact ⟨l, la, subset.trans Ioc_subset_Icc_self as⟩ }
1610
- end
1611
-
1612
1584
end order_topology
1613
1585
1614
1586
end linear_order
0 commit comments