Skip to content

Commit 930a6c1

Browse files
committed
feat: port Algebra.Order.Pointwise (#1533)
Co-authored-by: qawbecrdtey <qawbecrdtey@naver.com>
1 parent e609eeb commit 930a6c1

File tree

2 files changed

+297
-0
lines changed

2 files changed

+297
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -115,6 +115,7 @@ import Mathlib.Algebra.Order.Monoid.WithTop
115115
import Mathlib.Algebra.Order.Monoid.WithZero.Basic
116116
import Mathlib.Algebra.Order.Monoid.WithZero.Defs
117117
import Mathlib.Algebra.Order.Pi
118+
import Mathlib.Algebra.Order.Pointwise
118119
import Mathlib.Algebra.Order.Positive.Field
119120
import Mathlib.Algebra.Order.Positive.Ring
120121
import Mathlib.Algebra.Order.Ring.Abs
Lines changed: 296 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,296 @@
1+
/-
2+
Copyright (c) 2021 Alex J. Best. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Alex J. Best, Yaël Dillies
5+
6+
! This file was ported from Lean 3 source module algebra.order.pointwise
7+
! leanprover-community/mathlib commit 9003f28797c0664a49e4179487267c494477d853
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Algebra.Bounds
12+
import Mathlib.Algebra.Order.Field.Basic -- Porting note: `LinearOrderedField`, etc
13+
import Mathlib.Data.Set.Pointwise.SMul
14+
15+
/-!
16+
# Pointwise operations on ordered algebraic objects
17+
18+
This file contains lemmas about the effect of pointwise operations on sets with an order structure.
19+
20+
## TODO
21+
22+
`Sup (s • t) = Sup s • Sup t` and `Inf (s • t) = Inf s • Inf t` hold as well but
23+
`CovariantClass` is currently not polymorphic enough to state it.
24+
-/
25+
26+
27+
open Function Set
28+
29+
open Pointwise
30+
31+
variable {α : Type _}
32+
33+
-- Porting note : Swapped the place of `CompleteLattice` and `ConditionallyCompleteLattice`
34+
-- due to simpNF problem between `supₛ_xx` `csupₛ_xx`.
35+
36+
section CompleteLattice
37+
38+
variable [CompleteLattice α]
39+
40+
section One
41+
42+
variable [One α]
43+
44+
@[to_additive (attr := simp)]
45+
theorem supₛ_one : supₛ (1 : Set α) = 1 :=
46+
supₛ_singleton
47+
#align Sup_zero supₛ_zero
48+
#align Sup_one supₛ_one
49+
50+
@[to_additive (attr := simp)]
51+
theorem infₛ_one : infₛ (1 : Set α) = 1 :=
52+
infₛ_singleton
53+
#align Inf_zero infₛ_zero
54+
#align Inf_one infₛ_one
55+
56+
end One
57+
58+
section Group
59+
60+
variable [Group α] [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)]
61+
(s t : Set α)
62+
63+
@[to_additive]
64+
theorem supₛ_inv (s : Set α) : supₛ s⁻¹ = (infₛ s)⁻¹ := by
65+
rw [← image_inv, supₛ_image]
66+
exact ((OrderIso.inv α).map_infₛ _).symm
67+
#align Sup_inv supₛ_inv
68+
#align Sup_neg supₛ_neg
69+
70+
@[to_additive]
71+
theorem infₛ_inv (s : Set α) : infₛ s⁻¹ = (supₛ s)⁻¹ := by
72+
rw [← image_inv, infₛ_image]
73+
exact ((OrderIso.inv α).map_supₛ _).symm
74+
#align Inf_inv infₛ_inv
75+
#align Inf_neg infₛ_neg
76+
77+
@[to_additive]
78+
theorem supₛ_mul : supₛ (s * t) = supₛ s * supₛ t :=
79+
(supₛ_image2_eq_supₛ_supₛ fun _ => (OrderIso.mulRight _).to_galoisConnection) fun _ =>
80+
(OrderIso.mulLeft _).to_galoisConnection
81+
#align Sup_mul supₛ_mul
82+
#align Sup_add supₛ_add
83+
84+
@[to_additive]
85+
theorem infₛ_mul : infₛ (s * t) = infₛ s * infₛ t :=
86+
(infₛ_image2_eq_infₛ_infₛ fun _ => (OrderIso.mulRight _).symm.to_galoisConnection) fun _ =>
87+
(OrderIso.mulLeft _).symm.to_galoisConnection
88+
#align Inf_mul infₛ_mul
89+
#align Inf_add infₛ_add
90+
91+
@[to_additive]
92+
theorem supₛ_div : supₛ (s / t) = supₛ s / infₛ t := by simp_rw [div_eq_mul_inv, supₛ_mul, supₛ_inv]
93+
#align Sup_div supₛ_div
94+
#align Sup_sub supₛ_sub
95+
96+
@[to_additive]
97+
theorem infₛ_div : infₛ (s / t) = infₛ s / supₛ t := by simp_rw [div_eq_mul_inv, infₛ_mul, infₛ_inv]
98+
#align Inf_div infₛ_div
99+
#align Inf_sub infₛ_sub
100+
101+
end Group
102+
103+
end CompleteLattice
104+
105+
section ConditionallyCompleteLattice
106+
107+
variable [ConditionallyCompleteLattice α]
108+
109+
section One
110+
111+
variable [One α]
112+
113+
@[to_additive (attr := simp)]
114+
theorem csupₛ_one : supₛ (1 : Set α) = 1 :=
115+
csupₛ_singleton _
116+
#align cSup_zero csupₛ_zero
117+
#align cSup_one csupₛ_one
118+
119+
@[to_additive (attr := simp)]
120+
theorem cinfₛ_one : infₛ (1 : Set α) = 1 :=
121+
cinfₛ_singleton _
122+
#align cInf_zero cinfₛ_zero
123+
#align cInf_one cinfₛ_one
124+
125+
end One
126+
127+
section Group
128+
129+
variable [Group α] [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)]
130+
{s t : Set α}
131+
132+
@[to_additive]
133+
theorem csupₛ_inv (hs₀ : s.Nonempty) (hs₁ : BddBelow s) : supₛ s⁻¹ = (infₛ s)⁻¹ := by
134+
rw [← image_inv]
135+
exact ((OrderIso.inv α).map_cinfₛ' hs₀ hs₁).symm
136+
#align cSup_inv csupₛ_inv
137+
#align cSup_neg csupₛ_neg
138+
139+
@[to_additive]
140+
theorem cinfₛ_inv (hs₀ : s.Nonempty) (hs₁ : BddAbove s) : infₛ s⁻¹ = (supₛ s)⁻¹ := by
141+
rw [← image_inv]
142+
exact ((OrderIso.inv α).map_csupₛ' hs₀ hs₁).symm
143+
#align cInf_inv cinfₛ_inv
144+
#align cInf_neg cinfₛ_neg
145+
146+
@[to_additive]
147+
theorem csupₛ_mul (hs₀ : s.Nonempty) (hs₁ : BddAbove s) (ht₀ : t.Nonempty) (ht₁ : BddAbove t) :
148+
supₛ (s * t) = supₛ s * supₛ t :=
149+
csupₛ_image2_eq_csupₛ_csupₛ (fun _ => (OrderIso.mulRight _).to_galoisConnection)
150+
(fun _ => (OrderIso.mulLeft _).to_galoisConnection) hs₀ hs₁ ht₀ ht₁
151+
#align cSup_mul csupₛ_mul
152+
#align cSup_add csupₛ_add
153+
154+
@[to_additive]
155+
theorem cinfₛ_mul (hs₀ : s.Nonempty) (hs₁ : BddBelow s) (ht₀ : t.Nonempty) (ht₁ : BddBelow t) :
156+
infₛ (s * t) = infₛ s * infₛ t :=
157+
cinfₛ_image2_eq_cinfₛ_cinfₛ (fun _ => (OrderIso.mulRight _).symm.to_galoisConnection)
158+
(fun _ => (OrderIso.mulLeft _).symm.to_galoisConnection) hs₀ hs₁ ht₀ ht₁
159+
#align cInf_mul cinfₛ_mul
160+
#align cInf_add cinfₛ_add
161+
162+
@[to_additive]
163+
theorem csupₛ_div (hs₀ : s.Nonempty) (hs₁ : BddAbove s) (ht₀ : t.Nonempty) (ht₁ : BddBelow t) :
164+
supₛ (s / t) = supₛ s / infₛ t := by
165+
rw [div_eq_mul_inv, csupₛ_mul hs₀ hs₁ ht₀.inv ht₁.inv, csupₛ_inv ht₀ ht₁, div_eq_mul_inv]
166+
#align cSup_div csupₛ_div
167+
#align cSup_sub csupₛ_sub
168+
169+
@[to_additive]
170+
theorem cinfₛ_div (hs₀ : s.Nonempty) (hs₁ : BddBelow s) (ht₀ : t.Nonempty) (ht₁ : BddAbove t) :
171+
infₛ (s / t) = infₛ s / supₛ t := by
172+
rw [div_eq_mul_inv, cinfₛ_mul hs₀ hs₁ ht₀.inv ht₁.inv, cinfₛ_inv ht₀ ht₁, div_eq_mul_inv]
173+
#align cInf_div cinfₛ_div
174+
#align cInf_sub cinfₛ_sub
175+
176+
end Group
177+
178+
end ConditionallyCompleteLattice
179+
180+
namespace LinearOrderedField
181+
182+
variable {K : Type _} [LinearOrderedField K] {a b r : K} (hr : 0 < r)
183+
184+
open Set
185+
186+
-- Porting note: Removing `include hr`
187+
188+
theorem smul_Ioo : r • Ioo a b = Ioo (r • a) (r • b) := by
189+
ext x
190+
simp only [mem_smul_set, smul_eq_mul, mem_Ioo]
191+
constructor
192+
· rintro ⟨a, ⟨a_h_left_left, a_h_left_right⟩, rfl⟩
193+
constructor
194+
exact (mul_lt_mul_left hr).mpr a_h_left_left
195+
exact (mul_lt_mul_left hr).mpr a_h_left_right
196+
· rintro ⟨a_left, a_right⟩
197+
use x / r
198+
refine' ⟨⟨(lt_div_iff' hr).mpr a_left, (div_lt_iff' hr).mpr a_right⟩, _⟩
199+
rw [mul_div_cancel' _ (ne_of_gt hr)]
200+
#align linear_ordered_field.smul_Ioo LinearOrderedField.smul_Ioo
201+
202+
theorem smul_Icc : r • Icc a b = Icc (r • a) (r • b) := by
203+
ext x
204+
simp only [mem_smul_set, smul_eq_mul, mem_Icc]
205+
constructor
206+
· rintro ⟨a, ⟨a_h_left_left, a_h_left_right⟩, rfl⟩
207+
constructor
208+
exact (mul_le_mul_left hr).mpr a_h_left_left
209+
exact (mul_le_mul_left hr).mpr a_h_left_right
210+
· rintro ⟨a_left, a_right⟩
211+
use x / r
212+
refine' ⟨⟨(le_div_iff' hr).mpr a_left, (div_le_iff' hr).mpr a_right⟩, _⟩
213+
rw [mul_div_cancel' _ (ne_of_gt hr)]
214+
#align linear_ordered_field.smul_Icc LinearOrderedField.smul_Icc
215+
216+
theorem smul_Ico : r • Ico a b = Ico (r • a) (r • b) := by
217+
ext x
218+
simp only [mem_smul_set, smul_eq_mul, mem_Ico]
219+
constructor
220+
· rintro ⟨a, ⟨a_h_left_left, a_h_left_right⟩, rfl⟩
221+
constructor
222+
exact (mul_le_mul_left hr).mpr a_h_left_left
223+
exact (mul_lt_mul_left hr).mpr a_h_left_right
224+
· rintro ⟨a_left, a_right⟩
225+
use x / r
226+
refine' ⟨⟨(le_div_iff' hr).mpr a_left, (div_lt_iff' hr).mpr a_right⟩, _⟩
227+
rw [mul_div_cancel' _ (ne_of_gt hr)]
228+
#align linear_ordered_field.smul_Ico LinearOrderedField.smul_Ico
229+
230+
theorem smul_Ioc : r • Ioc a b = Ioc (r • a) (r • b) := by
231+
ext x
232+
simp only [mem_smul_set, smul_eq_mul, mem_Ioc]
233+
constructor
234+
· rintro ⟨a, ⟨a_h_left_left, a_h_left_right⟩, rfl⟩
235+
constructor
236+
exact (mul_lt_mul_left hr).mpr a_h_left_left
237+
exact (mul_le_mul_left hr).mpr a_h_left_right
238+
· rintro ⟨a_left, a_right⟩
239+
use x / r
240+
refine' ⟨⟨(lt_div_iff' hr).mpr a_left, (div_le_iff' hr).mpr a_right⟩, _⟩
241+
rw [mul_div_cancel' _ (ne_of_gt hr)]
242+
#align linear_ordered_field.smul_Ioc LinearOrderedField.smul_Ioc
243+
244+
theorem smul_Ioi : r • Ioi a = Ioi (r • a) := by
245+
ext x
246+
simp only [mem_smul_set, smul_eq_mul, mem_Ioi]
247+
constructor
248+
· rintro ⟨a_w, a_h_left, rfl⟩
249+
exact (mul_lt_mul_left hr).mpr a_h_left
250+
· rintro h
251+
use x / r
252+
constructor
253+
exact (lt_div_iff' hr).mpr h
254+
exact mul_div_cancel' _ (ne_of_gt hr)
255+
#align linear_ordered_field.smul_Ioi LinearOrderedField.smul_Ioi
256+
257+
theorem smul_Iio : r • Iio a = Iio (r • a) := by
258+
ext x
259+
simp only [mem_smul_set, smul_eq_mul, mem_Iio]
260+
constructor
261+
· rintro ⟨a_w, a_h_left, rfl⟩
262+
exact (mul_lt_mul_left hr).mpr a_h_left
263+
· rintro h
264+
use x / r
265+
constructor
266+
exact (div_lt_iff' hr).mpr h
267+
exact mul_div_cancel' _ (ne_of_gt hr)
268+
#align linear_ordered_field.smul_Iio LinearOrderedField.smul_Iio
269+
270+
theorem smul_Ici : r • Ici a = Ici (r • a) := by
271+
ext x
272+
simp only [mem_smul_set, smul_eq_mul, mem_Ioi]
273+
constructor
274+
· rintro ⟨a_w, a_h_left, rfl⟩
275+
exact (mul_le_mul_left hr).mpr a_h_left
276+
· rintro h
277+
use x / r
278+
constructor
279+
exact (le_div_iff' hr).mpr h
280+
exact mul_div_cancel' _ (ne_of_gt hr)
281+
#align linear_ordered_field.smul_Ici LinearOrderedField.smul_Ici
282+
283+
theorem smul_Iic : r • Iic a = Iic (r • a) := by
284+
ext x
285+
simp only [mem_smul_set, smul_eq_mul, mem_Iio]
286+
constructor
287+
· rintro ⟨a_w, a_h_left, rfl⟩
288+
exact (mul_le_mul_left hr).mpr a_h_left
289+
· rintro h
290+
use x / r
291+
constructor
292+
exact (div_le_iff' hr).mpr h
293+
exact mul_div_cancel' _ (ne_of_gt hr)
294+
#align linear_ordered_field.smul_Iic LinearOrderedField.smul_Iic
295+
296+
end LinearOrderedField

0 commit comments

Comments
 (0)