Remove need for explicit arguments in map derivatives
Signed-off-by: Danila Fedorin <danila.fedorin@gmail.com>
This commit is contained in:
parent
f21ebdcf46
commit
3305de4710
|
@ -26,7 +26,7 @@ data Expr : Set where
|
||||||
data Stmt : Set where
|
data Stmt : Set where
|
||||||
_←_ : String → Expr → Stmt
|
_←_ : String → Expr → Stmt
|
||||||
|
|
||||||
open import Lattice.MapSet String _≟ˢ_
|
open import Lattice.MapSet _≟ˢ_
|
||||||
renaming
|
renaming
|
||||||
( MapSet to StringSet
|
( MapSet to StringSet
|
||||||
; insert to insertˢ
|
; insert to insertˢ
|
||||||
|
|
|
@ -20,11 +20,11 @@ module FromFiniteHeightLattice (fhB : FiniteHeightLattice B)
|
||||||
)
|
)
|
||||||
|
|
||||||
import Lattice.FiniteMap
|
import Lattice.FiniteMap
|
||||||
module FM = Lattice.FiniteMap A B _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A isLattice₂
|
module FM = Lattice.FiniteMap ≡-dec-A isLattice₂
|
||||||
open FM.WithKeys ks public
|
open FM.WithKeys ks public
|
||||||
|
|
||||||
import Lattice.FiniteValueMap
|
import Lattice.FiniteValueMap
|
||||||
module FVM = Lattice.FiniteValueMap A B _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A isLattice₂
|
module FVM = Lattice.FiniteValueMap ≡-dec-A isLattice₂
|
||||||
open FVM.IterProdIsomorphism.WithUniqueKeysAndFixedHeight uks ≈₂-dec height₂ fixedHeight₂ public
|
open FVM.IterProdIsomorphism.WithUniqueKeysAndFixedHeight uks ≈₂-dec height₂ fixedHeight₂ public
|
||||||
|
|
||||||
≈-dec = ≈₂-dec⇒≈-dec ≈₂-dec
|
≈-dec = ≈₂-dec⇒≈-dec ≈₂-dec
|
||||||
|
|
|
@ -4,14 +4,14 @@ open import Relation.Binary.PropositionalEquality as Eq
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
open import Data.List using (List; _∷_; [])
|
open import Data.List using (List; _∷_; [])
|
||||||
|
|
||||||
module Lattice.FiniteMap {a b : Level} (A : Set a) (B : Set b)
|
module Lattice.FiniteMap {a b : Level} {A : Set a} {B : Set b}
|
||||||
(_≈₂_ : B → B → Set b)
|
{_≈₂_ : B → B → Set b}
|
||||||
(_⊔₂_ : B → B → B) (_⊓₂_ : B → B → B)
|
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
||||||
(≡-dec-A : IsDecidable (_≡_ {a} {A}))
|
(≡-dec-A : IsDecidable (_≡_ {a} {A}))
|
||||||
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
||||||
|
|
||||||
open IsLattice lB using () renaming (_≼_ to _≼₂_)
|
open IsLattice lB using () renaming (_≼_ to _≼₂_)
|
||||||
open import Lattice.Map A B _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A lB as Map
|
open import Lattice.Map ≡-dec-A lB as Map
|
||||||
using (Map; ⊔-equal-keys; ⊓-equal-keys; ∈k-dec; m₁≼m₂⇒m₁[k]≼m₂[k])
|
using (Map; ⊔-equal-keys; ⊓-equal-keys; ∈k-dec; m₁≼m₂⇒m₁[k]≼m₂[k])
|
||||||
renaming
|
renaming
|
||||||
( _≈_ to _≈ᵐ_
|
( _≈_ to _≈ᵐ_
|
||||||
|
|
|
@ -10,9 +10,9 @@ open import Relation.Binary.Definitions using (Decidable)
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
open import Function.Definitions using (Inverseˡ; Inverseʳ)
|
open import Function.Definitions using (Inverseˡ; Inverseʳ)
|
||||||
|
|
||||||
module Lattice.FiniteValueMap (A : Set) (B : Set)
|
module Lattice.FiniteValueMap {A : Set} {B : Set}
|
||||||
(_≈₂_ : B → B → Set)
|
{_≈₂_ : B → B → Set}
|
||||||
(_⊔₂_ : B → B → B) (_⊓₂_ : B → B → B)
|
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
||||||
(≡-dec-A : Decidable (_≡_ {_} {A}))
|
(≡-dec-A : Decidable (_≡_ {_} {A}))
|
||||||
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
||||||
|
|
||||||
|
@ -29,7 +29,7 @@ open import Data.List.Relation.Unary.Any using (Any; here; there)
|
||||||
open import Relation.Nullary using (¬_)
|
open import Relation.Nullary using (¬_)
|
||||||
open import Isomorphism using (IsInverseˡ; IsInverseʳ)
|
open import Isomorphism using (IsInverseˡ; IsInverseʳ)
|
||||||
|
|
||||||
open import Lattice.Map A B _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A lB
|
open import Lattice.Map ≡-dec-A lB
|
||||||
using
|
using
|
||||||
( subset-impl
|
( subset-impl
|
||||||
; locate; forget
|
; locate; forget
|
||||||
|
@ -40,7 +40,7 @@ open import Lattice.Map A B _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A lB
|
||||||
; in₁; in₂; bothᵘ; single
|
; in₁; in₂; bothᵘ; single
|
||||||
; ⊔-combines
|
; ⊔-combines
|
||||||
)
|
)
|
||||||
open import Lattice.FiniteMap A B _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A lB public
|
open import Lattice.FiniteMap ≡-dec-A lB public
|
||||||
|
|
||||||
module IterProdIsomorphism where
|
module IterProdIsomorphism where
|
||||||
open import Data.Unit using (⊤; tt)
|
open import Data.Unit using (⊤; tt)
|
||||||
|
|
|
@ -3,9 +3,9 @@ open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym;
|
||||||
open import Relation.Binary.Definitions using (Decidable)
|
open import Relation.Binary.Definitions using (Decidable)
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
|
|
||||||
module Lattice.Map {a b : Level} (A : Set a) (B : Set b)
|
module Lattice.Map {a b : Level} {A : Set a} {B : Set b}
|
||||||
(_≈₂_ : B → B → Set b)
|
{_≈₂_ : B → B → Set b}
|
||||||
(_⊔₂_ : B → B → B) (_⊓₂_ : B → B → B)
|
{_⊔₂_ : B → B → B} {_⊓₂_ : B → B → B}
|
||||||
(≡-dec-A : Decidable (_≡_ {a} {A}))
|
(≡-dec-A : Decidable (_≡_ {a} {A}))
|
||||||
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
(lB : IsLattice B _≈₂_ _⊔₂_ _⊓₂_) where
|
||||||
|
|
||||||
|
|
|
@ -3,7 +3,7 @@ open import Relation.Binary.PropositionalEquality as Eq using (_≡_; refl; sym;
|
||||||
open import Relation.Binary.Definitions using (Decidable)
|
open import Relation.Binary.Definitions using (Decidable)
|
||||||
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
open import Agda.Primitive using (Level) renaming (_⊔_ to _⊔ℓ_)
|
||||||
|
|
||||||
module Lattice.MapSet {a : Level} (A : Set a) (≡-dec-A : Decidable (_≡_ {a} {A})) where
|
module Lattice.MapSet {a : Level} {A : Set a} (≡-dec-A : Decidable (_≡_ {a} {A})) where
|
||||||
|
|
||||||
open import Data.List using (List; map)
|
open import Data.List using (List; map)
|
||||||
open import Data.Product using (_,_; proj₁)
|
open import Data.Product using (_,_; proj₁)
|
||||||
|
@ -12,7 +12,7 @@ open import Function using (_∘_)
|
||||||
open import Lattice.Unit using (⊤; tt) renaming (_≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_; isLattice to ⊤-isLattice)
|
open import Lattice.Unit using (⊤; tt) renaming (_≈_ to _≈₂_; _⊔_ to _⊔₂_; _⊓_ to _⊓₂_; isLattice to ⊤-isLattice)
|
||||||
import Lattice.Map
|
import Lattice.Map
|
||||||
|
|
||||||
private module UnitMap = Lattice.Map A ⊤ _≈₂_ _⊔₂_ _⊓₂_ ≡-dec-A ⊤-isLattice
|
private module UnitMap = Lattice.Map ≡-dec-A ⊤-isLattice
|
||||||
open UnitMap
|
open UnitMap
|
||||||
using (Map; Expr; ⟦_⟧)
|
using (Map; Expr; ⟦_⟧)
|
||||||
renaming
|
renaming
|
||||||
|
|
Loading…
Reference in New Issue
Block a user