2023-04-04 21:08:31 -07:00
|
|
|
|
module Lattice where
|
|
|
|
|
|
2023-09-02 20:36:12 -07:00
|
|
|
|
open import Equivalence
|
2023-09-03 21:05:57 -07:00
|
|
|
|
import Chain
|
2023-09-02 20:36:12 -07:00
|
|
|
|
|
2023-09-03 21:05:57 -07:00
|
|
|
|
open import Relation.Binary.Core using (_Preserves_⟶_ )
|
2023-09-23 16:39:11 -07:00
|
|
|
|
open import Relation.Nullary using (Dec; ¬_)
|
|
|
|
|
open import Data.Nat as Nat using (ℕ)
|
|
|
|
|
open import Data.Product using (_×_; Σ; _,_)
|
2023-07-13 21:50:27 -07:00
|
|
|
|
open import Data.Sum using (_⊎_; inj₁; inj₂)
|
2023-08-20 20:49:08 -07:00
|
|
|
|
open import Agda.Primitive using (lsuc; Level) renaming (_⊔_ to _⊔ℓ_)
|
|
|
|
|
open import Function.Definitions using (Injective)
|
2023-09-03 23:56:39 -07:00
|
|
|
|
|
2023-09-17 19:43:24 -07:00
|
|
|
|
IsDecidable : ∀ {a} {A : Set a} (R : A → A → Set a) → Set a
|
|
|
|
|
IsDecidable {a} {A} R = ∀ (a₁ a₂ : A) → Dec (R a₁ a₂)
|
2023-08-20 18:35:57 -07:00
|
|
|
|
|
2024-03-01 23:26:25 -08:00
|
|
|
|
module _ {a b} {A : Set a} {B : Set b}
|
|
|
|
|
(_≼₁_ : A → A → Set a) (_≼₂_ : B → B → Set b) where
|
|
|
|
|
|
|
|
|
|
Monotonic : (A → B) → Set (a ⊔ℓ b)
|
|
|
|
|
Monotonic f = ∀ {a₁ a₂ : A} → a₁ ≼₁ a₂ → f a₁ ≼₂ f a₂
|
|
|
|
|
|
2023-07-15 14:40:11 -07:00
|
|
|
|
record IsSemilattice {a} (A : Set a)
|
|
|
|
|
(_≈_ : A → A → Set a)
|
|
|
|
|
(_⊔_ : A → A → A) : Set a where
|
|
|
|
|
|
2023-08-20 19:57:26 -07:00
|
|
|
|
_≼_ : A → A → Set a
|
2024-02-07 21:04:13 -08:00
|
|
|
|
a ≼ b = (a ⊔ b) ≈ b
|
2023-08-20 19:57:26 -07:00
|
|
|
|
|
|
|
|
|
_≺_ : A → A → Set a
|
|
|
|
|
a ≺ b = (a ≼ b) × (¬ a ≈ b)
|
|
|
|
|
|
2023-04-06 23:08:49 -07:00
|
|
|
|
field
|
2023-07-15 15:16:51 -07:00
|
|
|
|
≈-equiv : IsEquivalence A _≈_
|
2023-09-03 17:08:37 -07:00
|
|
|
|
≈-⊔-cong : ∀ {a₁ a₂ a₃ a₄} → a₁ ≈ a₂ → a₃ ≈ a₄ → (a₁ ⊔ a₃) ≈ (a₂ ⊔ a₄)
|
2023-07-15 15:16:51 -07:00
|
|
|
|
|
2023-07-15 14:40:11 -07:00
|
|
|
|
⊔-assoc : (x y z : A) → ((x ⊔ y) ⊔ z) ≈ (x ⊔ (y ⊔ z))
|
|
|
|
|
⊔-comm : (x y : A) → (x ⊔ y) ≈ (y ⊔ x)
|
|
|
|
|
⊔-idemp : (x : A) → (x ⊔ x) ≈ x
|
|
|
|
|
|
2023-09-03 19:33:04 -07:00
|
|
|
|
open IsEquivalence ≈-equiv public
|
|
|
|
|
|
2024-02-07 21:04:13 -08:00
|
|
|
|
open import Relation.Binary.Reasoning.Base.Single _≈_ ≈-refl ≈-trans
|
|
|
|
|
|
2024-03-01 23:26:25 -08:00
|
|
|
|
⊔-Monotonicˡ : ∀ (a₁ : A) → Monotonic _≼_ _≼_ (λ a₂ → a₁ ⊔ a₂)
|
|
|
|
|
⊔-Monotonicˡ a {a₁} {a₂} a₁≼a₂ = ≈-trans (≈-sym lhs) (≈-⊔-cong (≈-refl {a}) a₁≼a₂)
|
|
|
|
|
where
|
|
|
|
|
lhs =
|
|
|
|
|
begin
|
|
|
|
|
a ⊔ (a₁ ⊔ a₂)
|
|
|
|
|
∼⟨ ≈-⊔-cong (≈-sym (⊔-idemp _)) ≈-refl ⟩
|
|
|
|
|
(a ⊔ a) ⊔ (a₁ ⊔ a₂)
|
|
|
|
|
∼⟨ ⊔-assoc _ _ _ ⟩
|
|
|
|
|
a ⊔ (a ⊔ (a₁ ⊔ a₂))
|
|
|
|
|
∼⟨ ≈-⊔-cong ≈-refl (≈-sym (⊔-assoc _ _ _)) ⟩
|
|
|
|
|
a ⊔ ((a ⊔ a₁) ⊔ a₂)
|
|
|
|
|
∼⟨ ≈-⊔-cong ≈-refl (≈-⊔-cong (⊔-comm _ _) ≈-refl) ⟩
|
|
|
|
|
a ⊔ ((a₁ ⊔ a) ⊔ a₂)
|
|
|
|
|
∼⟨ ≈-⊔-cong ≈-refl (⊔-assoc _ _ _) ⟩
|
|
|
|
|
a ⊔ (a₁ ⊔ (a ⊔ a₂))
|
|
|
|
|
∼⟨ ≈-sym (⊔-assoc _ _ _) ⟩
|
|
|
|
|
(a ⊔ a₁) ⊔ (a ⊔ a₂)
|
|
|
|
|
∎
|
|
|
|
|
|
2024-03-03 16:51:57 -08:00
|
|
|
|
⊔-Monotonicʳ : ∀ (a₂ : A) → Monotonic _≼_ _≼_ (λ a₁ → a₁ ⊔ a₂)
|
|
|
|
|
⊔-Monotonicʳ a {a₁} {a₂} a₁≼a₂ = ≈-trans (≈-sym lhs) (≈-⊔-cong a₁≼a₂ (≈-refl {a}))
|
|
|
|
|
where
|
|
|
|
|
lhs =
|
|
|
|
|
begin
|
|
|
|
|
(a₁ ⊔ a₂) ⊔ a
|
|
|
|
|
∼⟨ ≈-⊔-cong ≈-refl (≈-sym (⊔-idemp _)) ⟩
|
|
|
|
|
(a₁ ⊔ a₂) ⊔ (a ⊔ a)
|
|
|
|
|
∼⟨ ≈-sym (⊔-assoc _ _ _) ⟩
|
|
|
|
|
((a₁ ⊔ a₂) ⊔ a) ⊔ a
|
|
|
|
|
∼⟨ ≈-⊔-cong (⊔-assoc _ _ _) ≈-refl ⟩
|
|
|
|
|
(a₁ ⊔ (a₂ ⊔ a)) ⊔ a
|
|
|
|
|
∼⟨ ≈-⊔-cong (≈-⊔-cong ≈-refl (⊔-comm _ _)) ≈-refl ⟩
|
|
|
|
|
(a₁ ⊔ (a ⊔ a₂)) ⊔ a
|
|
|
|
|
∼⟨ ≈-⊔-cong (≈-sym (⊔-assoc _ _ _)) ≈-refl ⟩
|
|
|
|
|
((a₁ ⊔ a) ⊔ a₂) ⊔ a
|
|
|
|
|
∼⟨ ⊔-assoc _ _ _ ⟩
|
|
|
|
|
(a₁ ⊔ a) ⊔ (a₂ ⊔ a)
|
|
|
|
|
∎
|
|
|
|
|
|
2023-08-20 21:53:27 -07:00
|
|
|
|
≼-refl : ∀ (a : A) → a ≼ a
|
2024-02-07 21:04:13 -08:00
|
|
|
|
≼-refl a = ⊔-idemp a
|
2023-08-20 21:53:27 -07:00
|
|
|
|
|
2023-09-03 19:33:04 -07:00
|
|
|
|
≼-cong : ∀ {a₁ a₂ a₃ a₄ : A} → a₁ ≈ a₂ → a₃ ≈ a₄ → a₁ ≼ a₃ → a₂ ≼ a₄
|
2024-02-07 21:04:13 -08:00
|
|
|
|
≼-cong {a₁} {a₂} {a₃} {a₄} a₁≈a₂ a₃≈a₄ a₁⊔a₃≈a₃ =
|
|
|
|
|
begin
|
|
|
|
|
a₂ ⊔ a₄
|
|
|
|
|
∼⟨ ≈-⊔-cong (≈-sym a₁≈a₂) (≈-sym a₃≈a₄) ⟩
|
|
|
|
|
a₁ ⊔ a₃
|
|
|
|
|
∼⟨ a₁⊔a₃≈a₃ ⟩
|
|
|
|
|
a₃
|
|
|
|
|
∼⟨ a₃≈a₄ ⟩
|
|
|
|
|
a₄
|
|
|
|
|
∎
|
2023-09-03 19:33:04 -07:00
|
|
|
|
|
|
|
|
|
≺-cong : ∀ {a₁ a₂ a₃ a₄ : A} → a₁ ≈ a₂ → a₃ ≈ a₄ → a₁ ≺ a₃ → a₂ ≺ a₄
|
|
|
|
|
≺-cong a₁≈a₂ a₃≈a₄ (a₁≼a₃ , a₁̷≈a₃) =
|
|
|
|
|
( ≼-cong a₁≈a₂ a₃≈a₄ a₁≼a₃
|
|
|
|
|
, λ a₂≈a₄ → a₁̷≈a₃ (≈-trans a₁≈a₂ (≈-trans a₂≈a₄ (≈-sym a₃≈a₄)))
|
|
|
|
|
)
|
2023-07-15 15:16:51 -07:00
|
|
|
|
|
2024-03-06 00:35:06 -08:00
|
|
|
|
module _ {a} {A : Set a}
|
|
|
|
|
{_≈_ : A → A → Set a} {_⊔_ : A → A → A}
|
|
|
|
|
(lA : IsSemilattice A _≈_ _⊔_) where
|
|
|
|
|
|
|
|
|
|
open IsSemilattice lA using (_≼_)
|
|
|
|
|
|
|
|
|
|
id-Mono : Monotonic _≼_ _≼_ (λ x → x)
|
|
|
|
|
id-Mono a₁≼a₂ = a₁≼a₂
|
|
|
|
|
|
2024-03-03 17:23:57 -08:00
|
|
|
|
module _ {a b} {A : Set a} {B : Set b}
|
|
|
|
|
{_≈₁_ : A → A → Set a} {_⊔₁_ : A → A → A}
|
|
|
|
|
{_≈₂_ : B → B → Set b} {_⊔₂_ : B → B → B}
|
|
|
|
|
(lA : IsSemilattice A _≈₁_ _⊔₁_) (lB : IsSemilattice B _≈₂_ _⊔₂_) where
|
|
|
|
|
|
|
|
|
|
open IsSemilattice lA using () renaming (_≼_ to _≼₁_)
|
|
|
|
|
open IsSemilattice lB using () renaming (_≼_ to _≼₂_; ⊔-idemp to ⊔₂-idemp)
|
|
|
|
|
|
|
|
|
|
const-Mono : ∀ (x : B) → Monotonic _≼₁_ _≼₂_ (λ _ → x)
|
|
|
|
|
const-Mono x _ = ⊔₂-idemp x
|
|
|
|
|
|
2023-07-15 14:40:11 -07:00
|
|
|
|
record IsLattice {a} (A : Set a)
|
|
|
|
|
(_≈_ : A → A → Set a)
|
|
|
|
|
(_⊔_ : A → A → A)
|
|
|
|
|
(_⊓_ : A → A → A) : Set a where
|
2023-04-04 21:08:31 -07:00
|
|
|
|
|
|
|
|
|
field
|
2023-07-15 14:40:11 -07:00
|
|
|
|
joinSemilattice : IsSemilattice A _≈_ _⊔_
|
|
|
|
|
meetSemilattice : IsSemilattice A _≈_ _⊓_
|
2023-04-06 23:08:49 -07:00
|
|
|
|
|
2023-07-15 14:40:11 -07:00
|
|
|
|
absorb-⊔-⊓ : (x y : A) → (x ⊔ (x ⊓ y)) ≈ x
|
|
|
|
|
absorb-⊓-⊔ : (x y : A) → (x ⊓ (x ⊔ y)) ≈ x
|
2023-04-04 21:08:31 -07:00
|
|
|
|
|
2023-04-06 23:08:49 -07:00
|
|
|
|
open IsSemilattice joinSemilattice public
|
2023-08-20 19:57:26 -07:00
|
|
|
|
open IsSemilattice meetSemilattice public using () renaming
|
2023-04-06 23:08:49 -07:00
|
|
|
|
( ⊔-assoc to ⊓-assoc
|
|
|
|
|
; ⊔-comm to ⊓-comm
|
|
|
|
|
; ⊔-idemp to ⊓-idemp
|
2024-03-03 17:04:18 -08:00
|
|
|
|
; ⊔-Monotonicˡ to ⊓-Monotonicˡ
|
|
|
|
|
; ⊔-Monotonicʳ to ⊓-Monotonicʳ
|
2023-09-03 17:08:37 -07:00
|
|
|
|
; ≈-⊔-cong to ≈-⊓-cong
|
2023-08-20 21:53:27 -07:00
|
|
|
|
; _≼_ to _≽_
|
|
|
|
|
; _≺_ to _≻_
|
|
|
|
|
; ≼-refl to ≽-refl
|
2023-04-06 23:08:49 -07:00
|
|
|
|
)
|
|
|
|
|
|
2023-09-23 16:20:58 -07:00
|
|
|
|
FixedHeight : ∀ (h : ℕ) → Set a
|
|
|
|
|
FixedHeight h = Chain.Height (_≈_) ≈-equiv _≺_ ≺-cong h
|
|
|
|
|
|
2023-08-20 18:35:57 -07:00
|
|
|
|
record IsFiniteHeightLattice {a} (A : Set a)
|
|
|
|
|
(h : ℕ)
|
|
|
|
|
(_≈_ : A → A → Set a)
|
|
|
|
|
(_⊔_ : A → A → A)
|
|
|
|
|
(_⊓_ : A → A → A) : Set (lsuc a) where
|
|
|
|
|
|
|
|
|
|
field
|
|
|
|
|
isLattice : IsLattice A _≈_ _⊔_ _⊓_
|
|
|
|
|
|
|
|
|
|
open IsLattice isLattice public
|
|
|
|
|
|
2023-09-23 16:20:58 -07:00
|
|
|
|
field
|
|
|
|
|
fixedHeight : FixedHeight h
|
2023-09-03 19:33:04 -07:00
|
|
|
|
|
2023-08-20 21:53:27 -07:00
|
|
|
|
module ChainMapping {a b} {A : Set a} {B : Set b}
|
|
|
|
|
{_≈₁_ : A → A → Set a} {_≈₂_ : B → B → Set b}
|
|
|
|
|
{_⊔₁_ : A → A → A} {_⊔₂_ : B → B → B}
|
2023-08-20 20:49:08 -07:00
|
|
|
|
(slA : IsSemilattice A _≈₁_ _⊔₁_) (slB : IsSemilattice B _≈₂_ _⊔₂_) where
|
|
|
|
|
|
2023-09-03 21:05:57 -07:00
|
|
|
|
open IsSemilattice slA renaming (_≼_ to _≼₁_; _≺_ to _≺₁_; ≈-equiv to ≈₁-equiv; ≺-cong to ≺₁-cong)
|
|
|
|
|
open IsSemilattice slB renaming (_≼_ to _≼₂_; _≺_ to _≺₂_; ≈-equiv to ≈₂-equiv; ≺-cong to ≺₂-cong)
|
2023-08-20 20:49:08 -07:00
|
|
|
|
|
2023-09-03 21:05:57 -07:00
|
|
|
|
open Chain _≈₁_ ≈₁-equiv _≺₁_ ≺₁-cong using () renaming (Chain to Chain₁; step to step₁; done to done₁)
|
|
|
|
|
open Chain _≈₂_ ≈₂-equiv _≺₂_ ≺₂-cong using () renaming (Chain to Chain₂; step to step₂; done to done₂)
|
|
|
|
|
|
|
|
|
|
Chain-map : ∀ (f : A → B) →
|
|
|
|
|
Monotonic _≼₁_ _≼₂_ f →
|
|
|
|
|
Injective _≈₁_ _≈₂_ f →
|
|
|
|
|
f Preserves _≈₁_ ⟶ _≈₂_ →
|
|
|
|
|
∀ {a₁ a₂ : A} {n : ℕ} → Chain₁ a₁ a₂ n → Chain₂ (f a₁) (f a₂) n
|
|
|
|
|
Chain-map f Monotonicᶠ Injectiveᶠ Preservesᶠ (done₁ a₁≈a₂) =
|
|
|
|
|
done₂ (Preservesᶠ a₁≈a₂)
|
|
|
|
|
Chain-map f Monotonicᶠ Injectiveᶠ Preservesᶠ (step₁ (a₁≼₁a , a₁̷≈₁a) a≈₁a' a'a₂) =
|
2023-08-20 20:49:08 -07:00
|
|
|
|
let fa₁≺₂fa = (Monotonicᶠ a₁≼₁a , λ fa₁≈₂fa → a₁̷≈₁a (Injectiveᶠ fa₁≈₂fa))
|
2023-09-03 21:05:57 -07:00
|
|
|
|
fa≈fa' = Preservesᶠ a≈₁a'
|
|
|
|
|
in step₂ fa₁≺₂fa fa≈fa' (Chain-map f Monotonicᶠ Injectiveᶠ Preservesᶠ a'a₂)
|
2023-08-20 20:49:08 -07:00
|
|
|
|
|
2023-07-15 13:12:21 -07:00
|
|
|
|
record Semilattice {a} (A : Set a) : Set (lsuc a) where
|
|
|
|
|
field
|
2023-07-15 14:40:11 -07:00
|
|
|
|
_≈_ : A → A → Set a
|
2023-07-15 13:12:21 -07:00
|
|
|
|
_⊔_ : A → A → A
|
|
|
|
|
|
2023-07-15 14:40:11 -07:00
|
|
|
|
isSemilattice : IsSemilattice A _≈_ _⊔_
|
2023-07-15 13:12:21 -07:00
|
|
|
|
|
|
|
|
|
open IsSemilattice isSemilattice public
|
|
|
|
|
|
2023-04-06 23:08:49 -07:00
|
|
|
|
|
|
|
|
|
record Lattice {a} (A : Set a) : Set (lsuc a) where
|
|
|
|
|
field
|
2023-07-15 14:40:11 -07:00
|
|
|
|
_≈_ : A → A → Set a
|
2023-04-06 23:08:49 -07:00
|
|
|
|
_⊔_ : A → A → A
|
|
|
|
|
_⊓_ : A → A → A
|
|
|
|
|
|
2023-07-15 14:40:11 -07:00
|
|
|
|
isLattice : IsLattice A _≈_ _⊔_ _⊓_
|
2023-04-06 23:08:49 -07:00
|
|
|
|
|
|
|
|
|
open IsLattice isLattice public
|
2024-02-11 14:17:49 -08:00
|
|
|
|
|
|
|
|
|
record FiniteHeightLattice {a} (A : Set a) : Set (lsuc a) where
|
|
|
|
|
field
|
|
|
|
|
height : ℕ
|
|
|
|
|
_≈_ : A → A → Set a
|
|
|
|
|
_⊔_ : A → A → A
|
|
|
|
|
_⊓_ : A → A → A
|
|
|
|
|
|
2024-02-11 15:37:00 -08:00
|
|
|
|
isFiniteHeightLattice : IsFiniteHeightLattice A height _≈_ _⊔_ _⊓_
|
2024-02-11 14:17:49 -08:00
|
|
|
|
|
|
|
|
|
open IsFiniteHeightLattice isFiniteHeightLattice public
|