[Agda] Congruence and substitutivity rules in heterogeneous equality

Nils Anders Danielsson nad at Cs.Nott.AC.UK
Mon Jul 13 16:06:30 CEST 2009


On 2009-07-13 13:22, Liang-Ting Chen wrote:
> subst (P : A → Set) {x : A} {y : A} → x ≅ y → P x → P y

You could do something like the following:

  ≅⇒≡₁ : ∀ {A B} {x : A} {y : B} → x ≅ y → A ≡₁ B
  ≅⇒≡₁ refl = refl

  coerce : ∀ {A B} → A ≡₁ B → A → B
  coerce refl = id

  subst : ∀ {A B} (P : A → Set) {x : A} {y : B}
          (y≅x : y ≅ x) → P x → P (coerce (≅⇒≡₁ y≅x) y)
  subst P refl p = p

However, I suspect that you do not really want to prove
P (coerce (≅⇒≅₁ y≅x) y), but rather something slightly different.

> Furthermore, it also breaks the heterogeneous equational reasoning for
> the same reason.

In what way is this broken? The standard library requires homogeneity in
many places, but not for equational reasoning for _≅_ (in the latest
version).

-- 
/NAD


This message has been checked for viruses but the contents of an attachment
may still contain software viruses, which could damage your computer system:
you are advised to perform your own checks. Email communications with the
University of Nottingham may be monitored as permitted by UK legislation.



More information about the Agda mailing list