@@ -768,6 +768,21 @@ begin
768
768
exact ⟨hg.exists.some, hg.mono (λ y hy, is_glb.unique hy hg.exists.some_spec)⟩,
769
769
end
770
770
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
+
771
786
end linear_order
772
787
773
788
@[measurability]
@@ -879,7 +894,7 @@ end complete_linear_order
879
894
880
895
section conditionally_complete_linear_order
881
896
882
- variables [conditionally_complete_linear_order α] [second_countable_topology α] [order_topology α]
897
+ variables [conditionally_complete_linear_order α] [order_topology α] [second_countable_topology α]
883
898
884
899
lemma measurable_cSup {ι} {f : ι → δ → α} {s : set ι} (hs : s.countable)
885
900
(hf : ∀ i, measurable (f i)) (bdd : ∀ x, bdd_above ((λ i, f i x) '' s)) :
0 commit comments