Move proof of least element into FiniteHeightLattice
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
25
Lattice.agda
25
Lattice.agda
@@ -4,7 +4,8 @@ open import Equivalence
|
||||
import Chain
|
||||
|
||||
open import Relation.Binary.Core using (_Preserves_⟶_ )
|
||||
open import Relation.Nullary using (Dec; ¬_)
|
||||
open import Relation.Nullary using (Dec; ¬_; yes; no)
|
||||
open import Data.Empty using (⊥-elim)
|
||||
open import Data.Nat as Nat using (ℕ)
|
||||
open import Data.Product using (_×_; Σ; _,_)
|
||||
open import Data.Sum using (_⊎_; inj₁; inj₂)
|
||||
@@ -236,6 +237,28 @@ record IsFiniteHeightLattice {a} (A : Set a)
|
||||
field
|
||||
{{fixedHeight}} : FixedHeight h
|
||||
|
||||
-- If the equality is decidable, we can prove that the top and bottom
|
||||
-- elements of the chain are least and greatest elements of the lattice
|
||||
module _ {{≈-Decidable : IsDecidable _≈_}} where
|
||||
open IsDecidable ≈-Decidable using () renaming (R-dec to ≈-dec)
|
||||
|
||||
module MyChain = Chain _≈_ ≈-equiv _≺_ ≺-cong
|
||||
open MyChain.Height fixedHeight using (⊥; ⊤) public
|
||||
open MyChain.Height fixedHeight using (bounded; longestChain)
|
||||
|
||||
⊥≼ : ∀ (a : A) → ⊥ ≼ a
|
||||
⊥≼ a with ≈-dec a ⊥
|
||||
... | yes a≈⊥ = ≼-cong a≈⊥ ≈-refl (≼-refl a)
|
||||
... | no a̷≈⊥ with ≈-dec ⊥ (a ⊓ ⊥)
|
||||
... | yes ⊥≈a⊓⊥ = ≈-trans (⊔-comm ⊥ a) (≈-trans (≈-⊔-cong (≈-refl {a}) ⊥≈a⊓⊥) (absorb-⊔-⊓ a ⊥))
|
||||
... | no ⊥ᴬ̷≈a⊓⊥ = ⊥-elim (MyChain.Bounded-suc-n bounded (MyChain.step x≺⊥ ≈-refl longestChain))
|
||||
where
|
||||
⊥⊓a̷≈⊥ : ¬ (⊥ ⊓ a) ≈ ⊥
|
||||
⊥⊓a̷≈⊥ = λ ⊥⊓a≈⊥ → ⊥ᴬ̷≈a⊓⊥ (≈-trans (≈-sym ⊥⊓a≈⊥) (⊓-comm _ _))
|
||||
|
||||
x≺⊥ : (⊥ ⊓ a) ≺ ⊥
|
||||
x≺⊥ = (≈-trans (⊔-comm _ _) (≈-trans (≈-refl {⊥ ⊔ (⊥ ⊓ a)}) (absorb-⊔-⊓ ⊥ a)) , ⊥⊓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}
|
||||
|
||||
Reference in New Issue
Block a user