[Agda] why is this corecursion red? And indexed coinductive definitions.

Thorsten Altenkirch Thorsten.Altenkirch at nottingham.ac.uk
Sat May 2 14:55:50 CEST 2020


Thank you Guillaume! After updating to the latest version the definition goes through.

Thorsten

On 01/05/2020, 13:11, "Guillaume Allais" <guillaume.allais at ens-lyon.org> wrote:

    Hi Thorsten,

    What you are missing is that the type of the second field depends
    on the value of the first one which is not the case in your shrunk
    example. Andreas recently fixed a bug involving this type of scenario:
    https://github.com/agda/agda/pull/4424

    So you may want to test your example on the development version of Agda.

    guillaume

    Ps: note that your example file is not self-contained: the module
    "delta-i" is missing.

    On 01/05/2020 12:35, Thorsten Altenkirch wrote:
    > Hi everybody,
    > 
    > There are two issues here: the first is that I don’t understand why the termination checker accepts one definition and rejects another. The 2nd is the observation that the particular coinductive definition I am interested in doesn’t seem to be reducible to more primitive coinductive definitions.
    > 
    > I am trying to implement Reedy fibrations using coinduction and I have a problem with corecursion. Ok, I am not expecting that everybody already knows what Reedy fibrations are so I tried to cook down the problem to a simple case. But the simple case passes the termination checker while the one I am actually interested in doesn’t. ☹
    > 
    > I may be overlooking something obvious.
    > 
    > Here is my trivial example – a coinductive definition of (n : ℕ) → A n :
    > 
    > record Πℕ (A : ℕ → Set) : Set where
    > coinductive
    >   field
    >     hd : A 0
    >     tl : Πℕ (λ n → A (suc n))
    > 
    > and here is a function that translates a function into its coinductive representation using corecursion:
    > 
    > mkΠℕ : {A : ℕ → Set} → ((n : ℕ) → A n) → Πℕ A
    > hd (mkΠℕ {A} f) = f 0
    > tl (mkΠℕ {A} f) = mkΠℕ (λ n → f (suc n))
    > 
    > No problem. But I want to do something similar for directed categories (i.e.all the morphisms apart from identities point in one direction). Instead of composing with suc I use an operation
    > 
    > shift : (C : DCat)(X : Obj C 0 → Set) → DCat
    > 
    > which does something similar to directed categories. I attach a file for the details.
    > 
    > record ReedyFib (C : DCat) : Set₁ where
    >   coinductive
    >   field
    >     hd : Obj C 0 → Set
    >     tl : ReedyFib (shift C hd)
    > 
    > and then I want to create a ReedyFib from a strict presheaf using another operation
    > 
    > shiftPSh : {C : DCat}(F : PSh C) → PSh (shift C (Fzero F))
    > 
    > which shifts presheaves as before composition with suc shifted the function.
    > 
    > psh→rfib : {C : DCat} → PSh C → ReedyFib C
    > hd (psh→rfib {C} F) = Fzero F
    > tl (psh→rfib {C} F) = psh→rfib (shiftPSh F)
    > 
    > To me this looks very similar to the example above hence I don’t understand why the termination checker rejects one and accepts the other. I thought maybe it is the negative occurrence of Obj C 0 but I tried a simple example for this (see file) and there is no problem.
    > 
    > What am I missing?
    > 
    > I realize that all examples change a parameter in the coinductive definition hence it is not clear how to reduce this to basic coinductive types. Actually a better way would be to allow to define dependent records and coinductive types. To handle this we need destructors instead of fields, e.g. it should look like this:
    > 
    > record ReedyFib : DCat → Set₁ where
    >   coinductive
    >   destructors
    >     hd : (C : DCat) → ReedyFib C → Obj C 0 → Set
    >     tl : (C : DCat) → ReedyFib C → ReedyFib (shift C (hd C))
    > 
    > We had previously discussed the case of coinductive vectors which using destructors look like this:
    > 
    >   record Vec : ℕ → Set where
    >     coinductive
    >     destructor
    >       hd : {n : ℕ} → Vec (suc n) → A
    >       tl : {n : ℕ} → Vec (suc n) → Vec n
    > 
    > but in this case we can simulate destructors using equality:
    > 
    >   record Vec (m : ℕ) : Set where
    >     coinductive
    >     field
    >       hd-eq : {n : ℕ} → (m ≡ suc n) → A
    >       tl-eq : {n : ℕ} → (m ≡ suc n) → Vec n
    > 
    > However, this only works because Vec appears recursively “unconstrained” which is not the case in my examples.
    > 
    > Thorsten
    > 
    > 
    > 
    > 
    > 
    > 
    > 
    > 
    > 
    > 
    > 
    > 
    > 
    > This message and any attachment are intended solely for the addressee
    > and may contain confidential information. If you have received this
    > message in error, please contact the sender and delete the email and
    > attachment. 
    > 
    > Any views or opinions expressed by the author of this email do not
    > necessarily reflect the views of the University of Nottingham. Email
    > communications with the University of Nottingham may be monitored 
    > where permitted by law.
    > 
    > 
    > 
    > 
    > 
    > _______________________________________________
    > Agda mailing list
    > Agda at lists.chalmers.se
    > https://lists.chalmers.se/mailman/listinfo/agda
    > 




This message and any attachment are intended solely for the addressee
and may contain confidential information. If you have received this
message in error, please contact the sender and delete the email and
attachment. 

Any views or opinions expressed by the author of this email do not
necessarily reflect the views of the University of Nottingham. Email
communications with the University of Nottingham may be monitored 
where permitted by law.






More information about the Agda mailing list