Skip to content

Commit 3675299

Browse files
mcdollChrisHughes24
andcommitted
feat: port Order.LatticeIntervals (#1070)
Lots of renames, but easy. I've added names for all instances Co-authored-by: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com>
1 parent c5e582f commit 3675299

File tree

2 files changed

+184
-0
lines changed

2 files changed

+184
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -335,6 +335,7 @@ import Mathlib.Order.Hom.Set
335335
import Mathlib.Order.InitialSeg
336336
import Mathlib.Order.Iterate
337337
import Mathlib.Order.Lattice
338+
import Mathlib.Order.LatticeIntervals
338339
import Mathlib.Order.Max
339340
import Mathlib.Order.MinMax
340341
import Mathlib.Order.Monotone.Basic
Lines changed: 183 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,183 @@
1+
/-
2+
Copyright (c) 2020 Aaron Anderson. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Aaron Anderson
5+
! This file was ported from Lean 3 source module order.lattice_intervals
6+
! leanprover-community/mathlib commit d012cd09a9b256d870751284dd6a29882b0be105
7+
! Please do not edit these lines, except to modify the commit id
8+
! if you have ported upstream changes.
9+
-/
10+
import Mathlib.Order.Bounds.Basic
11+
12+
/-!
13+
# Intervals in Lattices
14+
15+
In this file, we provide instances of lattice structures on intervals within lattices.
16+
Some of them depend on the order of the endpoints of the interval, and thus are not made
17+
global instances. These are probably not all of the lattice instances that could be placed on these
18+
intervals, but more can be added easily along the same lines when needed.
19+
20+
## Main definitions
21+
22+
In the following, `*` can represent either `c`, `o`, or `i`.
23+
* `Set.Ic*.order_bot`
24+
* `Set.Ii*.semillatice_inf`
25+
* `Set.I*c.order_top`
26+
* `Set.I*c.semillatice_inf`
27+
* `Set.I**.lattice`
28+
* `Set.Iic.bounded_order`, within an `order_bot`
29+
* `Set.Ici.bounded_order`, within an `order_top`
30+
-/
31+
32+
33+
variable {α : Type _}
34+
35+
namespace Set
36+
37+
namespace Ico
38+
39+
instance semilatticeInf [SemilatticeInf α] {a b : α} : SemilatticeInf (Ico a b) :=
40+
Subtype.semilatticeInf fun _ _ hx hy => ⟨le_inf hx.1 hy.1, lt_of_le_of_lt inf_le_left hx.2
41+
42+
/-- `Ico a b` has a bottom element whenever `a < b`. -/
43+
@[reducible]
44+
protected def orderBot [PartialOrder α] {a b : α} (h : a < b) : OrderBot (Ico a b) :=
45+
(isLeast_Ico h).orderBot
46+
#align set.Ico.order_bot Set.Ico.orderBot
47+
48+
end Ico
49+
50+
namespace Iio
51+
52+
instance semilatticeInf [SemilatticeInf α] {a : α} : SemilatticeInf (Iio a) :=
53+
Subtype.semilatticeInf fun _ _ hx _ => lt_of_le_of_lt inf_le_left hx
54+
55+
end Iio
56+
57+
namespace Ioc
58+
59+
instance semilatticeSup [SemilatticeSup α] {a b : α} : SemilatticeSup (Ioc a b) :=
60+
Subtype.semilatticeSup fun _ _ hx hy => ⟨lt_of_lt_of_le hx.1 le_sup_left, sup_le hx.2 hy.2
61+
62+
/-- `Ioc a b` has a top element whenever `a < b`. -/
63+
@[reducible]
64+
protected def orderTop [PartialOrder α] {a b : α} (h : a < b) : OrderTop (Ioc a b) :=
65+
(isGreatest_Ioc h).orderTop
66+
#align set.Ioc.order_top Set.Ioc.orderTop
67+
68+
end Ioc
69+
70+
namespace Ioi
71+
72+
instance semilatticeSup [SemilatticeSup α] {a : α} : SemilatticeSup (Ioi a) :=
73+
Subtype.semilatticeSup fun _ _ hx _ => lt_of_lt_of_le hx le_sup_left
74+
75+
end Ioi
76+
77+
namespace Iic
78+
79+
instance semilatticeInf [SemilatticeInf α] {a : α} : SemilatticeInf (Iic a) :=
80+
Subtype.semilatticeInf fun _ _ hx _ => le_trans inf_le_left hx
81+
82+
instance semilatticeSup [SemilatticeSup α] {a : α} : SemilatticeSup (Iic a) :=
83+
Subtype.semilatticeSup fun _ _ hx hy => sup_le hx hy
84+
85+
instance [Lattice α] {a : α} : Lattice (Iic a) :=
86+
{ Iic.semilatticeInf, Iic.semilatticeSup with }
87+
88+
instance orderTop [Preorder α] {a : α} :
89+
OrderTop (Iic a) where
90+
top := ⟨a, le_refl a⟩
91+
le_top x := x.prop
92+
93+
@[simp]
94+
theorem coe_top [Preorder α] {a : α} : (⊤ : Iic a) = a :=
95+
rfl
96+
#align set.Iic.coe_top Set.Iic.coe_top
97+
98+
instance orderBot [Preorder α] [OrderBot α] {a : α} :
99+
OrderBot (Iic a) where
100+
bot := ⟨⊥, bot_le⟩
101+
bot_le := fun ⟨_, _⟩ => Subtype.mk_le_mk.2 bot_le
102+
103+
@[simp]
104+
theorem coe_bot [Preorder α] [OrderBot α] {a : α} : (⊥ : Iic a) = (⊥ : α) :=
105+
rfl
106+
#align set.Iic.coe_bot Set.Iic.coe_bot
107+
108+
instance [Preorder α] [OrderBot α] {a : α} : BoundedOrder (Iic a) :=
109+
{ Iic.orderTop, Iic.orderBot with }
110+
111+
end Iic
112+
113+
namespace Ici
114+
115+
instance semilatticeInf [SemilatticeInf α] {a : α} : SemilatticeInf (Ici a) :=
116+
Subtype.semilatticeInf fun _ _ hx hy => le_inf hx hy
117+
118+
instance semilatticeSup [SemilatticeSup α] {a : α} : SemilatticeSup (Ici a) :=
119+
Subtype.semilatticeSup fun _ _ hx _ => le_trans hx le_sup_left
120+
121+
instance lattice [Lattice α] {a : α} : Lattice (Ici a) :=
122+
{ Ici.semilatticeInf, Ici.semilatticeSup with }
123+
124+
instance distribLattice [DistribLattice α] {a : α} : DistribLattice (Ici a) :=
125+
{ Ici.lattice with le_sup_inf := fun _ _ _ => le_sup_inf }
126+
127+
instance orderBot [Preorder α] {a : α} :
128+
OrderBot (Ici a) where
129+
bot := ⟨a, le_refl a⟩
130+
bot_le x := x.prop
131+
132+
@[simp]
133+
theorem coe_bot [Preorder α] {a : α} : ↑(⊥ : Ici a) = a :=
134+
rfl
135+
#align set.Ici.coe_bot Set.Ici.coe_bot
136+
137+
instance orderTop [Preorder α] [OrderTop α] {a : α} :
138+
OrderTop (Ici a) where
139+
top := ⟨⊤, le_top⟩
140+
le_top := fun ⟨_, _⟩ => Subtype.mk_le_mk.2 le_top
141+
142+
@[simp]
143+
theorem coe_top [Preorder α] [OrderTop α] {a : α} : ↑(⊤ : Ici a) = (⊤ : α) :=
144+
rfl
145+
#align set.Ici.coe_top Set.Ici.coe_top
146+
147+
instance boundedOrder [Preorder α] [OrderTop α] {a : α} : BoundedOrder (Ici a) :=
148+
{ Ici.orderTop, Ici.orderBot with }
149+
150+
end Ici
151+
152+
namespace Icc
153+
154+
instance semilatticeInf [SemilatticeInf α] {a b : α} : SemilatticeInf (Icc a b) :=
155+
Subtype.semilatticeInf fun _ _ hx hy => ⟨le_inf hx.1 hy.1, le_trans inf_le_left hx.2
156+
157+
instance semilatticeSup [SemilatticeSup α] {a b : α} : SemilatticeSup (Icc a b) :=
158+
Subtype.semilatticeSup fun _ _ hx hy => ⟨le_trans hx.1 le_sup_left, sup_le hx.2 hy.2
159+
160+
instance lattice [Lattice α] {a b : α} : Lattice (Icc a b) :=
161+
{ Icc.semilatticeInf, Icc.semilatticeSup with }
162+
163+
/-- `Icc a b` has a bottom element whenever `a ≤ b`. -/
164+
@[reducible]
165+
protected def orderBot [Preorder α] {a b : α} (h : a ≤ b) : OrderBot (Icc a b) :=
166+
(isLeast_Icc h).orderBot
167+
#align set.Icc.order_bot Set.Icc.orderBot
168+
169+
/-- `Icc a b` has a top element whenever `a ≤ b`. -/
170+
@[reducible]
171+
protected def orderTop [Preorder α] {a b : α} (h : a ≤ b) : OrderTop (Icc a b) :=
172+
(isGreatest_Icc h).orderTop
173+
#align set.Icc.order_top Set.Icc.orderTop
174+
175+
/-- `Icc a b` is a `bounded_order` whenever `a ≤ b`. -/
176+
@[reducible]
177+
protected def boundedOrder [Preorder α] {a b : α} (h : a ≤ b) : BoundedOrder (Icc a b) :=
178+
{ Icc.orderTop h, Icc.orderBot h with }
179+
#align set.Icc.bounded_order Set.Icc.boundedOrder
180+
181+
end Icc
182+
183+
end Set

0 commit comments

Comments
 (0)