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

Commit 84b450d

Browse files
feat(topology): path connected spaces (#3627)
From the sphere eversion project. Co-authored-by: sgouezel <sebastien.gouezel@univ-rennes1.fr>
1 parent f4b2790 commit 84b450d

File tree

4 files changed

+643
-3
lines changed

4 files changed

+643
-3
lines changed

src/data/set/intervals/basic.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -58,6 +58,22 @@ def Ici (a : α) := {x | a ≤ x}
5858
/-- Left-open right-infinite interval -/
5959
def Ioi (a : α) := {x | a < x}
6060

61+
lemma Ioo_def (a b : α) : {x | a < x ∧ x < b} = Ioo a b := rfl
62+
63+
lemma Ico_def (a b : α) : {x | a ≤ x ∧ x < b} = Ico a b := rfl
64+
65+
lemma Iio_def (a : α) : {x | x < a} = Iio a := rfl
66+
67+
lemma Icc_def (a b : α) : {x | a ≤ x ∧ x ≤ b} = Icc a b := rfl
68+
69+
lemma Iic_def (b : α) : {x | x ≤ b} = Iic b := rfl
70+
71+
lemma Ioc_def (a b : α) : {x | a < x ∧ x ≤ b} = Ioc a b := rfl
72+
73+
lemma Ici_def (a : α) : {x | a ≤ x} = Ici a := rfl
74+
75+
lemma Ioi_def (a : α) : {x | a < x} = Ioi a := rfl
76+
6177
@[simp] lemma mem_Ioo : x ∈ Ioo a b ↔ a < x ∧ x < b := iff.rfl
6278
@[simp] lemma mem_Ico : x ∈ Ico a b ↔ a ≤ x ∧ x < b := iff.rfl
6379
@[simp] lemma mem_Iio : x ∈ Iio b ↔ x < b := iff.rfl

src/geometry/manifold/real_instances.lean

Lines changed: 0 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -378,7 +378,4 @@ local attribute [instance] fact_zero_lt_one
378378
instance : charted_space (euclidean_half_space 1) (Icc (0 : ℝ) 1) := by apply_instance
379379
instance : smooth_manifold_with_corners (𝓡∂ 1) (Icc (0 : ℝ) 1) := by apply_instance
380380

381-
instance : has_zero (Icc (0 : ℝ) 1) := ⟨⟨(0 : ℝ), ⟨le_refl _, zero_le_one⟩⟩⟩
382-
instance : has_one (Icc (0 : ℝ) 1) := ⟨⟨(1 : ℝ), ⟨zero_le_one, le_refl _⟩⟩⟩
383-
384381
end

src/order/filter/bases.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -244,6 +244,23 @@ by erw [(filter_basis.of_sets s).generate, ← (has_basis_generate s).filter_eq]
244244
lemma of_sets_filter_eq_generate (s : set (set α)) : (filter_basis.of_sets s).filter = generate s :=
245245
by rw [← (filter_basis.of_sets s).generate, generate_eq_generate_inter s] ; refl
246246

247+
lemma has_basis.to_has_basis (hl : l.has_basis p s) (h : ∀ i, p i → ∃ i', p' i' ∧ s' i' ⊆ s i)
248+
(h' : ∀ i', p' i' → ∃ i, p i ∧ s i ⊆ s' i') : l.has_basis p' s' :=
249+
begin
250+
constructor,
251+
intro t,
252+
rw hl.mem_iff,
253+
split,
254+
{ rintros ⟨i, pi, hi⟩,
255+
rcases h i pi with ⟨i', pi', hi'⟩,
256+
use [i', pi'],
257+
tauto },
258+
{ rintros ⟨i', pi', hi'⟩,
259+
rcases h' i' pi' with ⟨i, pi, hi⟩,
260+
use [i, pi],
261+
tauto },
262+
end
263+
247264
lemma has_basis.eventually_iff (hl : l.has_basis p s) {q : α → Prop} :
248265
(∀ᶠ x in l, q x) ↔ ∃ i, p i ∧ ∀ ⦃x⦄, x ∈ s i → q x :=
249266
by simpa using hl.mem_iff

0 commit comments

Comments
 (0)