Skip to content

Commit dc318ce

Browse files
joneugsterParcly-Taxeljcommelin
committed
feat: port Order.Category.SemilatCat (#4990)
Co-authored-by: Jeremy Tan Jie Rui <reddeloostw@gmail.com> Co-authored-by: Johan Commelin <johan@commelin.net>
1 parent a121cf8 commit dc318ce

File tree

2 files changed

+214
-0
lines changed

2 files changed

+214
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2257,6 +2257,7 @@ import Mathlib.Order.Category.LinOrdCat
22572257
import Mathlib.Order.Category.NonemptyFinLinOrdCat
22582258
import Mathlib.Order.Category.PartOrdCat
22592259
import Mathlib.Order.Category.PreordCat
2260+
import Mathlib.Order.Category.SemilatCat
22602261
import Mathlib.Order.Chain
22612262
import Mathlib.Order.Circular
22622263
import Mathlib.Order.Closure
Lines changed: 213 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,213 @@
1+
/-
2+
Copyright (c) 2022 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+
! This file was ported from Lean 3 source module order.category.Semilat
7+
! leanprover-community/mathlib commit e8ac6315bcfcbaf2d19a046719c3b553206dac75
8+
! Please do not edit these lines, except to modify the commit id
9+
! if you have ported upstream changes.
10+
-/
11+
import Mathlib.Order.Category.PartOrdCat
12+
import Mathlib.Order.Hom.Lattice
13+
14+
/-!
15+
# The categories of semilattices
16+
17+
This defines `SemilatSupCat` and `SemilatInfCat`, the categories of sup-semilattices with a bottom
18+
element and inf-semilattices with a top element.
19+
20+
## References
21+
22+
* [nLab, *semilattice*](https://ncatlab.org/nlab/show/semilattice)
23+
-/
24+
25+
set_option linter.uppercaseLean3 false
26+
27+
universe u
28+
29+
open CategoryTheory
30+
31+
/-- The category of sup-semilattices with a bottom element. -/
32+
structure SemilatSupCat : Type (u + 1) where
33+
protected X : Type u
34+
[isSemilatticeSup : SemilatticeSup X]
35+
[isOrderBot : OrderBot.{u} X]
36+
#align SemilatSup SemilatSupCat
37+
38+
/-- The category of inf-semilattices with a top element. -/
39+
structure SemilatInfCat : Type (u + 1) where
40+
protected X : Type u
41+
[isSemilatticeInf : SemilatticeInf X]
42+
[isOrderTop : OrderTop.{u} X]
43+
#align SemilatInf SemilatInfCat
44+
45+
namespace SemilatSupCat
46+
47+
instance : CoeSort SemilatSupCat (Type _) :=
48+
⟨SemilatSupCat.X⟩
49+
50+
attribute [instance] isSemilatticeSup isOrderBot
51+
52+
/-- Construct a bundled `SemilatSupCat` from a `SemilatticeSup`. -/
53+
def of (α : Type _) [SemilatticeSup α] [OrderBot α] : SemilatSupCat :=
54+
⟨α⟩
55+
#align SemilatSup.of SemilatSupCat.of
56+
57+
@[simp]
58+
theorem coe_of (α : Type _) [SemilatticeSup α] [OrderBot α] : ↥(of α) = α :=
59+
rfl
60+
#align SemilatSup.coe_of SemilatSupCat.coe_of
61+
62+
instance : Inhabited SemilatSupCat :=
63+
⟨of PUnit⟩
64+
65+
instance : LargeCategory.{u} SemilatSupCat where
66+
Hom X Y := SupBotHom X Y
67+
id X := SupBotHom.id X
68+
comp f g := g.comp f
69+
id_comp := SupBotHom.comp_id
70+
comp_id := SupBotHom.id_comp
71+
assoc _ _ _ := SupBotHom.comp_assoc _ _ _
72+
73+
-- Porting note: added
74+
instance instFunLike (X Y : SemilatSupCat) : FunLike (X ⟶ Y) X (fun _ => Y) :=
75+
show FunLike (SupBotHom X Y) X (fun _ => Y) from inferInstance
76+
77+
instance : ConcreteCategory SemilatSupCat where
78+
forget :=
79+
{ obj := SemilatSupCat.X
80+
map := FunLike.coe }
81+
forget_faithful := ⟨(FunLike.coe_injective ·)⟩
82+
83+
instance hasForgetToPartOrd : HasForget₂ SemilatSupCat PartOrdCat where
84+
forget₂ :=
85+
-- Porting note: was ⟨X⟩, see https://github.com/leanprover-community/mathlib4/issues/4998
86+
{ obj := fun X => {α := X}
87+
-- Porting note: was `map := fun f => f`
88+
map := fun f => ⟨f.toSupHom, OrderHomClass.mono f.toSupHom⟩ }
89+
#align SemilatSup.has_forget_to_PartOrd SemilatSupCat.hasForgetToPartOrd
90+
91+
@[simp]
92+
theorem coe_forget_to_partOrdCat (X : SemilatSupCat) :
93+
↥((forget₂ SemilatSupCat PartOrdCat).obj X) = ↥X :=
94+
rfl
95+
#align SemilatSup.coe_forget_to_PartOrd SemilatSupCat.coe_forget_to_partOrdCat
96+
97+
end SemilatSupCat
98+
99+
namespace SemilatInfCat
100+
101+
instance : CoeSort SemilatInfCat (Type _) :=
102+
⟨SemilatInfCat.X⟩
103+
104+
attribute [instance] isSemilatticeInf isOrderTop
105+
106+
/-- Construct a bundled `SemilatInfCat` from a `SemilatticeInf`. -/
107+
def of (α : Type _) [SemilatticeInf α] [OrderTop α] : SemilatInfCat :=
108+
⟨α⟩
109+
#align SemilatInf.of SemilatInfCat.of
110+
111+
@[simp]
112+
theorem coe_of (α : Type _) [SemilatticeInf α] [OrderTop α] : ↥(of α) = α :=
113+
rfl
114+
#align SemilatInf.coe_of SemilatInfCat.coe_of
115+
116+
instance : Inhabited SemilatInfCat :=
117+
⟨of PUnit⟩
118+
119+
instance : LargeCategory.{u} SemilatInfCat where
120+
Hom X Y := InfTopHom X Y
121+
id X := InfTopHom.id X
122+
comp f g := g.comp f
123+
id_comp := InfTopHom.comp_id
124+
comp_id := InfTopHom.id_comp
125+
assoc _ _ _ := InfTopHom.comp_assoc _ _ _
126+
127+
-- Porting note: added
128+
instance instFunLike (X Y : SemilatInfCat) : FunLike (X ⟶ Y) X (fun _ => Y) :=
129+
show FunLike (InfTopHom X Y) X (fun _ => Y) from inferInstance
130+
131+
instance : ConcreteCategory SemilatInfCat where
132+
forget :=
133+
{ obj := SemilatInfCat.X
134+
map := FunLike.coe }
135+
forget_faithful := ⟨(FunLike.coe_injective ·)⟩
136+
137+
instance hasForgetToPartOrd : HasForget₂ SemilatInfCat PartOrdCat where
138+
forget₂ :=
139+
{ obj := fun X => ⟨X, inferInstance⟩
140+
-- Porting note: was `map := fun f => f`
141+
map := fun f => ⟨f.toInfHom, OrderHomClass.mono f.toInfHom⟩ }
142+
#align SemilatInf.has_forget_to_PartOrd SemilatInfCat.hasForgetToPartOrd
143+
144+
@[simp]
145+
theorem coe_forget_to_partOrdCat (X : SemilatInfCat) :
146+
↥((forget₂ SemilatInfCat PartOrdCat).obj X) = ↥X :=
147+
rfl
148+
#align SemilatInf.coe_forget_to_PartOrd SemilatInfCat.coe_forget_to_partOrdCat
149+
150+
end SemilatInfCat
151+
152+
/-! ### Order dual -/
153+
154+
namespace SemilatSupCat
155+
156+
/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/
157+
@[simps]
158+
def Iso.mk {α β : SemilatSupCat.{u}} (e : α ≃o β) : α ≅ β where
159+
hom := (e : SupBotHom _ _)
160+
inv := (e.symm : SupBotHom _ _)
161+
hom_inv_id := by ext; exact e.symm_apply_apply _
162+
inv_hom_id := by ext; exact e.apply_symm_apply _
163+
#align SemilatSup.iso.mk SemilatSupCat.Iso.mk
164+
165+
/-- `order_dual` as a functor. -/
166+
@[simps]
167+
def dual : SemilatSupCat ⥤ SemilatInfCat where
168+
obj X := SemilatInfCat.of Xᵒᵈ
169+
map {X Y} := SupBotHom.dual
170+
#align SemilatSup.dual SemilatSupCat.dual
171+
172+
end SemilatSupCat
173+
174+
namespace SemilatInfCat
175+
176+
/-- Constructs an isomorphism of lattices from an order isomorphism between them. -/
177+
@[simps]
178+
def Iso.mk {α β : SemilatInfCat.{u}} (e : α ≃o β) : α ≅ β where
179+
hom := (e : InfTopHom _ _)
180+
inv := (e.symm : InfTopHom _ _)
181+
hom_inv_id := by ext; exact e.symm_apply_apply _
182+
inv_hom_id := by ext; exact e.apply_symm_apply _
183+
#align SemilatInf.iso.mk SemilatInfCat.Iso.mk
184+
185+
/-- `OrderDual` as a functor. -/
186+
@[simps]
187+
def dual : SemilatInfCat ⥤ SemilatSupCat where
188+
obj X := SemilatSupCat.of Xᵒᵈ
189+
map {X Y} := InfTopHom.dual
190+
#align SemilatInf.dual SemilatInfCat.dual
191+
192+
end SemilatInfCat
193+
194+
/-- The equivalence between `SemilatSupCat` and `SemilatInfCat` induced by `OrderDual` both ways. -/
195+
@[simps functor inverse]
196+
def SemilatSupCatEquivSemilatInfCat : SemilatSupCat ≌ SemilatInfCat where
197+
functor := SemilatSupCat.dual
198+
inverse := SemilatInfCat.dual
199+
unitIso := NatIso.ofComponents fun X => SemilatSupCat.Iso.mk <| OrderIso.dualDual X
200+
counitIso := NatIso.ofComponents fun X => SemilatInfCat.Iso.mk <| OrderIso.dualDual X
201+
#align SemilatSup_equiv_SemilatInf SemilatSupCatEquivSemilatInfCat
202+
203+
theorem SemilatSupCat_dual_comp_forget_to_partOrdCat :
204+
SemilatSupCat.dual ⋙ forget₂ SemilatInfCat PartOrdCat =
205+
forget₂ SemilatSupCat PartOrdCat ⋙ PartOrdCat.dual :=
206+
rfl
207+
#align SemilatSup_dual_comp_forget_to_PartOrd SemilatSupCat_dual_comp_forget_to_partOrdCat
208+
209+
theorem SemilatInfCat_dual_comp_forget_to_partOrdCat :
210+
SemilatInfCat.dual ⋙ forget₂ SemilatSupCat PartOrdCat =
211+
forget₂ SemilatInfCat PartOrdCat ⋙ PartOrdCat.dual :=
212+
rfl
213+
#align SemilatInf_dual_comp_forget_to_PartOrd SemilatInfCat_dual_comp_forget_to_partOrdCat

0 commit comments

Comments
 (0)