Skip to content

Commit ad83df7

Browse files
YaelDillieskim-em
authored andcommitted
feat: The lattice of complemented elements (#5194)
Match leanprover-community/mathlib3#16267
1 parent 2d9691c commit ad83df7

File tree

1 file changed

+138
-1
lines changed

1 file changed

+138
-1
lines changed

Mathlib/Order/Disjoint.lean

Lines changed: 138 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Johannes Hölzl
55
66
! This file was ported from Lean 3 source module order.disjoint
7-
! leanprover-community/mathlib commit 70d50ecfd4900dd6d328da39ab7ebd516abe4025
7+
! leanprover-community/mathlib commit 22c4d2ff43714b6ff724b2745ccfdc0f236a4a76
88
! Please do not edit these lines, except to modify the commit id
99
! if you have ported upstream changes.
1010
-/
@@ -25,6 +25,7 @@ This file defines `Disjoint`, `Codisjoint`, and the `IsCompl` predicate.
2525
2626
-/
2727

28+
open Function
2829

2930
variable {α : Type _}
3031

@@ -644,6 +645,39 @@ theorem eq_bot_of_top_isCompl (h : IsCompl ⊤ x) : x = ⊥ :=
644645

645646
end
646647

648+
section IsComplemented
649+
650+
section Lattice
651+
652+
variable [Lattice α] [BoundedOrder α]
653+
654+
/-- An element is *complemented* if it has a complement. -/
655+
def IsComplemented (a : α) : Prop :=
656+
∃ b, IsCompl a b
657+
#align is_complemented IsComplemented
658+
659+
theorem isComplemented_bot : IsComplemented (⊥ : α) :=
660+
⟨⊤, isCompl_bot_top⟩
661+
#align is_complemented_bot isComplemented_bot
662+
663+
theorem isComplemented_top : IsComplemented (⊤ : α) :=
664+
⟨⊥, isCompl_top_bot⟩
665+
#align is_complemented_top isComplemented_top
666+
667+
end Lattice
668+
669+
variable [DistribLattice α] [BoundedOrder α] {a b : α}
670+
671+
theorem IsComplemented.sup : IsComplemented a → IsComplemented b → IsComplemented (a ⊔ b) :=
672+
fun ⟨a', ha⟩ ⟨b', hb⟩ => ⟨a' ⊓ b', ha.sup_inf hb⟩
673+
#align is_complemented.sup IsComplemented.sup
674+
675+
theorem IsComplemented.inf : IsComplemented a → IsComplemented b → IsComplemented (a ⊓ b) :=
676+
fun ⟨a', ha⟩ ⟨b', hb⟩ => ⟨a' ⊔ b', ha.inf_sup hb⟩
677+
#align is_complemented.inf IsComplemented.inf
678+
679+
end IsComplemented
680+
647681
/-- A complemented bounded lattice is one where every element has a (not necessarily unique)
648682
complement. -/
649683
class ComplementedLattice (α) [Lattice α] [BoundedOrder α] : Prop where
@@ -664,4 +698,107 @@ instance : ComplementedLattice αᵒᵈ :=
664698

665699
end ComplementedLattice
666700

701+
-- TODO: Define as a sublattice?
702+
/-- The sublattice of complemented elements. -/
703+
@[reducible]
704+
def Complementeds (α : Type _) [Lattice α] [BoundedOrder α] : Type _ := {a : α // IsComplemented a}
705+
#align complementeds Complementeds
706+
707+
namespace Complementeds
708+
709+
section Lattice
710+
711+
variable [Lattice α] [BoundedOrder α] {a b : Complementeds α}
712+
713+
instance hasCoeT : CoeTC (Complementeds α) α := ⟨Subtype.val⟩
714+
#align complementeds.has_coe_t Complementeds.hasCoeT
715+
716+
theorem coe_injective : Injective ((↑) : Complementeds α → α) := Subtype.coe_injective
717+
#align complementeds.coe_injective Complementeds.coe_injective
718+
719+
@[simp, norm_cast]
720+
theorem coe_inj : (a : α) = b ↔ a = b := Subtype.coe_inj
721+
#align complementeds.coe_inj Complementeds.coe_inj
722+
723+
-- porting note: removing `simp` because `Subtype.coe_le_coe` already proves it
724+
@[norm_cast]
725+
theorem coe_le_coe : (a : α) ≤ b ↔ a ≤ b := by simp
726+
#align complementeds.coe_le_coe Complementeds.coe_le_coe
727+
728+
-- porting note: removing `simp` because `Subtype.coe_lt_coe` already proves it
729+
@[norm_cast]
730+
theorem coe_lt_coe : (a : α) < b ↔ a < b := Iff.rfl
731+
#align complementeds.coe_lt_coe Complementeds.coe_lt_coe
732+
733+
instance : BoundedOrder (Complementeds α) :=
734+
Subtype.boundedOrder isComplemented_bot isComplemented_top
735+
736+
@[simp, norm_cast]
737+
theorem coe_bot : ((⊥ : Complementeds α) : α) = ⊥ := rfl
738+
#align complementeds.coe_bot Complementeds.coe_bot
739+
740+
@[simp, norm_cast]
741+
theorem coe_top : ((⊤ : Complementeds α) : α) = ⊤ := rfl
742+
#align complementeds.coe_top Complementeds.coe_top
743+
744+
-- porting note: removing `simp` because `Subtype.mk_bot` already proves it
745+
theorem mk_bot : (⟨⊥, isComplemented_bot⟩ : Complementeds α) = ⊥ := rfl
746+
#align complementeds.mk_bot Complementeds.mk_bot
747+
748+
-- porting note: removing `simp` because `Subtype.mk_top` already proves it
749+
theorem mk_top : (⟨⊤, isComplemented_top⟩ : Complementeds α) = ⊤ := rfl
750+
#align complementeds.mk_top Complementeds.mk_top
751+
752+
instance : Inhabited (Complementeds α) := ⟨⊥⟩
753+
754+
end Lattice
755+
756+
variable [DistribLattice α] [BoundedOrder α] {a b : Complementeds α}
757+
758+
instance : Sup (Complementeds α) :=
759+
fun a b => ⟨a ⊔ b, a.2.sup b.2⟩⟩
760+
761+
instance : Inf (Complementeds α) :=
762+
fun a b => ⟨a ⊓ b, a.2.inf b.2⟩⟩
763+
764+
@[simp, norm_cast]
765+
theorem coe_sup (a b : Complementeds α) : ↑(a ⊔ b) = (a : α) ⊔ b := rfl
766+
#align complementeds.coe_sup Complementeds.coe_sup
767+
768+
@[simp, norm_cast]
769+
theorem coe_inf (a b : Complementeds α) : ↑(a ⊓ b) = (a : α) ⊓ b := rfl
770+
#align complementeds.coe_inf Complementeds.coe_inf
771+
772+
@[simp]
773+
theorem mk_sup_mk {a b : α} (ha : IsComplemented a) (hb : IsComplemented b) :
774+
(⟨a, ha⟩ ⊔ ⟨b, hb⟩ : Complementeds α) = ⟨a ⊔ b, ha.sup hb⟩ := rfl
775+
#align complementeds.mk_sup_mk Complementeds.mk_sup_mk
776+
777+
@[simp]
778+
theorem mk_inf_mk {a b : α} (ha : IsComplemented a) (hb : IsComplemented b) :
779+
(⟨a, ha⟩ ⊓ ⟨b, hb⟩ : Complementeds α) = ⟨a ⊓ b, ha.inf hb⟩ := rfl
780+
#align complementeds.mk_inf_mk Complementeds.mk_inf_mk
781+
782+
instance : DistribLattice (Complementeds α) :=
783+
Complementeds.coe_injective.distribLattice _ coe_sup coe_inf
784+
785+
@[simp, norm_cast]
786+
theorem disjoint_coe : Disjoint (a : α) b ↔ Disjoint a b := by
787+
rw [disjoint_iff, disjoint_iff, ← coe_inf, ← coe_bot, coe_inj]
788+
#align complementeds.disjoint_coe Complementeds.disjoint_coe
789+
790+
@[simp, norm_cast]
791+
theorem codisjoint_coe : Codisjoint (a : α) b ↔ Codisjoint a b := by
792+
rw [codisjoint_iff, codisjoint_iff, ← coe_sup, ← coe_top, coe_inj]
793+
#align complementeds.codisjoint_coe Complementeds.codisjoint_coe
794+
795+
@[simp, norm_cast]
796+
theorem isCompl_coe : IsCompl (a : α) b ↔ IsCompl a b := by
797+
simp_rw [isCompl_iff, disjoint_coe, codisjoint_coe]
798+
#align complementeds.is_compl_coe Complementeds.isCompl_coe
799+
800+
instance : ComplementedLattice (Complementeds α) :=
801+
fun ⟨a, b, h⟩ => ⟨⟨b, a, h.symm⟩, isCompl_coe.1 h⟩⟩
802+
803+
end Complementeds
667804
end IsCompl

0 commit comments

Comments
 (0)