<div dir="ltr"><div>Hi Andrew,</div><div><br></div><div>This seems to be an unfortunate interaction between --rewriting and an optimization performed by Agda, the detection of "projection-like" functions. Agda detects that the function `Cset` is projection-like, which allows it to erase the first argument. Unfortunately this means the argument is no longer available for the rewrite rule either.</div><div><br></div><div>If you want, you can turn off the optimization using the `--no-projection-like` flag. I'm not sure if there's currently an option to turn it off on a case-by-case basis.</div><div><br></div><div>-- Jesper<br></div></div><br><div class="gmail_quote"><div dir="ltr" class="gmail_attr">On Sun, Oct 4, 2020 at 3:07 PM Andrew M. Pitts <<a href="mailto:amp12@cam.ac.uk">amp12@cam.ac.uk</a>> wrote:<br></div><blockquote class="gmail_quote" style="margin:0px 0px 0px 0.8ex;border-left:1px solid rgb(204,204,204);padding-left:1ex">I’m struggling to understand an aspect of --rewriting in Agda (2.6.1) and wonder if someone can help?<br>
<br>
The following artificial example illustrates my problem (which is more complicated and less artificial).<br>
<br>
---<br>
{-# OPTIONS --rewriting #-}<br>
<br>
open import Agda.Builtin.Equality<br>
open import Agda.Builtin.Equality.Rewrite<br>
<br>
data C (S : Set) : Set₁ where<br>
  mkC :<br>
    (A : Set)<br>
    (e : A ≡ S)<br>
    → ---------<br>
    C S<br>
<br>
Cset : (S : Set) → C S → Set<br>
Cset S (mkC A e) = A<br>
<br>
Cequ : (S : Set)(c : C S) → Cset S c ≡ S<br>
Cequ S (mkC A e) = e<br>
<br>
{-# REWRITE Cequ #-} — this is not accepted!<br>
<br>
——<br>
<br>
Given S : Set, the datatype C S contains values given by a set A and a proof that A equals S. Can we make that proof definitional rather than propositional, using a REWRITE? Apparently not — if you check the above Agda code, Agda 2.6.1 complains<br>
<br>
-- Cequ is not a legal rewrite rule, since the following variables are<br>
-- not bound by the left hand side: S<br>
-- when checking the pragma REWRITE Cequ<br>
<br>
The “left-hand side” it refers to is Cset S c, which certainly contains an occurrence of S. Why is S not “bound by” this left-hand side? I declared the type of Cequ to be<br>
<br>
(S : Set)(c : C S) → Cset S c ≡ S<br>
<br>
bu Agda 2.6.1 reports it to be<br>
<br>
(S : Set) (c : C S) → Cset _ c ≡ S<br>
<br>
Note the underscore. Presumably this is why REWRITE Cequ is not accepted.<br>
<br>
Where did that underscore come from? And is there some way to avoid it and have the REWRITE accepted?<br>
<br>
Andy<br>
<br>
<br>
<br>
<br>
<br>
<br>
_______________________________________________<br>
Agda mailing list<br>
<a href="mailto:Agda@lists.chalmers.se" target="_blank">Agda@lists.chalmers.se</a><br>
<a href="https://lists.chalmers.se/mailman/listinfo/agda" rel="noreferrer" target="_blank">https://lists.chalmers.se/mailman/listinfo/agda</a><br>
</blockquote></div>