module Main where open import IO open import Level open import Data.Nat.Show using (show) open import Data.List using (List; _∷_; []) open import Data.String using (String) renaming (_≟_ to _≟ˢ_) open import Data.Unit using (⊤; tt) renaming (_≟_ to _≟ᵘ_) open import Data.List.Relation.Unary.All using (_∷_; []) open import Relation.Binary.PropositionalEquality as Eq using (_≡_; sym; subst; refl; trans) open import Relation.Nullary using (¬_) open import Utils using (Unique; push; empty) xyzw : List String xyzw = "x" ∷ "y" ∷ "z" ∷ "w" ∷ [] xyzw-Unique : Unique xyzw xyzw-Unique = push ((λ ()) ∷ (λ ()) ∷ (λ ()) ∷ []) (push ((λ ()) ∷ (λ ()) ∷ []) (push ((λ ()) ∷ []) (push [] empty))) open import Lattice using (IsFiniteHeightLattice; FiniteHeightLattice) open import Lattice.AboveBelow ⊤ _≡_ (record { ≈-refl = refl; ≈-sym = sym; ≈-trans = trans }) _≟ᵘ_ as AB renaming (≈-dec to ≈ᵘ-dec) open AB.Plain renaming (finiteHeightLattice to finiteHeightLatticeᵘ) open import Lattice.Bundles.FiniteValueMap String AB.AboveBelow _≟ˢ_ renaming (finiteHeightLattice to finiteHeightLatticeᵐ) fhlᵘ = finiteHeightLatticeᵘ (Data.Unit.tt) fhlⁱᵖ = finiteHeightLatticeᵐ fhlᵘ xyzw-Unique ≈ᵘ-dec main = run {0ℓ} (putStrLn (show (FiniteHeightLattice.height fhlⁱᵖ)))