Skip to content

Commit

Permalink
feat(measure_theory/borel_space): a preconnected set is measurable (#…
Browse files Browse the repository at this point in the history
…8044)

In a conditionally complete linear order equipped with the order topology and the corresponding borel σ-algebra, any preconnected set is measurable.
  • Loading branch information
ADedecker committed Jul 9, 2021
1 parent bcd61b1 commit a444e81
Show file tree
Hide file tree
Showing 3 changed files with 60 additions and 12 deletions.
27 changes: 26 additions & 1 deletion src/data/set/finite.lean
Expand Up @@ -3,7 +3,7 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Johannes Hölzl, Mario Carneiro
-/
import data.fintype.basic
import data.finset.sort

/-!
# Finite sets
Expand Down Expand Up @@ -810,3 +810,28 @@ lemma card_compl_eq_card_compl (h : fintype.card {x // p x} = fintype.card {x //
by simp only [card_subtype_compl, h]

end fintype

/--
If a set `s` does not contain any elements between any pair of elements `x, z ∈ s` with `x ≤ z`
(i.e if given `x, y, z ∈ s` such that `x ≤ y ≤ z`, then `y` is either `x` or `z`), then `s` is
finite.
-/
lemma set.finite_of_forall_between_eq_endpoints {α : Type*} [linear_order α] (s : set α)
(h : ∀ (x ∈ s) (y ∈ s) (z ∈ s), x ≤ y → y ≤ z → x = y ∨ y = z) :
set.finite s :=
begin
by_contra hinf,
change s.infinite at hinf,
rcases hinf.exists_subset_card_eq 3 with ⟨t, hts, ht⟩,
let f := t.order_iso_of_fin ht,
let x := f 0,
let y := f 1,
let z := f 2,
have := h x (hts x.2) y (hts y.2) z (hts z.2)
(f.monotone $ by dec_trivial) (f.monotone $ by dec_trivial),
have key₁ : (0 : fin 3) ≠ 1 := by dec_trivial,
have key₂ : (1 : fin 3) ≠ 2 := by dec_trivial,
cases this,
{ dsimp only [x, y] at this, exact key₁ (f.injective $ subtype.coe_injective this) },
{ dsimp only [y, z] at this, exact key₂ (f.injective $ subtype.coe_injective this) }
end
29 changes: 24 additions & 5 deletions src/measure_theory/borel_space.lean
Expand Up @@ -351,17 +351,36 @@ instance nhds_within_Iio_is_measurably_generated :
(𝓝[Iio b] a).is_measurably_generated :=
measurable_set_Iio.nhds_within_is_measurably_generated _

variables [second_countable_topology α]

@[measurability]
lemma measurable_set_lt' : measurable_set {p : α × α | p.1 < p.2} :=
lemma measurable_set_lt' [second_countable_topology α] : measurable_set {p : α × α | p.1 < p.2} :=
(is_open_lt continuous_fst continuous_snd).measurable_set

@[measurability]
lemma measurable_set_lt {f g : δ → α} (hf : measurable f) (hg : measurable g) :
measurable_set {a | f a < g a} :=
lemma measurable_set_lt [second_countable_topology α] {f g : δ → α} (hf : measurable f)
(hg : measurable g) : measurable_set {a | f a < g a} :=
hf.prod_mk hg measurable_set_lt'

lemma set.ord_connected.measurable_set (h : ord_connected s) : measurable_set s :=
begin
let u := ⋃ (x ∈ s) (y ∈ s), Ioo x y,
have huopen : is_open u := is_open_bUnion (λ x hx, is_open_bUnion (λ y hy, is_open_Ioo)),
have humeas : measurable_set u := huopen.measurable_set,
have hfinite : (s \ u).finite,
{ refine set.finite_of_forall_between_eq_endpoints (s \ u) (λ x hx y hy z hz hxy hyz, _),
by_contra h,
push_neg at h,
exact hy.2 (mem_bUnion_iff.mpr ⟨x, hx.1,
mem_bUnion_iff.mpr ⟨z, hz.1, lt_of_le_of_ne hxy h.1, lt_of_le_of_ne hyz h.2⟩⟩) },
have : u ⊆ s :=
bUnion_subset (λ x hx, bUnion_subset (λ y hy, Ioo_subset_Icc_self.trans (h.out hx hy))),
rw ← union_diff_cancel this,
exact humeas.union hfinite.measurable_set
end

lemma is_preconnected.measurable_set
(h : is_preconnected s) : measurable_set s :=
h.ord_connected.measurable_set

end linear_order

section linear_order
Expand Down
16 changes: 10 additions & 6 deletions src/topology/algebra/ordered/basic.lean
Expand Up @@ -679,6 +679,10 @@ lemma filter.tendsto.min {b : filter β} {a₁ a₂ : α} (hf : tendsto f b (
tendsto (λb, min (f b) (g b)) b (𝓝 (min a₁ a₂)) :=
(continuous_min.tendsto (a₁, a₂)).comp (hf.prod_mk_nhds hg)

lemma is_preconnected.ord_connected {s : set α} (h : is_preconnected s) :
ord_connected s :=
⟨λ x hx y hy, h.Icc_subset hx hy⟩

end linear_order

end order_closed_topology
Expand Down Expand Up @@ -2688,14 +2692,14 @@ end

lemma is_preconnected_interval : is_preconnected (interval a b) := is_preconnected_Icc

lemma set.ord_connected.is_preconnected {s : set α} (h : s.ord_connected) :
is_preconnected s :=
is_preconnected_of_forall_pair $ λ x y hx hy, ⟨interval x y, h.interval_subset hx hy,
left_mem_interval, right_mem_interval, is_preconnected_interval⟩

lemma is_preconnected_iff_ord_connected {s : set α} :
is_preconnected s ↔ ord_connected s :=
⟨λ h, ⟨λ x hx y hy, h.Icc_subset hx hy⟩, λ h, is_preconnected_of_forall_pair $ λ x y hx hy,
⟨interval x y, h.interval_subset hx hy, left_mem_interval, right_mem_interval,
is_preconnected_interval⟩⟩

alias is_preconnected_iff_ord_connected ↔
is_preconnected.ord_connected set.ord_connected.is_preconnected
⟨is_preconnected.ord_connected, set.ord_connected.is_preconnected⟩

lemma is_preconnected_Ici : is_preconnected (Ici a) := ord_connected_Ici.is_preconnected
lemma is_preconnected_Iic : is_preconnected (Iic a) := ord_connected_Iic.is_preconnected
Expand Down

0 comments on commit a444e81

Please sign in to comment.