|
| 1 | +/- |
| 2 | +Copyright (c) 2023 Yaël Dillies. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Yaël Dillies |
| 5 | +-/ |
| 6 | +import analysis.convex.function |
| 7 | +import data.set.intervals.proj_Icc |
| 8 | + |
| 9 | +/-! |
| 10 | +# Convexity of extension from intervals |
| 11 | +
|
| 12 | +This file proves that constantly extending monotone/antitone functions preserves their convexity. |
| 13 | +
|
| 14 | +## TODO |
| 15 | +
|
| 16 | +We could deduplicate the proofs if we had a typeclass stating that `segment 𝕜 x y = [x -[𝕜] y]` as |
| 17 | +`𝕜ᵒᵈ` respects it if `𝕜` does, while `𝕜ᵒᵈ` isn't a `linear_ordered_field` if `𝕜` is. |
| 18 | +-/ |
| 19 | + |
| 20 | +open set |
| 21 | + |
| 22 | +variables {𝕜 β : Type*} [linear_ordered_field 𝕜] [linear_ordered_add_comm_monoid β] [has_smul 𝕜 β] |
| 23 | + {s : set 𝕜} {f : 𝕜 → β} {z : 𝕜} |
| 24 | + |
| 25 | +/-- A convex set extended towards minus infinity is convex. -/ |
| 26 | +protected lemma convex.Ici_extend (hf : convex 𝕜 s) : |
| 27 | + convex 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} := |
| 28 | +by { rw convex_iff_ord_connected at ⊢ hf, exact hf.restrict.Ici_extend } |
| 29 | + |
| 30 | +/-- A convex set extended towards infinity is convex. -/ |
| 31 | +protected lemma convex.Iic_extend (hf : convex 𝕜 s) : |
| 32 | + convex 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} := |
| 33 | +by { rw convex_iff_ord_connected at ⊢ hf, exact hf.restrict.Iic_extend } |
| 34 | + |
| 35 | +/-- A convex monotone function extended constantly towards minus infinity is convex. -/ |
| 36 | +protected lemma convex_on.Ici_extend (hf : convex_on 𝕜 s f) (hf' : monotone_on f s) : |
| 37 | + convex_on 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} (Ici_extend $ restrict (Ici z) f) := |
| 38 | +begin |
| 39 | + refine ⟨hf.1.Ici_extend, λ x hx y hy a b ha hb hab, _⟩, |
| 40 | + dsimp [Ici_extend_apply] at ⊢ hx hy, |
| 41 | + refine (hf' (hf.1.ord_connected.uIcc_subset hx hy $ monotone.image_uIcc_subset (λ _ _, max_le_max |
| 42 | + le_rfl) $ mem_image_of_mem _ $ convex_uIcc _ _ left_mem_uIcc right_mem_uIcc ha hb hab) |
| 43 | + (hf.1 hx hy ha hb hab) _).trans (hf.2 hx hy ha hb hab), |
| 44 | + rw [smul_max ha z, smul_max hb z], |
| 45 | + refine le_trans _ max_add_add_le_max_add_max, |
| 46 | + rw [convex.combo_self hab, smul_eq_mul, smul_eq_mul], |
| 47 | +end |
| 48 | + |
| 49 | +/-- A convex antitone function extended constantly towards infinity is convex. -/ |
| 50 | +protected lemma convex_on.Iic_extend (hf : convex_on 𝕜 s f) (hf' : antitone_on f s) : |
| 51 | + convex_on 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} (Iic_extend $ restrict (Iic z) f) := |
| 52 | +begin |
| 53 | + refine ⟨hf.1.Iic_extend, λ x hx y hy a b ha hb hab, _⟩, |
| 54 | + dsimp [Iic_extend_apply] at ⊢ hx hy, |
| 55 | + refine (hf' (hf.1 hx hy ha hb hab) (hf.1.ord_connected.uIcc_subset hx hy $ |
| 56 | + monotone.image_uIcc_subset (λ _ _, min_le_min le_rfl) $ mem_image_of_mem _ $ |
| 57 | + convex_uIcc _ _ left_mem_uIcc right_mem_uIcc ha hb hab) _).trans (hf.2 hx hy ha hb hab), |
| 58 | + rw [smul_min ha z, smul_min hb z], |
| 59 | + refine min_add_min_le_min_add_add.trans _ , |
| 60 | + rw [convex.combo_self hab, smul_eq_mul, smul_eq_mul], |
| 61 | +end |
| 62 | + |
| 63 | +/-- A concave antitone function extended constantly minus towards infinity is concave. -/ |
| 64 | +protected lemma concave_on.Ici_extend (hf : concave_on 𝕜 s f) (hf' : antitone_on f s) : |
| 65 | + concave_on 𝕜 {x | Ici_extend (restrict (Ici z) (∈ s)) x} (Ici_extend $ restrict (Ici z) f) := |
| 66 | +hf.dual.Ici_extend hf'.dual_right |
| 67 | + |
| 68 | +/-- A concave monotone function extended constantly towards infinity is concave. -/ |
| 69 | +protected lemma concave_on.Iic_extend (hf : concave_on 𝕜 s f) (hf' : monotone_on f s) : |
| 70 | + concave_on 𝕜 {x | Iic_extend (restrict (Iic z) (∈ s)) x} (Iic_extend $ restrict (Iic z) f) := |
| 71 | +hf.dual.Iic_extend hf'.dual_right |
| 72 | + |
| 73 | +/-- A convex monotone function extended constantly towards minus infinity is convex. -/ |
| 74 | +protected lemma convex_on.Ici_extend_of_monotone (hf : convex_on 𝕜 univ f) (hf' : monotone f) : |
| 75 | + convex_on 𝕜 univ (Ici_extend $ restrict (Ici z) f) := |
| 76 | +hf.Ici_extend $ hf'.monotone_on _ |
| 77 | + |
| 78 | +/-- A convex antitone function extended constantly towards infinity is convex. -/ |
| 79 | +protected lemma convex_on.Iic_extend_of_antitone (hf : convex_on 𝕜 univ f) (hf' : antitone f) : |
| 80 | + convex_on 𝕜 univ (Iic_extend $ restrict (Iic z) f) := |
| 81 | +hf.Iic_extend $ hf'.antitone_on _ |
| 82 | + |
| 83 | +/-- A concave antitone function extended constantly minus towards infinity is concave. -/ |
| 84 | +protected lemma concave_on.Ici_extend_of_antitone (hf : concave_on 𝕜 univ f) (hf' : antitone f) : |
| 85 | + concave_on 𝕜 univ (Ici_extend $ restrict (Ici z) f) := |
| 86 | +hf.Ici_extend $ hf'.antitone_on _ |
| 87 | + |
| 88 | +/-- A concave monotone function extended constantly towards infinity is concave. -/ |
| 89 | +protected lemma concave_on.Iic_extend_of_monotone (hf : concave_on 𝕜 univ f) (hf' : monotone f) : |
| 90 | + concave_on 𝕜 univ (Iic_extend $ restrict (Iic z) f) := |
| 91 | +hf.Iic_extend $ hf'.monotone_on _ |
0 commit comments