Make 'isLattice' for simple types be an instance

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
Danila Fedorin 2025-01-04 17:27:38 -08:00
parent 4da9b6d3cd
commit 8abf6f8670
2 changed files with 92 additions and 92 deletions

View File

@ -18,8 +18,9 @@ private
≡-⊓-cong : {a₁ a₂ a₃ a₄} a₁ a₂ a₃ a₄ (a₁ a₃) (a₂ a₄) ≡-⊓-cong : {a₁ a₂ a₃ a₄} a₁ a₂ a₃ a₄ (a₁ a₃) (a₂ a₄)
≡-⊓-cong a₁≡a₂ a₃≡a₄ rewrite a₁≡a₂ rewrite a₃≡a₄ = refl ≡-⊓-cong a₁≡a₂ a₃≡a₄ rewrite a₁≡a₂ rewrite a₃≡a₄ = refl
isMaxSemilattice : IsSemilattice _≡_ _⊔_ instance
isMaxSemilattice = record isMaxSemilattice : IsSemilattice _≡_ _⊔_
isMaxSemilattice = record
{ ≈-equiv = record { ≈-equiv = record
{ ≈-refl = refl { ≈-refl = refl
; ≈-sym = sym ; ≈-sym = sym
@ -31,8 +32,8 @@ isMaxSemilattice = record
; ⊔-idemp = ⊔-idem ; ⊔-idemp = ⊔-idem
} }
isMinSemilattice : IsSemilattice _≡_ _⊓_ isMinSemilattice : IsSemilattice _≡_ _⊓_
isMinSemilattice = record isMinSemilattice = record
{ ≈-equiv = record { ≈-equiv = record
{ ≈-refl = refl { ≈-refl = refl
; ≈-sym = sym ; ≈-sym = sym
@ -74,16 +75,17 @@ private
helper : x (x y) x x x x x x (x y) x helper : x (x y) x x x x x x (x y) x
helper x⊔x⊓y≤x⊔x x⊔x≡x rewrite x⊔x≡x = x⊔x⊓y≤x⊔x helper x⊔x⊓y≤x⊔x x⊔x≡x rewrite x⊔x≡x = x⊔x⊓y≤x⊔x
isLattice : IsLattice _≡_ _⊔_ _⊓_ instance
isLattice = record isLattice : IsLattice _≡_ _⊔_ _⊓_
isLattice = record
{ joinSemilattice = isMaxSemilattice { joinSemilattice = isMaxSemilattice
; meetSemilattice = isMinSemilattice ; meetSemilattice = isMinSemilattice
; absorb-⊔-⊓ = λ x y maxmin-absorb {x} {y} ; absorb-⊔-⊓ = λ x y maxmin-absorb {x} {y}
; absorb-⊓-⊔ = λ x y minmax-absorb {x} {y} ; absorb-⊓-⊔ = λ x y minmax-absorb {x} {y}
} }
lattice : Lattice lattice : Lattice
lattice = record lattice = record
{ _≈_ = _≡_ { _≈_ = _≡_
; _⊔_ = _⊔_ ; _⊔_ = _⊔_
; _⊓_ = _⊓_ ; _⊓_ = _⊓_

View File

@ -45,8 +45,9 @@ tt ⊓ tt = tt
⊔-idemp : (x : ) (x x) x ⊔-idemp : (x : ) (x x) x
⊔-idemp tt = Eq.refl ⊔-idemp tt = Eq.refl
isJoinSemilattice : IsSemilattice _≈_ _⊔_ instance
isJoinSemilattice = record isJoinSemilattice : IsSemilattice _≈_ _⊔_
isJoinSemilattice = record
{ ≈-equiv = ≈-equiv { ≈-equiv = ≈-equiv
; ≈-⊔-cong = ≈-⊔-cong ; ≈-⊔-cong = ≈-⊔-cong
; ⊔-assoc = ⊔-assoc ; ⊔-assoc = ⊔-assoc
@ -66,8 +67,9 @@ isJoinSemilattice = record
⊓-idemp : (x : ) (x x) x ⊓-idemp : (x : ) (x x) x
⊓-idemp tt = Eq.refl ⊓-idemp tt = Eq.refl
isMeetSemilattice : IsSemilattice _≈_ _⊓_ instance
isMeetSemilattice = record isMeetSemilattice : IsSemilattice _≈_ _⊓_
isMeetSemilattice = record
{ ≈-equiv = ≈-equiv { ≈-equiv = ≈-equiv
; ≈-⊔-cong = ≈-⊓-cong ; ≈-⊔-cong = ≈-⊓-cong
; ⊔-assoc = ⊓-assoc ; ⊔-assoc = ⊓-assoc
@ -75,22 +77,17 @@ isMeetSemilattice = record
; ⊔-idemp = ⊓-idemp ; ⊔-idemp = ⊓-idemp
} }
absorb-⊔-⊓ : (x y : ) (x (x y)) x instance
absorb-⊔-⊓ tt tt = Eq.refl isLattice : IsLattice _≈_ _⊔_ _⊓_
isLattice = record
absorb-⊓-⊔ : (x y : ) (x (x y)) x
absorb-⊓-⊔ tt tt = Eq.refl
isLattice : IsLattice _≈_ _⊔_ _⊓_
isLattice = record
{ joinSemilattice = isJoinSemilattice { joinSemilattice = isJoinSemilattice
; meetSemilattice = isMeetSemilattice ; meetSemilattice = isMeetSemilattice
; absorb-⊔-⊓ = absorb-⊔-⊓ ; absorb-⊔-⊓ = λ { tt tt Eq.refl }
; absorb-⊓-⊔ = absorb-⊓-⊔ ; absorb-⊓-⊔ = λ { tt tt Eq.refl }
} }
lattice : Lattice lattice : Lattice
lattice = record lattice = record
{ _≈_ = _≈_ { _≈_ = _≈_
; _⊔_ = _⊔_ ; _⊔_ = _⊔_
; _⊓_ = _⊓_ ; _⊓_ = _⊓_
@ -107,22 +104,23 @@ private
isLongest {tt} {tt} (step (tt⊔tt≈tt , tt̷≈tt) _ _) = ⊥-elim (tt̷≈tt refl) isLongest {tt} {tt} (step (tt⊔tt≈tt , tt̷≈tt) _ _) = ⊥-elim (tt̷≈tt refl)
isLongest (done _) = z≤n isLongest (done _) = z≤n
fixedHeight : IsLattice.FixedHeight isLattice 0 instance
fixedHeight = record fixedHeight : IsLattice.FixedHeight isLattice 0
fixedHeight = record
{ = tt { = tt
; = tt ; = tt
; longestChain = longestChain ; longestChain = longestChain
; bounded = isLongest ; bounded = isLongest
} }
isFiniteHeightLattice : IsFiniteHeightLattice 0 _≈_ _⊔_ _⊓_ isFiniteHeightLattice : IsFiniteHeightLattice 0 _≈_ _⊔_ _⊓_
isFiniteHeightLattice = record isFiniteHeightLattice = record
{ isLattice = isLattice { isLattice = isLattice
; fixedHeight = fixedHeight ; fixedHeight = fixedHeight
} }
finiteHeightLattice : FiniteHeightLattice finiteHeightLattice : FiniteHeightLattice
finiteHeightLattice = record finiteHeightLattice = record
{ height = 0 { height = 0
; _≈_ = _≈_ ; _≈_ = _≈_
; _⊔_ = _⊔_ ; _⊔_ = _⊔_