]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_lsuba.ma
- probe: critical bug fixed (all objects were deleted due to wrong test)
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lsubc_lsuba.ma
index b77900e2f9e95c20d292c336b84e967f651a1e0c..255de428b7d0a7609e7c05f84fdc2a6e193fedec 100644 (file)
@@ -18,10 +18,9 @@ include "basic_2/computation/acp_aaa.ma".
 (* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****)
 
 (* properties concerning lenv refinement for atomic arity assignment ********)
-(*
-lamma lsubc_lsuba: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
+
+lemma lsuba_lsubc: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
                    ∀L1,L2. L1 ⁝⊑ L2 → L1 ⊑[RP] L2.
-#RR #RS #RP #H1RP #H2RP #L1 #L2 #H elim H -L1 -L2
-// /2 width=1/ /3 width=4/
+#RR #RS #RP #H1RP #H2RP #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/
 qed.
-*)
\ No newline at end of file