]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/syntax/tweq.ma
milestone in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / syntax / tweq.ma
index fd31eb05e380088b9caa71efe8ca0692e7d04ad0..91234b9cef188722b5b6efb2ac71fac40855320f 100644 (file)
 include "static_2/notation/relations/approxeq_2.ma".
 include "static_2/syntax/term_weight.ma".
 
-lemma tw_le_pair_dx (I): ∀V,T. ♯{T} < ♯{②{I}V.T}.
-#I #V #T /2 width=1 by le_S_S/
-qed.
-
 (* SORT-IRRELEVANT WHD EQUIVALENCE ON TERMS *********************************)
 
 inductive tweq: relation term ≝
@@ -211,7 +207,7 @@ elim (tweq_inv_cast_sn … H) -H #W2 #U2 #HVW #HTU #H destruct
 /2 width=1 by conj/
 qed-.
 
-lemma tweq_inv_cast_sn_XY_Y: ∀T,V. ⓝV.T ≅ T → ⊥.
+lemma tweq_inv_cast_xy_y: ∀T,V. ⓝV.T ≅ T → ⊥.
 @(f_ind … tw) #n #IH #T #Hn #V #H destruct
 elim (tweq_inv_cast_sn … H) -H #X1 #X2 #_ #HX2 #H destruct -V
 /2 width=4 by/