2024-03-01 21:58:58 -08:00
|
|
|
|
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)
|
|
|
|
|
|
2024-03-09 21:46:15 -08:00
|
|
|
|
module FromFiniteHeightLattice (fhB : FiniteHeightLattice B)
|
|
|
|
|
{ks : List A} (uks : Unique ks)
|
|
|
|
|
(≈₂-dec : Decidable (FiniteHeightLattice._≈_ fhB)) where
|
|
|
|
|
|
2024-03-01 21:58:58 -08:00
|
|
|
|
open Lattice.FiniteHeightLattice fhB using () renaming
|
|
|
|
|
( _≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_
|
|
|
|
|
; height to height₂
|
|
|
|
|
; isLattice to isLattice₂
|
|
|
|
|
; fixedHeight to fixedHeight₂
|
|
|
|
|
)
|
|
|
|
|
|
2024-03-09 21:46:15 -08:00
|
|
|
|
import Lattice.FiniteMap
|
2024-03-10 18:35:29 -07:00
|
|
|
|
module FM = Lattice.FiniteMap ≡-dec-A isLattice₂
|
2024-03-09 21:46:15 -08:00
|
|
|
|
open FM.WithKeys ks public
|
2024-03-01 23:27:49 -08:00
|
|
|
|
|
2024-03-09 21:46:15 -08:00
|
|
|
|
import Lattice.FiniteValueMap
|
2024-03-10 18:35:29 -07:00
|
|
|
|
module FVM = Lattice.FiniteValueMap ≡-dec-A isLattice₂
|
2024-03-09 21:46:15 -08:00
|
|
|
|
open FVM.IterProdIsomorphism.WithUniqueKeysAndFixedHeight uks ≈₂-dec height₂ fixedHeight₂ public
|
2024-03-01 23:27:49 -08:00
|
|
|
|
|
2024-03-09 21:46:15 -08:00
|
|
|
|
≈-dec = ≈₂-dec⇒≈-dec ≈₂-dec
|