Skip to content

Proofs in Data.Vec.Properties take general properties as inputs #2421

Open
@mildsunrise

Description

@mildsunrise

Looking at Data.Vec.Properties, I see many properties accept an irrelevant eq parameter that is actually a well-known property on naturals and has nothing relevant to the particular input. For example:

++-∷ʳ :  .(eq : suc (m + n) ≡ m + suc n) z (xs : Vec A m) (ys : Vec A n) 
        cast eq ((xs ++ ys) ∷ʳ z) ≡ xs ++ (ys ∷ʳ z)
++-∷ʳ {m = zero}  eq z []       []       = refl
++-∷ʳ {m = zero}  eq z []       (y ∷ ys) = cong (y ∷_) (++-∷ʳ refl z [] ys)
++-∷ʳ {m = suc m} eq z (x ∷ xs) ys       = cong (x ∷_) (++-∷ʳ (cong pred eq) z xs ys)

Instead of requesting a proof for suc (m + n) ≡ m + suc n, we could fill it in ourselves with +-suc which is already in scope:

++-∷ʳ :  z (xs : Vec A m) (ys : Vec A n)  let eq = sym (+-suc m n) in
        cast eq ((xs ++ ys) ∷ʳ z) ≡ xs ++ (ys ∷ʳ z)
++-∷ʳ {m = zero}  z []       []       = refl
++-∷ʳ {m = zero}  z []       (y ∷ ys) = cong (y ∷_) (++-∷ʳ′ z [] ys)
++-∷ʳ {m = suc m} z (x ∷ xs) ys       = cong (x ∷_) (++-∷ʳ′ z xs ys)

This not only makes ++-∷ʳ easier to use, it also simplifies the code for the proof itself (no extra eq argument, no need for refl and cong pred eq). And because eq is irrelevant in cast, we don't lose generality -- the original statement is just const ++-∷ʳ:

++-∷ʳ′ :  .(eq : suc (m + n) ≡ m + suc n) z (xs : Vec A m) (ys : Vec A n) 
         cast eq ((xs ++ ys) ∷ʳ z) ≡ xs ++ (ys ∷ʳ z)
++-∷ʳ′ _ = ++-∷ʳ

So I wanted to ask, what's the motivation for writing properties the current way? Was it because of limitations of Agda back when these proofs were written (wrt irrelevancy and so on)? Am I myself missing limitations of irrelevancy or something else? Later in the file I can see properties that use the second style:

fromList-reverse : (xs : List A)  cast (List.length-reverse xs) (fromList (List.reverse xs)) ≡ reverse (fromList xs)

And the follow up question is: would it be a good idea to change the properties to the second style in the next compat-breaking version?

Metadata

Metadata

Assignees

No one assigned

    Type

    No type

    Projects

    No projects

    Milestone

    No milestone

    Relationships

    None yet

    Development

    No branches or pull requests

    Issue actions