lemma discr_tpair_xy_x: ∀I,T,V. 𝕔{I} V. T = V → False.
#I #T #V elim V -V
[ #J #H destruct
-| #J #W #U #IHW #_ #H destruct -I /2/ (**) (* improve context after destruct *)
+| #J #W #U #IHW #_ #H destruct
+(*
+ (generalize in match e1) -e1 >e0 normalize
+*) -I /2/ (**) (* destruct: one quality is not simplified, the destucted equality is not erased *)
]
qed.
lemma discr_tpair_xy_y: ∀I,V,T. 𝕔{I} V. T = T → False.
#I #V #T elim T -T
[ #J #H destruct
-| #J #W #U #_ #IHU #H destruct -I V /2/ (**) (* improve context after destruct *)
+| #J #W #U #_ #IHU #H destruct -I V /2/ (**) (* destruct: the destucted equality is not erased *)
]
qed.
axiom tnf_inv_abbr: ∀V,T. ℕ[𝕔{Abbr}V.T] → False.
-axiom tnf_inv_cast: ∀V,T. ℕ[𝕔{Cast}V.T] → False.
+lemma tnf_inv_cast: ∀V,T. ℕ[𝕔{Cast}V.T] → False.
+#V #T #H lapply (H T ?) -H /2/
+qed.
(* Basic properties *********************************************************)
∀A.∀a.∀P: ∀x:A. eq ? x a → Type[0]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p.
#A #a #P #H #x #p (generalize in match H) (generalize in match P)
cases p; //; qed.
-
+
+lemma eq_rect_Type1_r:
+ ∀A.∀a.∀P: ∀x:A. eq ? x a → Type[1]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p.
+ #A #a #P #H #x #p (generalize in match H) (generalize in match P)
+ cases p; //; qed.
+
lemma eq_rect_Type2_r:
∀A.∀a.∀P: ∀x:A. eq ? x a → Type[2]. P a (refl A a) → ∀x.∀p:eq ? x a.P x p.
#A #a #P #H #x #p (generalize in match H) (generalize in match P)