<div dir="ltr">By the way, I&#39;ve made a feature request for Coq to have something like this, but more powerful: <a href="https://coq.inria.fr/bugs/show_bug.cgi?id=3927">https://coq.inria.fr/bugs/show_bug.cgi?id=3927</a>.  Perhaps Agda could use a similar extension?</div><div class="gmail_extra"><br><div class="gmail_quote">On Sat, Jan 17, 2015 at 12:09 PM, Jason Gross <span dir="ltr">&lt;<a href="mailto:jasongross9@gmail.com" target="_blank">jasongross9@gmail.com</a>&gt;</span> wrote:<br><blockquote class="gmail_quote" style="margin:0 0 0 .8ex;border-left:1px #ccc solid;padding-left:1ex"><div dir="ltr">Is it possible to push this even further?  For example, define an alternate definition of, say, ap, relying on primTrustMe in some fancy way such that when it gets either univalence or refl, it computes?<div><br></div><div>-Jason</div></div><div class="gmail_extra"><br><div class="gmail_quote"><div><div class="h5">On Sat, Jan 17, 2015 at 2:33 AM, Alan Jeffrey <span dir="ltr">&lt;<a href="mailto:ajeffrey@bell-labs.com" target="_blank">ajeffrey@bell-labs.com</a>&gt;</span> wrote:<br></div></div><blockquote class="gmail_quote" style="margin:0 0 0 .8ex;border-left:1px #ccc solid;padding-left:1ex"><div><div class="h5">Hi everyone,<br>
<br>
In the Agda development of Homotopy Type Theory at <a href="https://github.com/HoTT/HoTT-Agda/" target="_blank">https://github.com/HoTT/HoTT-<u></u>Agda/</a> the univalence axiom is given by three postulates (the map from (A ≃ B) to (A ≡ B) and its β and η rules).<br>
<br>
I wonder whether these postulates could be replaced by uses of primTrustMe?<br>
<br>
As a reminder, primTrustMe is a trusted function which inhabits the type (M ≡ N) for any M and N. It is possible to introduce contradictions (e.g. 0 ≡ 1) in the same way as with postulates, so it has to be handled with care. The semantics is as for postulates, but with an extra beta reduction:<br>
<br>
  primTrustMe M M → refl<br>
<br>
In the attached Agda code, primTrustMe is used to define:<br>
<br>
  private<br>
    trustme : ∀ {ℓ} {A B : Set ℓ} (p : A ≃ B) → (∃ q ∙ ((≡-to-≃ q) ≡ p))<br>
    trustme p = ⟨ primTrustMe , primTrustMe ⟩<br>
<br>
from which we get the map from (A ≃ B) to (A ≡ B) and its β rule:<br>
<br>
  ≃-to-≡ : ∀ {ℓ} {A B : Set ℓ} → (A ≃ B) → (A ≡ B)<br>
  ≃-to-≡ p with trustme p<br>
  ≃-to-≡ .(≡-to-≃ refl) | ⟨ refl , refl ⟩ = refl<br>
<br>
  ≃-to-≡-β : ∀ {ℓ} {A B : Set ℓ} (p : A ≃ B) → (≡-to-≃ (≃-to-≡ p) ≡ p)<br>
  ≃-to-≡-β p with trustme p<br>
  ≃-to-≡-β .(≡-to-≃ refl) | ⟨ refl , refl ⟩ = refl<br>
<br>
Interestingly, the η rule and the coherence property for β and η then become trivial:<br>
<br>
  ≃-to-≡-η : ∀ {ℓ} {A B : Set ℓ} (p : A ≡ B) → (≃-to-≡ (≡-to-≃ p) ≡ p)<br>
  ≃-to-≡-η refl = refl<br>
<br>
  ≃-to-≡-τ : ∀ {ℓ} {A B : Set ℓ} (p : A ≡ B) →<br>
    (cong ≡-to-≃ (≃-to-≡-η p) ≡ ≃-to-≡-β (≡-to-≃ p))<br>
  ≃-to-≡-τ refl = refl<br>
<br>
Note there&#39;s some hoop-jumping with private declarations to hide trustme from users, because:<br>
<br>
  (fst (trustme p)) → refl  (for any p : (A ≃ A))<br>
<br>
that is, all proofs of (A ≃ A) would be identified if we were allowed unfettered access to trustme. Instead, we only allow (≃-to-≡ p) to reduce to refl when (trustme p) reduces to ⟨ refl , refl ⟩, that is not only do we have (A ≃ A) but also that p must be the trivial proof that (A ≃ A).<br>
<br>
Now, this isn&#39;t a conservative extension of HOTT because it introduces extra beta reductions that were previously just propositional equalities, in particular:<br>
<br>
  (≃-to-≡ (≡-to-≃ refl)) → refl<br>
  (≃-to-≡-β (≡-to-≃ refl)) → refl<br>
  (≃-to-≡-η refl) → refl<br>
  (≃-to-≡-τ refl) → refl<br>
<br>
So questions... a) Is this re-inventing the wheel? b) Is this sound? c) Is there a connection between this and a computational interpretation of univalence?<span><font color="#888888"><br>
<br>
Alan.<br>
</font></span><br></div></div>_______________________________________________<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" target="_blank">https://lists.chalmers.se/mailman/listinfo/agda</a><br>
<br></blockquote></div><br></div>
</blockquote></div><br></div>