agda-spa/Lattice/MapSet.agda
Danila Fedorin 3c346dcd15 Add a 'set' lattice backed by maps
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
2023-09-23 17:12:12 -07:00

20 lines
779 B
Agda
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.

open import Lattice
open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym; trans; cong; subst)
open import Relation.Binary.Definitions using (Decidable)
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔_)
module Lattice.MapSet {a : Level} (A : Set a) (≡-dec-A : Decidable (_≡_ {a} {A})) where
open import Lattice.Unit using () renaming (_≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_; isLattice to -isLattice)
import Lattice.Map
private module UnitMap = Lattice.Map A _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A -isLattice
open UnitMap using (Map)
open UnitMap using
( _⊆_; _≈_; ≈-equiv; _⊔_; _⊓_
; isUnionSemilattice; isIntersectSemilattice; isLattice
) public
MapSet : Set a
MapSet = Map