(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* Type) (x:Q p) (h:p = p), x = eq_rect p Q x p h. Proof. intros; rewrite M.proof_irrelevance with (p1:=h) (p2:=eq_refl p). reflexivity. Qed. End Eq_rect_eq. (** Export the theory of injective dependent elimination *) Module EqdepTheory := EqdepTheory(Eq_rect_eq). Export EqdepTheory. Scheme eq_indd := Induction for eq Sort Prop. (** We derive the irrelevance of the membership property for subsets *) Lemma subset_eq_compat : forall (U:Type) (P:U->Prop) (x y:U) (p:P x) (q:P y), x = y -> exist P x p = exist P y q. Proof. intros. rewrite M.proof_irrelevance with (p1:=q) (p2:=eq_rect x P p y H). elim H using eq_indd. reflexivity. Qed. Lemma subsetT_eq_compat : forall (U:Type) (P:U->Prop) (x y:U) (p:P x) (q:P y), x = y -> existT P x p = existT P y q. Proof. intros. rewrite M.proof_irrelevance with (p1:=q) (p2:=eq_rect x P p y H). elim H using eq_indd. reflexivity. Qed. End ProofIrrelevanceTheory.