]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_lsuba.ma
universary milestone in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lsubc_lsuba.ma
index b4f5be28952efa0d91e31ef5a83f45485e05bcd9..3a17f0e5b42c7eb1db6810352e88bc9f7e2dcf18 100644 (file)
 (**************************************************************************)
 
 include "basic_2/static/lsuba.ma".
-include "basic_2/computation/acp_aaa.ma".
+include "basic_2/computation/gcp_aaa.ma".
 
-(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****)
+(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************)
 
 (* properties concerning lenv refinement for atomic arity assignment ********)
 
-lemma lsuba_lsubc: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λG,L,T. RP G L T) →
-                   â\88\80G,L1,L2. G â\8a¢ L1 â\81\9dâ\8a\91 L2 â\86\92 G â\8a¢ L1 â\8a\91[RP] L2.
-#RR #RS #RP #H1RP #H2RP #G #L1 #L2 #H elim H -L1 -L2 // /2 width=1/
-#L1 #L2 #W #V #A #H elim (aaa_inv_cast … H) -H /3 width=4/
+lemma lsuba_lsubc: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
+                   â\88\80G,L1,L2. G â\8a¢ L1 â«\83â\81\9d L2 â\86\92 G â\8a¢ L1 â«\83[RP] L2.
+#RR #RS #RP #H1RP #H2RP #G #L1 #L2 #H elim H -L1 -L2 /2 width=1 by lsubc_pair/
+#L1 #L2 #W #V #A #H elim (aaa_inv_cast … H) -H /3 width=4 by acr_aaa, lsubc_beta/
 qed.