open import Relation.Binary.PropositionalEquality using (_≡_) open import Relation.Binary.Definitions using (Decidable) module Lattice.Bundles.FiniteValueMap (A B : Set) (≡-dec-A : Decidable (_≡_ {_} {A})) where open import Lattice open import Data.List using (List) open import Data.Nat using (ℕ) open import Utils using (Unique) module FromFiniteHeightLattice (fhB : FiniteHeightLattice B) {ks : List A} (uks : Unique ks) (≈₂-dec : Decidable (FiniteHeightLattice._≈_ fhB)) where open Lattice.FiniteHeightLattice fhB using () renaming ( _≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_ ; height to height₂ ; isLattice to isLattice₂ ; fixedHeight to fixedHeight₂ ) import Lattice.FiniteMap module FM = Lattice.FiniteMap ≡-dec-A isLattice₂ open FM.WithKeys ks public import Lattice.FiniteValueMap module FVM = Lattice.FiniteValueMap ≡-dec-A isLattice₂ open FVM.IterProdIsomorphism.WithUniqueKeysAndFixedHeight uks ≈₂-dec height₂ fixedHeight₂ public ≈-dec = ≈₂-dec⇒≈-dec ≈₂-dec