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

Commit 51cc43e

Browse files
committed
feat(measure_theory/borel_space): a monotone function is measurable (#8045)
1 parent 8e5af43 commit 51cc43e

File tree

1 file changed

+16
-1
lines changed

1 file changed

+16
-1
lines changed

src/measure_theory/borel_space.lean

Lines changed: 16 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -768,6 +768,21 @@ begin
768768
exact ⟨hg.exists.some, hg.mono (λ y hy, is_glb.unique hy hg.exists.some_spec)⟩,
769769
end
770770

771+
lemma measurable_of_monotone [linear_order β] [order_topology β] {f : β → α} (hf : monotone f) :
772+
measurable f :=
773+
suffices h : ∀ x, ord_connected (f ⁻¹' Ioi x),
774+
from measurable_of_Ioi (λ x, (h x).measurable_set),
775+
λ x, ord_connected_def.mpr (λ a ha b hb c hc, lt_of_lt_of_le ha (hf hc.1))
776+
777+
alias measurable_of_monotone ← monotone.measurable
778+
779+
lemma measurable_of_antimono [linear_order β] [order_topology β] {f : β → α}
780+
(hf : ∀ ⦃x y : β⦄, x ≤ y → f y ≤ f x) :
781+
measurable f :=
782+
suffices h : ∀ x, ord_connected (f ⁻¹' Ioi x),
783+
from measurable_of_Ioi (λ x, (h x).measurable_set),
784+
λ x, ord_connected_def.mpr (λ a ha b hb c hc, lt_of_lt_of_le hb (hf hc.2))
785+
771786
end linear_order
772787

773788
@[measurability]
@@ -879,7 +894,7 @@ end complete_linear_order
879894

880895
section conditionally_complete_linear_order
881896

882-
variables [conditionally_complete_linear_order α] [second_countable_topology α] [order_topology α]
897+
variables [conditionally_complete_linear_order α] [order_topology α] [second_countable_topology α]
883898

884899
lemma measurable_cSup {ι} {f : ι → δ → α} {s : set ι} (hs : s.countable)
885900
(hf : ∀ i, measurable (f i)) (bdd : ∀ x, bdd_above ((λ i, f i x) '' s)) :

0 commit comments

Comments
 (0)