]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/computation/acp_aaa.ma
- first properties of strongly normalizing terms
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / computation / acp_aaa.ma
index 0f70ef703e4de14a9369ca32f7bacd1e97d44556..1cab5d4b8f2b2a04c143d648dc3e190375732165 100644 (file)
@@ -22,6 +22,7 @@ include "Basic_2/computation/lsubc_ldrops.ma".
 
 (* Main propertis ***********************************************************)
 
+(* Basic_1: was: sc3_arity_csubc *)
 theorem aacr_aaa_csubc_lifts: ∀RR,RS,RP.
                               acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
                               ∀L1,T,A. L1 ⊢ T ÷ A → ∀L0,des. ⇩*[des] L0 ≡ L1 →
@@ -44,7 +45,7 @@ theorem aacr_aaa_csubc_lifts: ∀RR,RS,RP.
   [ #K2 #HK20 #H destruct
     generalize in match HLK2; generalize in match I; -HLK2 -I * #HLK2
     [ elim (lift_total V0 0 (i0 +1)) #V #HV0
-      elim (lifts_lift_trans … HV10 … HV0 … Hi0 Hdes0) -HV10 #V2 #HV12 #HV2
+      elim (lifts_lift_trans  … Hi0 … Hdes0 … HV10 … HV0) -HV10 #V2 #HV12 #HV2
       @(s4 … HB … ◊ … HV0 HLK2) /3 width=7/ (* uses IHB HL20 V2 HV0 *)
     | @(s2 … HB … ◊) // /2 width=3/
     ]
@@ -87,6 +88,7 @@ theorem aacr_aaa_csubc_lifts: ∀RR,RS,RP.
 ]
 qed.
 
+(* Basic_1: was: sc3_arity *)
 lemma aacr_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
                 ∀L,T,A. L ⊢ T ÷ A → ⦃L, T⦄ [RP] ϵ 〚A〛.
 /2 width=8/ qed.