X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Flfpxs_fqup.ma;h=0f90e90e21377c2cd42fc92c6357069186ab38c0;hb=f7296f9cf2ee73465a374942c46b138f35c42ccb;hp=ce2496b91c419a27a63cf6f5ec228018a23cdbf0;hpb=10b733131aa2667d8ba4318d517f0ba3cf137359;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_fqup.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_fqup.ma index ce2496b91..0f90e90e2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_fqup.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_fqup.ma @@ -12,31 +12,39 @@ (* *) (**************************************************************************) -include "basic_2/rt_transition/lfpx_fqup.ma". +include "basic_2/i_static/tc_lfxs_fqup.ma". include "basic_2/rt_computation/lfpxs.ma". (* UNCOUNTED PARALLEL RT-COMPUTATION FOR LOCAL ENV.S ON REFERRED ENTRIES ****) +(* Advanced properties ******************************************************) + +(* Basic_2A1: uses: lpxs_refl *) +lemma lfpxs_refl: ∀h,G,T. reflexive … (lfpxs h G T). +/2 width=1 by tc_lfxs_refl/ qed. + +(* Advanced forward lemmas **************************************************) + +lemma lfpxs_fwd_bind_dx: ∀h,p,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈*[h, ⓑ{p,I}V.T] L2 → + ⦃G, L1.ⓑ{I}V⦄ ⊢ ⬈*[h, T] L2.ⓑ{I}V. +/2 width=2 by tc_lfxs_fwd_bind_dx/ qed-. + +lemma lfpxs_fwd_bind_dx_void: ∀h,p,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈*[h, ⓑ{p,I}V.T] L2 → + ⦃G, L1.ⓧ⦄ ⊢ ⬈*[h, T] L2.ⓧ. +/2 width=4 by tc_lfxs_fwd_bind_dx_void/ qed-. + (* Advanced eliminators *****************************************************) -(* Basic_2A1: was just: lpxs_ind *) -lemma lfpxs_ind: ∀h,G,L1,T. ∀R:predicate lenv. R L1 → - (∀L,L2. ⦃G, L1⦄ ⊢ ⬈*[h, T] L → ⦃G, L⦄ ⊢ ⬈[h, T] L2 → R L → R L2) → - ∀L2. ⦃G, L1⦄ ⊢ ⬈*[h, T] L2 → R L2. -#h #G #L1 #T #R #HL1 #IHL1 #L2 #HL12 -@(TC_star_ind … HL1 IHL1 … HL12) // -qed-. +(* Basic_2A1: uses: lpxs_ind *) +lemma lfpxs_ind_sn: ∀h,G,L1,T. ∀R:predicate lenv. R L1 → + (∀L,L2. ⦃G, L1⦄ ⊢ ⬈*[h, T] L → ⦃G, L⦄ ⊢ ⬈[h, T] L2 → R L → R L2) → + ∀L2. ⦃G, L1⦄ ⊢ ⬈*[h, T] L2 → R L2. +#h #G @tc_lfxs_ind_sn // (**) (* auto fails *) +qed-. -(* Basic_2A1: was just: lpxs_ind_dx *) +(* Basic_2A1: uses: lpxs_ind_dx *) lemma lfpxs_ind_dx: ∀h,G,L2,T. ∀R:predicate lenv. R L2 → (∀L1,L. ⦃G, L1⦄ ⊢ ⬈[h, T] L → ⦃G, L⦄ ⊢ ⬈*[h, T] L2 → R L → R L1) → ∀L1. ⦃G, L1⦄ ⊢ ⬈*[h, T] L2 → R L1. -#h #o #G #L2 #R #HL2 #IHL2 #L1 #HL12 -@(TC_star_ind_dx … HL2 IHL2 … HL12) // -qed-. - -(* Advanced properties ******************************************************) - -(* Basic_2A1: was just: lpxs_refl *) -lemma lfpxs_refl: ∀h,G,L,T. ⦃G, L⦄ ⊢ ⬈*[h, T] L. -/2 width=1 by lfpx_lfpxs/ qed. +#h #G @tc_lfxs_ind_dx // (**) (* auto fails *) +qed-.