]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_aaa.ma
partial commit: "computation" component ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lpxs_aaa.ma
index e866102d4835ab2fe1f60ab397da6d2ccfd387c6..80c43c731c5b0a4347544ee8ee9ee574a3e51c44 100644 (file)
@@ -19,13 +19,14 @@ include "basic_2/computation/lpxs.ma".
 
 (* Properties about atomic arity assignment on terms ************************)
 
-lemma aaa_lpxs_conf: ∀h,g,L1,T,A.
-                     L1 ⊢ T ⁝ A → ∀L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → L2 ⊢ T ⁝ A.
-#h #g #L1 #T #A #HT #L2 #HL12
+lemma aaa_lpxs_conf: ∀h,g,G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A →
+                     ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
+#h #g #G #L1 #T #A #HT #L2 #HL12
 @(TC_Conf3 … (λL,A. ⦃G, L⦄ ⊢ T ⁝ A) … HT ? HL12) /2 width=5 by aaa_lpx_conf/
 qed-.
 
-lemma aaa_lprs_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. L1 ⊢ ➡* L2 → L2 ⊢ T ⁝ A.
-#L1 #T #A #HT #L2 #HL12
+lemma aaa_lprs_conf: ∀G,L1,T,A. ⦃G, L1⦄ ⊢ T ⁝ A →
+                     ∀L2. ⦃G, L1⦄ ⊢ ➡* L2 → ⦃G, L2⦄ ⊢ T ⁝ A.
+#G #L1 #T #A #HT #L2 #HL12
 @(aaa_lpxs_conf … HT) -A -T /2 width=3/
 qed-.