[Agda] trouble with with

Thorsten Altenkirch txa at Cs.Nott.AC.UK
Sun May 4 17:14:39 CEST 2008


Hi guys,

inspired by Nisse's recent elaborations http://sneezy.cs.nott.ac.uk/fplunch/weblog/?p=96 
  I have been playing around with Agda (again) but in the moment I  
have a problem with with I don't really understand. Maybe somebody can  
advise me - it could be something obvious.

I define the finite sets as usual:

  data Fin : ℕ -> Set where
    fz : forall {n} -> Fin (succ n)
    fs : forall {n} -> Fin n -> Fin (succ n)

and implementations of the injections - sums are sums after all (I  
realize that it is arguable whether the first two arguments should be  
implicit):

  finl : forall {m n} -> Fin m -> Fin (m + n)
  finl fz     = fz
  finl (fs i) = fs (finl i)

  finr : forall {m n} -> Fin n -> Fin (m + n)
  finr {zero}   {n} i = i
  finr {succ m} {n} i = fs (finr {m} i)

and then I wanted to implement the view from section 3 of http://www.cs.nott.ac.uk/~txa/publ/ssgp06.pdf
but something goes wrong:

  data PlusView : forall {m n} -> Fin (m + n) -> Set where
    inl : forall {m n} -> (i : Fin m) -> PlusView {m} {n} (finl {m}  
{n} i)
    inr : forall {m n} -> (i : Fin n) -> PlusView {m} {n} (finr {m}  
{n} i)

  plusView : forall {m n} -> (i : Fin (m + n)) -> PlusView {m} {n} i
  plusView {zero}   {n} i      = inr {zero} {n} i
  plusView {succ m} {n} fz     = inl {succ m} {n} (fz {m})
  plusView {succ m} {n} (fs i) with plusView {m} {n} i
  plusView {succ m} {n} .(fs (finl {m} {n} j)) | (inl {.m} {.n} j) = ?
  plusView {succ m} {n} .(fs (finr {m} {n} j)) | (inr {.m} {.n} j) = ?

Agda isn't happy about my with-pattern:

/Users/txa/current/example.agda:30,3-32,71
With clause pattern .fs finl {m} {n} j is not an instance of its
parent pattern
when checking that the clause
plusView {(succ m)} {n} (fs i) with plusView {m} {n} i
plusView {(succ m')} {n'} .fs finl {m} {n} j | (inl {.m} {.n} j)
  = ?
plusView {(succ m')} {n'} .fs finr {m} {n} j | (inr {.m} {.n} j)
  = ?
has type {m : ℕ} {n : ℕ} (i : Fin (m + n)) -> PlusView i

I am using
Loading package Agda-2.1.2 ... linking ... done.
btw. Here is the file.
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.

-------------- next part --------------
A non-text attachment was scrubbed...
Name: example.agda
Type: application/octet-stream
Size: 1067 bytes
Desc: not available
Url : https://lists.chalmers.se/mailman/private/agda/attachments/20080504/244a0ad8/example.obj
-------------- next part --------------


Cheers,
Thorsten



More information about the Agda mailing list