Skip to content

Commit 7732480

Browse files
committed
chore: relax typeclass assumptions in two modular lattice lemmas (#7555)
Co-authored-by: Junyan Xu <junyanxumath@gmail.com> @alreadydone
1 parent fb3884a commit 7732480

File tree

1 file changed

+2
-2
lines changed

1 file changed

+2
-2
lines changed

Mathlib/Order/ModularLattice.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -409,12 +409,12 @@ theorem disjoint_sup_left_of_disjoint_sup_right [Lattice α] [OrderBot α]
409409
rwa [sup_comm, disjoint_comm] at hsup
410410
#align disjoint.disjoint_sup_left_of_disjoint_sup_right Disjoint.disjoint_sup_left_of_disjoint_sup_right
411411

412-
theorem isCompl_sup_right_of_isCompl_sup_left [CompleteLattice α] [IsModularLattice α]
412+
theorem isCompl_sup_right_of_isCompl_sup_left [Lattice α] [BoundedOrder α] [IsModularLattice α]
413413
(h : Disjoint a b) (hcomp : IsCompl (a ⊔ b) c) :
414414
IsCompl a (b ⊔ c) :=
415415
⟨h.disjoint_sup_right_of_disjoint_sup_left hcomp.disjoint, codisjoint_assoc.mp hcomp.codisjoint⟩
416416

417-
theorem isCompl_sup_left_of_isCompl_sup_right [CompleteLattice α] [IsModularLattice α]
417+
theorem isCompl_sup_left_of_isCompl_sup_right [Lattice α] [BoundedOrder α] [IsModularLattice α]
418418
(h : Disjoint b c) (hcomp : IsCompl a (b ⊔ c)) :
419419
IsCompl (a ⊔ b) c :=
420420
⟨h.disjoint_sup_left_of_disjoint_sup_right hcomp.disjoint, codisjoint_assoc.mpr hcomp.codisjoint⟩

0 commit comments

Comments
 (0)