]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma
advances on cpxs and cnx (cnxa removed) ,,,
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / syntax / tdeq.ma
index 4b077059d51a06a5f7ce9b56f38616dff747b2db..6e08f744662721eedc90405c94d9abec182ddc24 100644 (file)
@@ -171,11 +171,11 @@ qed-.
 
 (* Negated inversion lemmas *************************************************)
 
-lemma tdeq_false_inv_pair: ∀h,o,I1,I2,V1,V2,T1,T2.
-                           (②{I1}V1.T1 ≡[h, o] ②{I2}V2.T2 → ⊥) → 
-                           ∨∨ I1 = I2 → ⊥
-                           |  V1 ≡[h, o] V2 → ⊥
-                           |  (T1 ≡[h, o] T2 → ⊥).
+lemma tdneq_inv_pair: ∀h,o,I1,I2,V1,V2,T1,T2.
+                      (②{I1}V1.T1 ≡[h, o] ②{I2}V2.T2 → ⊥) → 
+                      ∨∨ I1 = I2 → ⊥
+                      |  V1 ≡[h, o] V2 → ⊥
+                      |  (T1 ≡[h, o] T2 → ⊥).
 #h #o #I1 #I2 #V1 #V2 #T1 #T2 #H12
 elim (eq_item2_dec I1 I2) /3 width=1 by or3_intro0/ #H destruct
 elim (tdeq_dec h o V1 V2) /3 width=1 by or3_intro1/