]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/static/aaa_rdeq.ma
milestone in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / static / aaa_rdeq.ma
index 23760c8bd3d829e0565a748dfdfc88dbcc175456..265ed7f5c38895bd31a8b40bd6419dd29f8f6cef 100644 (file)
@@ -17,11 +17,11 @@ include "static_2/static/aaa.ma".
 
 (* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************)
 
-(* Properties with degree-based equivalence on referred entries *************)
+(* Properties with sort-irrelevant equivalence on referred entries **********)
 
-lemma aaa_tdeq_conf_rdeq: ∀h,o,G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀T2. T1 ≛[h, o] T2 →
-                          ∀L2. L1 ≛[h, o, T1] L2 → ⦃G, L2⦄ ⊢ T2 ⁝ A.
-#h #o #G #L1 #T1 #A #H elim H -G -L1 -T1 -A
+lemma aaa_tdeq_conf_rdeq: ∀G,L1,T1,A. ⦃G, L1⦄ ⊢ T1 ⁝ A → ∀T2. T1 ≛ T2 →
+                          ∀L2. L1 ≛[T1] L2 → ⦃G, L2⦄ ⊢ T2 ⁝ A.
+#G #L1 #T1 #A #H elim H -G -L1 -T1 -A
 [ #G #L1 #s1 #X #H1 elim (tdeq_inv_sort1 … H1) -H1 //
 | #I #G #L1 #V1 #B #_ #IH #X #H1 >(tdeq_inv_lref1 … H1) -H1
   #Y #H2 elim (rdeq_inv_zero_pair_sn … H2) -H2