[Agda] why coinductive?

Thorsten Altenkirch Thorsten.Altenkirch at nottingham.ac.uk
Sun Mar 1 22:32:10 CET 2020


Hi Henning,



    I think this is about η-expansion, which is disabled by default for

    recursive records:

    https://agda.readthedocs.io/en/v2.5.4/language/record-types.html#eta-expansion



You are right, I forgot. Btw, one could have eta-expansion for sums but it seems that eta-expansion for records is more popular.



    Having η for recursive coinductive types would make definitional

    equality undecidable: https://cronfa.swan.ac.uk/Record/cronfa38822



    About the destructors, or rather what I would call observations,



I prefer destructors because it is nicely dual to constructors.



this

    asymmetry is there because dependencies in the domain of observations of

    coinductive types usually don't change because people use an explicit

    equality instead. For instance, instead of



    record List (A : Set) : ℕ → Set where

    coinductive

      hd : {n : ℕ} → List A (s n) → A

      tl : {n : ℕ} → List A (s n) → List A n



    people would write



    record List (A : Set) (n : ℕ) : Set where

    coinductive

      hd : {k : ℕ} → (n == s k) → A

      tl : {k : ℕ} → (n == s k) → List A k



    (Note that, if you change ℕ to conatural numbers, then this example

    becomes more interesting).



Indeed. I don’t understand the argument because one could apply the dual trick for inductive types as you say yourself below.



    These two types are isomorphic via an adjunction. Thus, it does not

    matter if we make the domain explicit or not, assuming that this is what

    your question was getting at.



    For inductive types, we prefer the first version, where dependencies in

    the codomain of constructors can change, as this simplifies pattern

    matching. The standard example there would be



    data Vec (A : Set) : ℕ → Set where

    coinductive

      []   : Vec A 0

      _::_ : A → {n : ℕ} → Vec A n → List A (s n)



    vs



    data Vec (A : Set) (n : ℕ) : Set where

    coinductive

      []   : {k : ℕ} → (n == 0) → Vec A n

      _::_ : A → {k : ℕ} → (n == s k) → Vec A k → List A n



My proposal is that we can should be able to write



record List (A : Set) : ℕ → Set where

  coinductive

  destructor

    hd : {n : ℕ} →  List A (suc n) →  A

    tl : {n : ℕ} → List A (suc n) →  List A n



That is “destructor” is an alternative to field which would be still allowed to ensure backwards compatibility.



As far as I can see copatternmatching should work nicely and we should be able to write:



map : {A B : Set} → (A → B) → {n : ℕ} → List A n → List B n

hd (map f x) = f (hd x)

tl (map f x) = map f (tl x)



There should be no need to abuse any equality types.



Thorsten





    On 01/03/2020 14:19, Thorsten Altenkirch wrote:

    > Hi,

    >

    >

    >

    > I am just writing some lecture notes hence I noticed something else. Why

    > do we have to say coinductive when defining a recursive record, aka a

    > coinductive type? This seems asymmetric because we don’t have t say

    > anything when defining a recursive sum (like the natural numbers)? Is

    > this just a relic from the days that records couldn’t be recursive or is

    > there a technical reason?

    >

    >

    >

    > Btw, I really would like to have destructors for coinductive types

   > (another asymmetry).

    >

    >

    >

    > 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.




-------------- next part --------------
An HTML attachment was scrubbed...
URL: <http://lists.chalmers.se/pipermail/agda/attachments/20200301/3602c608/attachment.html>


More information about the Agda mailing list