<html xmlns:o="urn:schemas-microsoft-com:office:office" xmlns:w="urn:schemas-microsoft-com:office:word" xmlns:m="http://schemas.microsoft.com/office/2004/12/omml" xmlns="http://www.w3.org/TR/REC-html40">
<head>
<meta http-equiv="Content-Type" content="text/html; charset=utf-8">
<meta name="Generator" content="Microsoft Word 15 (filtered medium)">
<style><!--
/* Font Definitions */
@font-face
        {font-family:"Cambria Math";
        panose-1:2 4 5 3 5 4 6 3 2 4;}
@font-face
        {font-family:Calibri;
        panose-1:2 15 5 2 2 2 4 3 2 4;}
@font-face
        {font-family:"Apple Color Emoji";
        panose-1:0 0 0 0 0 0 0 0 0 0;}
/* Style Definitions */
p.MsoNormal, li.MsoNormal, div.MsoNormal
        {margin:0cm;
        margin-bottom:.0001pt;
        font-size:12.0pt;
        font-family:"Calibri",sans-serif;
        mso-fareast-language:EN-US;}
.MsoChpDefault
        {mso-style-type:export-only;
        font-size:10.0pt;}
@page WordSection1
        {size:612.0pt 792.0pt;
        margin:72.0pt 72.0pt 72.0pt 72.0pt;}
div.WordSection1
        {page:WordSection1;}
--></style>
</head>
<body lang="EN-GB" link="#0563C1" vlink="#954F72">
<div class="WordSection1">
<p class="MsoNormal"><span style="font-size:11.0pt">Hi everybody,<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">There are two issues here: the first is that I don’t understand why the termination checker accepts one definition and rejects another. The 2<sup>nd</sup> is the observation that the particular coinductive
 definition I am interested in doesn’t seem to be reducible to more primitive coinductive definitions.<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">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.
</span><span style="font-size:11.0pt;font-family:"Apple Color Emoji"">☹</span><span style="font-size:11.0pt">
<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">I may be overlooking something obvious.<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">Here is my trivial example – a coinductive definition of (n :
</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt">) → A n :<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">record Π</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> (A :
</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> → Set) : Set where<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">coinductive<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">  field<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    hd : A 0<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    tl : Π</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> (λ n → A (suc n))<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">and here is a function that translates a function into its coinductive representation using corecursion:<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">mkΠ</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> : {A :
</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> → Set} → ((n :
</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt">) → A n) → Π</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> A<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">hd (mkΠ</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> {A} f) = f 0<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">tl (mkΠ</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> {A} f) = mkΠ</span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt">
 (λ n → f (suc n))<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">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<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">shift : (C : DCat)(X : Obj C 0 → Set) → DCat<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">which does something similar to directed categories. I attach a file for the details.<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">record ReedyFib (C : DCat) : Set₁ where<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">  coinductive<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">  field<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    hd : Obj C 0 → Set<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    tl : ReedyFib (shift C hd)<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">and then I want to create a ReedyFib from a strict presheaf using another operation
<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">shiftPSh : {C : DCat}(F : PSh C) → PSh (shift C (Fzero F))<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">which shifts presheaves as before composition with suc shifted the function.<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">psh→rfib : {C : DCat} → PSh C → ReedyFib C<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">hd (psh→rfib {C} F) = Fzero F<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">tl (psh→rfib {C} F) = psh→rfib (shiftPSh F)<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">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.<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">What am I missing? <o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">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:<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">record ReedyFib : DCat → Set₁ where<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">  coinductive<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">  destructors<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    hd : (C : DCat) → ReedyFib C → Obj C 0 → Set<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    tl : (C : DCat) → ReedyFib C → ReedyFib (shift C (hd C))<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">We had previously discussed the case of coinductive vectors which using destructors look like this:<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">  record Vec : </span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt"> → Set where<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    coinductive<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    destructor<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">      hd : {n : </span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt">} → Vec (suc n) → A<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">      tl : {n : </span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt">} → Vec (suc n) → Vec n<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">but in this case we can simulate destructors using equality:<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">  record Vec (m : </span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt">) : Set where<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    coinductive<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">    field<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">      hd-eq : {n : </span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt">} → (m ≡ suc n) → A<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">      tl-eq : {n : </span><span style="font-size:11.0pt;font-family:"Cambria Math",serif">ℕ</span><span style="font-size:11.0pt">} → (m ≡ suc n) → Vec n<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">However, this only works because Vec appears recursively “unconstrained” which is not the case in my examples.
<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt">Thorsten<o:p></o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
<p class="MsoNormal"><span style="font-size:11.0pt"><o:p> </o:p></span></p>
</div>
<PRE>

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.



</PRE></body>
</html>