module NatMap where open import Agda.Primitive using (Level) open import Data.Nat using (ℕ; _ NatMap A insert n a [] = (n , a) ∷ [] insert n a l@(x@(n' , a') ∷ xs) with n A -> A) -> NatMap A -> NatMap A -> NatMap A merge _ [] m₂ = m₂ merge _ m₁ [] = m₁ merge f m₁@(x₁@(n₁ , a₁) ∷ xs₁) m₂@(x₂@(n₂ , a₂) ∷ xs₂) with <-cmp n₁ n₂ ... | Tri.tri< _ _ _ = x₁ ∷ merge f xs₁ m₂ ... | Tri.tri> _ _ _ = x₂ ∷ merge f m₁ xs₂ ... | Tri.tri≈ _ _ _ = (n₁ , f a₁ a₂) ∷ merge f xs₁ xs₂ testMerge : merge (_++_) ((1 , "one") ∷ (2 , "two") ∷ []) ((2 , "two") ∷ (3 , "three") ∷ []) ≡ (1 , "one") ∷ (2 , "twotwo") ∷ (3 , "three") ∷ [] testMerge = refl -- Note to self: will it be tricky to define a Semilattice instance for this? -- I'm worried because proofs will likely require knowing the ordering -- merge certainly relies on it, but -- we can't prove commutativity if the NatMap type includes a proof of ordering, because then we'll need -- to compare `<=` for equality, which are going to be defined in terms of `=`... -- -- We might need to generalize Lattice and friends to use ≈. But then, -- pain in the ass for products because they'll need to lift ≈ ...