Skip to content

Commit 34a3845

Browse files
feat: port Order.Category.CompleteLatCat (#5019)
Co-authored-by: Yaël Dillies <yael.dillies@gmail.com>
1 parent b75e009 commit 34a3845

File tree

4 files changed

+104
-0
lines changed

4 files changed

+104
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2273,6 +2273,7 @@ import Mathlib.Order.Bounds.OrderIso
22732273
import Mathlib.Order.Category.BddDistLatCat
22742274
import Mathlib.Order.Category.BddLatCat
22752275
import Mathlib.Order.Category.BddOrdCat
2276+
import Mathlib.Order.Category.CompleteLatCat
22762277
import Mathlib.Order.Category.DistLatCat
22772278
import Mathlib.Order.Category.FinPartOrd
22782279
import Mathlib.Order.Category.FrmCat

Mathlib/Order/Category/BddOrdCat.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -65,6 +65,7 @@ instance largeCategory : LargeCategory.{u} BddOrdCat where
6565
#align BddOrd.large_category BddOrdCat.largeCategory
6666

6767
-- Porting note: added.
68+
-- see https://github.com/leanprover-community/mathlib4/issues/5017
6869
instance instFunLike (X Y : BddOrdCat) : FunLike (X ⟶ Y) X (fun _ => Y) :=
6970
show FunLike (BoundedOrderHom X Y) X (fun _ => Y) from inferInstance
7071

Lines changed: 101 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,101 @@
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.CompleteLat
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.BddLatCat
12+
import Mathlib.Order.Hom.CompleteLattice
13+
14+
/-!
15+
# The category of complete lattices
16+
17+
This file defines `CompleteLatCat`, the category of complete lattices.
18+
-/
19+
20+
set_option linter.uppercaseLean3 false
21+
22+
universe u
23+
24+
open CategoryTheory
25+
26+
/-- The category of complete lattices. -/
27+
def CompleteLatCat :=
28+
Bundled CompleteLattice
29+
#align CompleteLat CompleteLatCat
30+
31+
namespace CompleteLatCat
32+
33+
instance : CoeSort CompleteLatCat (Type _) :=
34+
Bundled.coeSort
35+
36+
instance (X : CompleteLatCat) : CompleteLattice X :=
37+
X.str
38+
39+
/-- Construct a bundled `CompleteLatCat` from a `CompleteLattice`. -/
40+
def of (α : Type _) [CompleteLattice α] : CompleteLatCat :=
41+
Bundled.of α
42+
#align CompleteLat.of CompleteLatCat.of
43+
44+
@[simp]
45+
theorem coe_of (α : Type _) [CompleteLattice α] : ↥(of α) = α :=
46+
rfl
47+
#align CompleteLat.coe_of CompleteLatCat.coe_of
48+
49+
instance : Inhabited CompleteLatCat :=
50+
⟨of PUnit⟩
51+
52+
instance : BundledHom @CompleteLatticeHom where
53+
toFun _ _ f := f.toFun
54+
id := @CompleteLatticeHom.id
55+
comp := @CompleteLatticeHom.comp
56+
hom_ext _ _ _ _ h := FunLike.coe_injective h
57+
58+
deriving instance LargeCategory for CompleteLatCat
59+
60+
instance : ConcreteCategory CompleteLatCat :=
61+
by dsimp [CompleteLatCat]; infer_instance
62+
63+
instance hasForgetToBddLat : HasForget₂ CompleteLatCat BddLatCat where
64+
forget₂ :=
65+
{ obj := fun X => BddLatCat.of X
66+
map := fun {X Y} => CompleteLatticeHom.toBoundedLatticeHom }
67+
forget_comp := rfl
68+
#align CompleteLat.has_forget_to_BddLat CompleteLatCat.hasForgetToBddLat
69+
70+
/-- Constructs an isomorphism of complete lattices from an order isomorphism between them. -/
71+
@[simps]
72+
def Iso.mk {α β : CompleteLatCat.{u}} (e : α ≃o β) : α ≅ β where
73+
hom := (e : CompleteLatticeHom _ _) -- Porting note: TODO, wrong?
74+
inv := (e.symm : CompleteLatticeHom _ _)
75+
hom_inv_id := by ext; exact e.symm_apply_apply _
76+
inv_hom_id := by ext; exact e.apply_symm_apply _
77+
#align CompleteLat.iso.mk CompleteLatCat.Iso.mk
78+
79+
/-- `OrderDual` as a functor. -/
80+
@[simps]
81+
def dual : CompleteLatCat ⥤ CompleteLatCat where
82+
obj X := of Xᵒᵈ
83+
map {X Y} := CompleteLatticeHom.dual
84+
#align CompleteLat.dual CompleteLatCat.dual
85+
86+
/-- The equivalence between `CompleteLatCat` and itself induced by `OrderDual` both ways. -/
87+
@[simps functor inverse]
88+
def dualEquiv : CompleteLatCat ≌ CompleteLatCat where
89+
functor := dual
90+
inverse := dual
91+
unitIso := NatIso.ofComponents fun X => Iso.mk <| OrderIso.dualDual X
92+
counitIso := NatIso.ofComponents fun X => Iso.mk <| OrderIso.dualDual X
93+
#align CompleteLat.dual_equiv CompleteLatCat.dualEquiv
94+
95+
end CompleteLatCat
96+
97+
theorem completeLatCat_dual_comp_forget_to_bddLatCat :
98+
CompleteLatCat.dual ⋙ forget₂ CompleteLatCat BddLatCat =
99+
forget₂ CompleteLatCat BddLatCat ⋙ BddLatCat.dual :=
100+
rfl
101+
#align CompleteLat_dual_comp_forget_to_BddLat completeLatCat_dual_comp_forget_to_bddLatCat

Mathlib/Order/Category/SemilatCat.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -71,6 +71,7 @@ instance : LargeCategory.{u} SemilatSupCat where
7171
assoc _ _ _ := SupBotHom.comp_assoc _ _ _
7272

7373
-- Porting note: added
74+
-- see https://github.com/leanprover-community/mathlib4/issues/5017
7475
instance instFunLike (X Y : SemilatSupCat) : FunLike (X ⟶ Y) X (fun _ => Y) :=
7576
show FunLike (SupBotHom X Y) X (fun _ => Y) from inferInstance
7677

0 commit comments

Comments
 (0)