Finish up the Nat semilattices

Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
Danila Fedorin 2023-07-13 21:50:27 -07:00
parent 97a3f25fd2
commit 422ea93edb

View File

@ -5,7 +5,8 @@ open import Relation.Binary.PropositionalEquality as Eq using (_≡_; sym)
open import Relation.Binary.Definitions open import Relation.Binary.Definitions
open import Data.Nat as Nat using (; _≤_) open import Data.Nat as Nat using (; _≤_)
open import Data.Product using (_×_; _,_) open import Data.Product using (_×_; _,_)
open import Agda.Primitive using (lsuc) open import Data.Sum using (_⊎_; inj₁; inj₂)
open import Agda.Primitive using (lsuc; Level)
open import NatMap using (NatMap) open import NatMap using (NatMap)
@ -15,6 +16,13 @@ record IsPreorder {a} (A : Set a) (_≼_ : A → A → Set a) : Set a where
≼-trans : Transitive (_≼_) ≼-trans : Transitive (_≼_)
≼-antisym : Antisymmetric (_≡_) (_≼_) ≼-antisym : Antisymmetric (_≡_) (_≼_)
isPreorderFlip : {a : Level} {A : Set a} {_≼_ : A A Set a} IsPreorder A _≼_ IsPreorder A (λ x y y x)
isPreorderFlip isPreorder = record
{ ≼-refl = IsPreorder.≼-refl isPreorder
; ≼-trans = λ {x} {y} {z} x≽y y≽z IsPreorder.≼-trans isPreorder y≽z x≽y
; ≼-antisym = λ {x} {y} x≽y y≽x IsPreorder.≼-antisym isPreorder y≽x x≽y
}
record Preorder {a} (A : Set a) : Set (lsuc a) where record Preorder {a} (A : Set a) : Set (lsuc a) where
field field
_≼_ : A A Set a _≼_ : A A Set a
@ -27,13 +35,12 @@ record IsSemilattice {a} (A : Set a) (_≼_ : A → A → Set a) (_⊔_ : A →
field field
isPreorder : IsPreorder A _≼_ isPreorder : IsPreorder A _≼_
⊔-assoc : (x : A) (y : A) (z : A) (x y) z x (y z) ⊔-assoc : (x y z : A) (x y) z x (y z)
⊔-comm : (x : A) (y : A) x y y x ⊔-comm : (x y : A) x y y x
⊔-idemp : (x : A) x x x ⊔-idemp : (x : A) x x x
⊔-bound : (x : A) (y : A) (z : A) x y z (x z × y z) ⊔-bound : (x y z : A) x y z (x z × y z)
⊔-least : (x : A) (y : A) (z : A) x y z ⊔-least : (x y z : A) x y z (z' : A) (x z' × y z') z z'
(z' : A) (x z' × y z') z z'
open IsPreorder isPreorder public open IsPreorder isPreorder public
@ -55,8 +62,8 @@ record IsLattice {a} (A : Set a) (_≼_ : A → A → Set a) (_⊔_ : A → A
joinSemilattice : IsSemilattice A _≼_ _⊔_ joinSemilattice : IsSemilattice A _≼_ _⊔_
meetSemilattice : IsSemilattice A _≽_ _⊓_ meetSemilattice : IsSemilattice A _≽_ _⊓_
absorb-⊔-⊓ : (x : A) (y : A) x (x y) x absorb-⊔-⊓ : (x y : A) x (x y) x
absorb-⊓-⊔ : (x : A) (y : A) x (x y) x absorb-⊓-⊔ : (x y : A) x (x y) x
open IsSemilattice joinSemilattice public open IsSemilattice joinSemilattice public
open IsSemilattice meetSemilattice public renaming open IsSemilattice meetSemilattice public renaming
@ -82,6 +89,7 @@ private module NatInstances where
open Nat open Nat
open NatProps open NatProps
open Eq open Eq
open Data.Sum
NatPreorder : Preorder NatPreorder : Preorder
NatPreorder = record NatPreorder = record
@ -94,14 +102,19 @@ private module NatInstances where
} }
private private
max-bound₁ : (x : ) (y : ) (z : ) x y z x z max-bound₁ : {x y z : } x y z x z
max-bound₁ x y z x⊔y≡z rewrite sym x⊔y≡z rewrite ⊔-comm x y = m≤n⇒m≤o⊔n y (≤-refl) max-bound₁ {x} {y} {z} x⊔y≡z rewrite sym x⊔y≡z rewrite ⊔-comm x y = m≤n⇒m≤o⊔n y (≤-refl)
max-bound₂ : (x : ) (y : ) (z : ) x y z y z max-bound₂ : {x y z : } x y z y z
max-bound₂ x y z x⊔y≡z rewrite sym x⊔y≡z = m≤n⇒m≤o⊔n x (≤-refl) max-bound₂ {x} {y} {z} x⊔y≡z rewrite sym x⊔y≡z = m≤n⇒m≤o⊔n x (≤-refl)
NatMinSemilattice : Semilattice max-least : (x y z : ) x y z (z' : ) (x z' × y z') z z'
NatMinSemilattice = record max-least x y z x⊔y≡z z' (x≤z' , y≤z') with (⊔-sel x y)
... | inj₁ x⊔y≡x rewrite trans (sym x⊔y≡z) (x⊔y≡x) = x≤z'
... | inj₂ x⊔y≡y rewrite trans (sym x⊔y≡z) (x⊔y≡y) = y≤z'
NatMaxSemilattice : Semilattice
NatMaxSemilattice = record
{ _≼_ = _≤_ { _≼_ = _≤_
; _⊔_ = _⊔_ ; _⊔_ = _⊔_
; isSemilattice = record ; isSemilattice = record
@ -109,6 +122,34 @@ private module NatInstances where
; ⊔-assoc = ⊔-assoc ; ⊔-assoc = ⊔-assoc
; ⊔-comm = ⊔-comm ; ⊔-comm = ⊔-comm
; ⊔-idemp = ⊔-idem ; ⊔-idemp = ⊔-idem
; ⊔-bound = λ x y z x⊔y≡z (max-bound₁ x y z x⊔y≡z , max-bound₂ x y z x⊔y≡z) ; ⊔-bound = λ x y z x⊔y≡z (max-bound₁ x⊔y≡z , max-bound₂ x⊔y≡z)
; ⊔-least = max-least
}
}
private
min-bound₁ : {x y z : } x y z z x
min-bound₁ {x} {y} {z} x⊓y≡z rewrite sym x⊓y≡z = m≤n⇒m⊓o≤n y (≤-refl)
min-bound₂ : {x y z : } x y z z y
min-bound₂ {x} {y} {z} x⊓y≡z rewrite sym x⊓y≡z rewrite ⊓-comm x y = m≤n⇒m⊓o≤n x (≤-refl)
min-greatest : (x y z : ) x y z (z' : ) (z' x × z' y) z' z
min-greatest x y z x⊓y≡z z' (z'≤x , z'≤y) with (⊓-sel x y)
... | inj₁ x⊓y≡x rewrite trans (sym x⊓y≡z) (x⊓y≡x) = z'≤x
... | inj₂ x⊓y≡y rewrite trans (sym x⊓y≡z) (x⊓y≡y) = z'≤y
NatMinSemilattice : Semilattice
NatMinSemilattice = record
{ _≼_ = _≥_
; _⊔_ = _⊓_
; isSemilattice = record
{ isPreorder = isPreorderFlip (Preorder.isPreorder NatPreorder)
; ⊔-assoc = ⊓-assoc
; ⊔-comm = ⊓-comm
; ⊔-idemp = ⊓-idem
; ⊔-bound = λ x y z x⊓y≡z (min-bound₁ x⊓y≡z , min-bound₂ x⊓y≡z)
; ⊔-least = min-greatest
} }
} }