agda-spa/Lattice.agda

148 lines
5.5 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module Lattice where
open import Equivalence
import Chain
open import Relation.Binary.Core using (_Preserves_⟶_ )
open import Relation.Nullary using (Dec; ¬_)
open import Data.Nat as Nat using ()
open import Data.Product using (_×_; Σ; _,_)
open import Data.Sum using (_⊎_; inj₁; inj₂)
open import Agda.Primitive using (lsuc; Level) renaming (_⊔_ to _⊔_)
open import Function.Definitions using (Injective)
IsDecidable : ∀ {a} {A : Set a} (R : A → A → Set a) → Set a
IsDecidable {a} {A} R = ∀ (a₁ a₂ : A) → Dec (R a₁ a₂)
record IsSemilattice {a} (A : Set a)
(_≈_ : A → A → Set a)
(_⊔_ : A → A → A) : Set a where
_≼_ : A → A → Set a
a ≼ b = (a ⊔ b) ≈ b
_≺_ : A → A → Set a
a ≺ b = (a ≼ b) × (¬ a ≈ b)
field
≈-equiv : IsEquivalence A _≈_
≈-⊔-cong : ∀ {a₁ a₂ a₃ a₄} → a₁ ≈ a₂ → a₃ ≈ a₄ → (a₁ ⊔ a₃) ≈ (a₂ ⊔ a₄)
⊔-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
open IsEquivalence ≈-equiv public
open import Relation.Binary.Reasoning.Base.Single _≈_ ≈-refl ≈-trans
≼-refl : ∀ (a : A) → a ≼ a
≼-refl a = ⊔-idemp a
≼-cong : ∀ {a₁ a₂ a₃ a₄ : A} → a₁ ≈ a₂ → a₃ ≈ a₄ → a₁ ≼ a₃ → a₂ ≼ a₄
≼-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₄
≺-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₄)))
)
record IsLattice {a} (A : Set a)
(_≈_ : A → A → Set a)
(_⊔_ : A → A → A)
(_⊓_ : A → A → A) : Set a where
field
joinSemilattice : IsSemilattice A _≈_ _⊔_
meetSemilattice : IsSemilattice A _≈_ _⊓_
absorb-⊔-⊓ : (x y : A) → (x ⊔ (x ⊓ y)) ≈ x
absorb-⊓-⊔ : (x y : A) → (x ⊓ (x ⊔ y)) ≈ x
open IsSemilattice joinSemilattice public
open IsSemilattice meetSemilattice public using () renaming
( ⊔-assoc to ⊓-assoc
; ⊔-comm to ⊓-comm
; ⊔-idemp to ⊓-idemp
; ≈-⊔-cong to ≈-⊓-cong
; _≼_ to _≽_
; _≺_ to _≻_
; ≼-refl to ≽-refl
)
FixedHeight : ∀ (h : ) → Set a
FixedHeight h = Chain.Height (_≈_) ≈-equiv _≺_ ≺-cong h
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
field
fixedHeight : FixedHeight h
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₂
module ChainMapping {a b} {A : Set a} {B : Set b}
{_≈₁_ : A → A → Set a} {_≈₂_ : B → B → Set b}
{_⊔₁_ : A → A → A} {_⊔₂_ : B → B → B}
(slA : IsSemilattice A _≈₁_ _⊔₁_) (slB : IsSemilattice B _≈₂_ _⊔₂_) where
open IsSemilattice slA renaming (_≼_ to _≼₁_; _≺_ to _≺₁_; ≈-equiv to ≈₁-equiv; ≺-cong to ≺₁-cong)
open IsSemilattice slB renaming (_≼_ to _≼₂_; _≺_ to _≺₂_; ≈-equiv to ≈₂-equiv; ≺-cong to ≺₂-cong)
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₂) =
let fa₁≺₂fa = (Monotonicᶠ a₁≼₁a , λ fa₁≈₂fa → a₁̷≈₁a (Injectiveᶠ fa₁≈₂fa))
fa≈fa' = Preservesᶠ a≈₁a'
in step₂ fa₁≺₂fa fa≈fa' (Chain-map f Monotonicᶠ Injectiveᶠ Preservesᶠ a'a₂)
record Semilattice {a} (A : Set a) : Set (lsuc a) where
field
_≈_ : A → A → Set a
_⊔_ : A → A → A
isSemilattice : IsSemilattice A _≈_ _⊔_
open IsSemilattice isSemilattice public
record Lattice {a} (A : Set a) : Set (lsuc a) where
field
_≈_ : A → A → Set a
_⊔_ : A → A → A
_⊓_ : A → A → A
isLattice : IsLattice A _≈_ _⊔_ _⊓_
open IsLattice isLattice public