[Agda] Proof involving associativity of ++ of vectors

Nils Anders Danielsson nad at Cs.Nott.AC.UK
Fri Oct 16 16:48:33 CEST 2009


On 2009-10-16 14:38, Shin-Cheng Mu wrote:
> So it seems that I need to perform a type cast:
> 
>    thm (x ∷ xs) zs pxs qzsxs =
>        subst (λ ws → Q ws) eq1 (thm xs (zs ++ [ x ]) pxs)   -- (*)
> 
> [...]
> 
> But I am not even allowed to do this type cast -- Agda complains
> about the type of ws in (*). Indeed, when we look at the type
> of subst:
> 
>   subst : ∀ {a} (Q : a → Set) → ∀ {x y} → x ≅ y → Q x → Q y
>
> [...]

I believe that the discussion in the following thread (Congruence and
substitutivity rules in heterogeneous equality) is relevant:

  http://thread.gmane.org/gmane.comp.lang.agda/954

In this case I suspect that I would define a custom "cast" function for
Q, or use "with" like Jean-Philippe proposed.

  cast : ∀ {A m n} {xs : Vec A m} {ys : Vec A n} →
         xs ≅ ys → Q xs → Q ys
  cast refl = id

-- 
/NAD


This message has been checked for viruses but the contents of an attachment
may still contain software viruses, which could damage your computer system:
you are advised to perform your own checks. Email communications with the
University of Nottingham may be monitored as permitted by UK legislation.



More information about the Agda mailing list