X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcsx_fqus.ma;h=acef276fa2851d4a9691671a0e068afa62c45fae;hp=03ae7cc3797aa2fc3bbd0c1b67014159bbb28141;hb=a454837a256907d2f83d42ced7be847e10361ea9;hpb=b4283c079ed7069016b8d924bbc7e08872440829 diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fqus.ma index 03ae7cc37..acef276fa 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fqus.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_fqus.ma @@ -19,7 +19,7 @@ include "basic_2/rt_computation/csx_lsubr.ma". (* Properties with extended supclosure **************************************) -lemma csx_fqu_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⊐[b] ⦃G2,L2,T2⦄ → +lemma csx_fqu_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⬂[b] ⦃G2,L2,T2⦄ → ⦃G1,L1⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄ → ⦃G2,L2⦄ ⊢ ⬈*[h] 𝐒⦃T2⦄. #h #b #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 [ /3 width=5 by csx_inv_lref_pair, drops_refl/ @@ -31,19 +31,19 @@ lemma csx_fqu_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⊐[b] ⦃G2,L2,T2 ] qed-. -lemma csx_fquq_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⊐⸮[b] ⦃G2,L2,T2⦄ → +lemma csx_fquq_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⬂⸮[b] ⦃G2,L2,T2⦄ → ⦃G1,L1⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄ → ⦃G2,L2⦄ ⊢ ⬈*[h] 𝐒⦃T2⦄. #h #b #G1 #G2 #L1 #L2 #T1 #T2 * /2 width=6 by csx_fqu_conf/ * #HG #HL #HT destruct // qed-. -lemma csx_fqup_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⊐+[b] ⦃G2,L2,T2⦄ → +lemma csx_fqup_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⬂+[b] ⦃G2,L2,T2⦄ → ⦃G1,L1⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄ → ⦃G2,L2⦄ ⊢ ⬈*[h] 𝐒⦃T2⦄. #h #b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqup_ind … H) -G2 -L2 -T2 /3 width=6 by csx_fqu_conf/ qed-. -lemma csx_fqus_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⊐*[b] ⦃G2,L2,T2⦄ → +lemma csx_fqus_conf: ∀h,b,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ⬂*[b] ⦃G2,L2,T2⦄ → ⦃G1,L1⦄ ⊢ ⬈*[h] 𝐒⦃T1⦄ → ⦃G2,L2⦄ ⊢ ⬈*[h] 𝐒⦃T2⦄. #h #b #G1 #G2 #L1 #L2 #T1 #T2 #H @(fqus_ind … H) -H /3 width=6 by csx_fquq_conf/