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

Commit 3ba1516

Browse files
committed
feat(analysis/convex/proj_Icc): Extending convex functions (#18797)
Constantly extending monotone/antitone functions preserves their convexity.
1 parent 63721b2 commit 3ba1516

File tree

7 files changed

+217
-3
lines changed

7 files changed

+217
-3
lines changed

src/algebra/order/module.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -209,6 +209,18 @@ lemma bdd_above.smul_of_nonpos (hc : c ≤ 0) (hs : bdd_above s) : bdd_below (c
209209

210210
end ordered_ring
211211

212+
section linear_ordered_ring
213+
variables [linear_ordered_ring k] [linear_ordered_add_comm_group M] [module k M] [ordered_smul k M]
214+
{a : k}
215+
216+
lemma smul_max_of_nonpos (ha : a ≤ 0) (b₁ b₂ : M) : a • max b₁ b₂ = min (a • b₁) (a • b₂) :=
217+
(antitone_smul_left ha : antitone (_ : M → M)).map_max
218+
219+
lemma smul_min_of_nonpos (ha : a ≤ 0) (b₁ b₂ : M) : a • min b₁ b₂ = max (a • b₁) (a • b₂) :=
220+
(antitone_smul_left ha : antitone (_ : M → M)).map_min
221+
222+
end linear_ordered_ring
223+
212224
section linear_ordered_field
213225
variables [linear_ordered_field k] [ordered_add_comm_group M] [module k M] [ordered_smul k M]
214226
{s : set M} {c : k}

src/algebra/order/monoid/lemmas.lean

Lines changed: 17 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -244,6 +244,23 @@ variables [linear_order α] {a b c d : α} [covariant_class α α (*) (<)]
244244
@[to_additive] lemma min_le_max_of_mul_le_mul (h : a * b ≤ c * d) : min a b ≤ max c d :=
245245
by { simp_rw [min_le_iff, le_max_iff], contrapose! h, exact mul_lt_mul_of_lt_of_lt h.1.1 h.2.2 }
246246

247+
end linear_order
248+
249+
section linear_order
250+
variables [linear_order α] [covariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (≤)]
251+
{a b c d : α}
252+
253+
@[to_additive max_add_add_le_max_add_max] lemma max_mul_mul_le_max_mul_max' :
254+
max (a * b) (c * d) ≤ max a c * max b d :=
255+
max_le (mul_le_mul' (le_max_left _ _) $ le_max_left _ _) $
256+
mul_le_mul' (le_max_right _ _) $ le_max_right _ _
257+
258+
--TODO: Also missing `min_mul_min_le_min_mul_mul`
259+
@[to_additive min_add_min_le_min_add_add] lemma min_mul_min_le_min_mul_mul' :
260+
min a c * min b d ≤ min (a * b) (c * d) :=
261+
le_min (mul_le_mul' (min_le_left _ _) $ min_le_left _ _) $
262+
mul_le_mul' (min_le_right _ _) $ min_le_right _ _
263+
247264
end linear_order
248265
end has_mul
249266

src/algebra/order/smul.lean

Lines changed: 13 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -170,12 +170,23 @@ ordered_smul.mk'' $ λ n hn, begin
170170
{ cases (int.neg_succ_not_pos _).1 hn }
171171
end
172172

173+
section linear_ordered_semiring
174+
variables [linear_ordered_semiring R] [linear_ordered_add_comm_monoid M] [smul_with_zero R M]
175+
[ordered_smul R M] {a : R}
176+
173177
-- TODO: `linear_ordered_field M → ordered_smul ℚ M`
174178

175-
instance linear_ordered_semiring.to_ordered_smul {R : Type*} [linear_ordered_semiring R] :
176-
ordered_smul R R :=
179+
instance linear_ordered_semiring.to_ordered_smul : ordered_smul R R :=
177180
ordered_smul.mk'' $ λ c, strict_mono_mul_left_of_pos
178181

182+
lemma smul_max (ha : 0 ≤ a) (b₁ b₂ : M) : a • max b₁ b₂ = max (a • b₁) (a • b₂) :=
183+
(monotone_smul_left ha : monotone (_ : M → M)).map_max
184+
185+
lemma smul_min (ha : 0 ≤ a) (b₁ b₂ : M) : a • min b₁ b₂ = min (a • b₁) (a • b₂) :=
186+
(monotone_smul_left ha : monotone (_ : M → M)).map_min
187+
188+
end linear_ordered_semiring
189+
179190
section linear_ordered_semifield
180191
variables [linear_ordered_semifield 𝕜] [ordered_add_comm_monoid M] [ordered_add_comm_monoid N]
181192
[mul_action_with_zero 𝕜 M] [mul_action_with_zero 𝕜 N]

src/analysis/convex/proj_Icc.lean

Lines changed: 91 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,91 @@
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 _

src/data/set/intervals/basic.lean

Lines changed: 20 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1164,6 +1164,26 @@ begin
11641164
le_of_lt h₂, le_of_lt h₁] },
11651165
end
11661166

1167+
section lattice
1168+
variables [lattice β] {f : α → β}
1169+
1170+
lemma _root_.monotone_on.image_Icc_subset (hf : monotone_on f (Icc a b)) :
1171+
f '' Icc a b ⊆ Icc (f a) (f b) :=
1172+
image_subset_iff.2 $ λ c hc,
1173+
⟨hf (left_mem_Icc.2 $ hc.1.trans hc.2) hc hc.1, hf hc (right_mem_Icc.2 $ hc.1.trans hc.2) hc.2
1174+
1175+
lemma _root_.antitone_on.image_Icc_subset (hf : antitone_on f (Icc a b)) :
1176+
f '' Icc a b ⊆ Icc (f b) (f a) :=
1177+
image_subset_iff.2 $ λ c hc,
1178+
⟨hf hc (right_mem_Icc.2 $ hc.1.trans hc.2) hc.2, hf (left_mem_Icc.2 $ hc.1.trans hc.2) hc hc.1
1179+
1180+
lemma _root_.monotone.image_Icc_subset (hf : monotone f) : f '' Icc a b ⊆ Icc (f a) (f b) :=
1181+
(hf.monotone_on _).image_Icc_subset
1182+
1183+
lemma _root_.antitone.image_Icc_subset (hf : antitone f) : f '' Icc a b ⊆ Icc (f b) (f a) :=
1184+
(hf.antitone_on _).image_Icc_subset
1185+
1186+
end lattice
11671187
end linear_order
11681188

11691189
section lattice

src/data/set/intervals/unordered_interval.lean

Lines changed: 24 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -129,7 +129,30 @@ by simpa only [uIcc_comm] using uIcc_injective_right a
129129
end distrib_lattice
130130

131131
section linear_order
132-
variables [linear_order α] [linear_order β] {f : α → β} {s : set α} {a a₁ a₂ b b₁ b₂ c d x : α}
132+
variables [linear_order α]
133+
134+
section lattice
135+
variables [lattice β] {f : α → β} {s : set α} {a b : α}
136+
137+
lemma _root_.monotone_on.image_uIcc_subset (hf : monotone_on f (uIcc a b)) :
138+
f '' uIcc a b ⊆ uIcc (f a) (f b) :=
139+
hf.image_Icc_subset.trans $
140+
by rw [hf.map_sup left_mem_uIcc right_mem_uIcc, hf.map_inf left_mem_uIcc right_mem_uIcc, uIcc]
141+
142+
lemma _root_.antitone_on.image_uIcc_subset (hf : antitone_on f (uIcc a b)) :
143+
f '' uIcc a b ⊆ uIcc (f a) (f b) :=
144+
hf.image_Icc_subset.trans $
145+
by rw [hf.map_sup left_mem_uIcc right_mem_uIcc, hf.map_inf left_mem_uIcc right_mem_uIcc, uIcc]
146+
147+
lemma _root_.monotone.image_uIcc_subset (hf : monotone f) : f '' uIcc a b ⊆ uIcc (f a) (f b) :=
148+
(hf.monotone_on _).image_uIcc_subset
149+
150+
lemma _root_.antitone.image_uIcc_subset (hf : antitone f) : f '' uIcc a b ⊆ uIcc (f a) (f b) :=
151+
(hf.antitone_on _).image_uIcc_subset
152+
153+
end lattice
154+
155+
variables [linear_order β] {f : α → β} {s : set α} {a a₁ a₂ b b₁ b₂ c d x : α}
133156

134157
lemma Icc_min_max : Icc (min a b) (max a b) = [a, b] := rfl
135158

src/order/lattice.lean

Lines changed: 40 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -846,6 +846,7 @@ hf.dual.map_sup _ _
846846
end monotone
847847

848848
namespace monotone_on
849+
variables {f : α → β} {s : set α} {x y : α}
849850

850851
/-- Pointwise supremum of two monotone functions is a monotone function. -/
851852
protected lemma sup [preorder α] [semilattice_sup β] {f g : α → β} {s : set α}
@@ -867,6 +868,25 @@ protected lemma min [preorder α] [linear_order β] {f g : α → β} {s : set
867868
(hf : monotone_on f s) (hg : monotone_on g s) : monotone_on (λ x, min (f x) (g x)) s :=
868869
hf.inf hg
869870

871+
lemma of_map_inf [semilattice_inf α] [semilattice_inf β]
872+
(h : ∀ (x ∈ s) (y ∈ s), f (x ⊓ y) = f x ⊓ f y) : monotone_on f s :=
873+
λ x hx y hy hxy, inf_eq_left.1 $ by rw [←h _ hx _ hy, inf_eq_left.2 hxy]
874+
875+
lemma of_map_sup [semilattice_sup α] [semilattice_sup β]
876+
(h : ∀ (x ∈ s) (y ∈ s), f (x ⊔ y) = f x ⊔ f y) : monotone_on f s :=
877+
(@of_map_inf αᵒᵈ βᵒᵈ _ _ _ _ h).dual
878+
879+
variables [linear_order α]
880+
881+
lemma map_sup [semilattice_sup β] (hf : monotone_on f s) (hx : x ∈ s) (hy : y ∈ s) :
882+
f (x ⊔ y) = f x ⊔ f y :=
883+
by cases le_total x y; have := hf _ _ h;
884+
assumption <|> simp only [h, this, sup_of_le_left, sup_of_le_right]
885+
886+
lemma map_inf [semilattice_inf β] (hf : monotone_on f s) (hx : x ∈ s) (hy : y ∈ s) :
887+
f (x ⊓ y) = f x ⊓ f y :=
888+
hf.dual.map_sup hx hy
889+
870890
end monotone_on
871891

872892
namespace antitone
@@ -912,6 +932,7 @@ hf.dual_right.map_inf x y
912932
end antitone
913933

914934
namespace antitone_on
935+
variables {f : α → β} {s : set α} {x y : α}
915936

916937
/-- Pointwise supremum of two antitone functions is a antitone function. -/
917938
protected lemma sup [preorder α] [semilattice_sup β] {f g : α → β} {s : set α}
@@ -933,6 +954,25 @@ protected lemma min [preorder α] [linear_order β] {f g : α → β} {s : set
933954
(hf : antitone_on f s) (hg : antitone_on g s) : antitone_on (λ x, min (f x) (g x)) s :=
934955
hf.inf hg
935956

957+
lemma of_map_inf [semilattice_inf α] [semilattice_sup β]
958+
(h : ∀ (x ∈ s) (y ∈ s), f (x ⊓ y) = f x ⊔ f y) : antitone_on f s :=
959+
λ x hx y hy hxy, sup_eq_left.1 $ by rw [←h _ hx _ hy, inf_eq_left.2 hxy]
960+
961+
lemma of_map_sup [semilattice_sup α] [semilattice_inf β]
962+
(h : ∀ (x ∈ s) (y ∈ s), f (x ⊔ y) = f x ⊓ f y) : antitone_on f s :=
963+
(@of_map_inf αᵒᵈ βᵒᵈ _ _ _ _ h).dual
964+
965+
variables [linear_order α]
966+
967+
lemma map_sup [semilattice_inf β] (hf : antitone_on f s) (hx : x ∈ s) (hy : y ∈ s) :
968+
f (x ⊔ y) = f x ⊓ f y :=
969+
by cases le_total x y; have := hf _ _ h; assumption <|>
970+
simp only [h, this, sup_of_le_left, sup_of_le_right, inf_of_le_left, inf_of_le_right]
971+
972+
lemma map_inf [semilattice_sup β] (hf : antitone_on f s) (hx : x ∈ s) (hy : y ∈ s) :
973+
f (x ⊓ y) = f x ⊔ f y :=
974+
hf.dual.map_sup hx hy
975+
936976
end antitone_on
937977

938978
/-!

0 commit comments

Comments
 (0)