X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Ffpbs_lfpxs.ma;h=87818e45b58c0fd54f9db5bdc539c0e4653b9f74;hb=6b35f96790b871aa06b22045b4e8e8dd7bba6590;hp=3c0e8e372658c02211a91f7af6b883d2e2e795b8;hpb=cead5946e630e5357fb26141252e10868c96a14d;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_lfpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_lfpxs.ma index 3c0e8e372..87818e45b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_lfpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_lfpxs.ma @@ -22,7 +22,7 @@ include "basic_2/rt_computation/fpbs_cpxs.ma". (* PARALLEL RST-COMPUTATION FOR CLOSURES ************************************) -(* Properties with uncounted parallel rt-computation on referred entries ****) +(* Properties with unbound parallel rt-computation on referred entries ******) (* Basic_2A1: uses: lpxs_fpbs *) lemma lfpxs_fpbs: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈*[h, T] L2 → ⦃G, L1, T⦄ ≥[h, o] ⦃G, L2, T⦄. @@ -49,6 +49,10 @@ lemma lfpxs_ffdeq_fpbs: ∀h,o,G1,L1,L,T1. ⦃G1, L1⦄ ⊢ ⬈*[h, T1] L → ∀G2,L2,T2. ⦃G1, L, T1⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. /3 width=3 by lfpxs_fpbs_trans, ffdeq_fpbs/ qed. +lemma fpbs_lfpx_trans: ∀h,o,G1,G2,L1,L,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L, T2⦄ → + ∀L2. ⦃G2, L⦄ ⊢ ⬈[h, T2] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. +/3 width=3 by fpbs_lfpxs_trans, lfpx_lfpxs/ qed-. + (* Properties with star-iterated structural successor for closures **********) (* Basic_2A1: uses: fqus_lpxs_fpbs *) @@ -56,7 +60,7 @@ lemma fqus_lfpxs_fpbs: ∀h,o,G1,G2,L1,L,T1,T2. ⦃G1, L1, T1⦄ ⊐* ⦃G2, L, ∀L2. ⦃G2, L⦄ ⊢ ⬈*[h, T2] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. /3 width=3 by fpbs_lfpxs_trans, fqus_fpbs/ qed. -(* Properties with uncounted context-sensitive parallel rt-computation ******) +(* Properties with unbound context-sensitive parallel rt-computation ********) (* Basic_2A1: uses: cpxs_fqus_lpxs_fpbs *) lemma cpxs_fqus_lfpxs_fpbs: ∀h,o,G1,L1,T1,T. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T → @@ -64,21 +68,29 @@ lemma cpxs_fqus_lfpxs_fpbs: ∀h,o,G1,L1,T1,T. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T → ∀L2.⦃G2, L⦄ ⊢ ⬈*[h, T2] L2 → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄. /3 width=5 by cpxs_fqus_fpbs, fpbs_lfpxs_trans/ qed. +lemma fpbs_cpxs_tdeq_fqup_lfpx_trans: ∀h,o,G1,G3,L1,L3,T1,T3. ⦃G1, L1, T1⦄ ≥ [h, o] ⦃G3, L3, T3⦄ → + ∀T4. ⦃G3, L3⦄ ⊢ T3 ⬈*[h] T4 → ∀T5. T4 ≛[h, o] T5 → + ∀G2,L4,T2. ⦃G3, L3, T5⦄ ⊐+ ⦃G2, L4, T2⦄ → + ∀L2. ⦃G2, L4⦄ ⊢ ⬈[h, T2] L2 → ⦃G1, L1, T1⦄ ≥ [h, o] ⦃G2, L2, T2⦄. +#h #o #G1 #G3 #L1 #L3 #T1 #T3 #H13 #T4 #HT34 #T5 #HT45 #G2 #L4 #T2 #H34 #L2 #HL42 +@(fpbs_lfpx_trans … HL42) -L2 (**) (* full auto too slow *) +@(fpbs_fqup_trans … H34) -G2 -L4 -T2 +/3 width=3 by fpbs_cpxs_trans, fpbs_tdeq_trans/ +qed-. + (* Advanced properties ******************************************************) -(* Basic_2A1: uses: fpbs_intro_alt *) -lemma fpbs_intro_star: ∀h,o,G1,L1,T1,T. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T → - ∀G,L,T0. ⦃G1, L1, T⦄ ⊐* ⦃G, L, T0⦄ → - ∀L0. ⦃G, L⦄ ⊢ ⬈*[h, T0] L0 → - ∀G2,L2,T2. ⦃G, L0, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ . +lemma fpbs_intro_fstar: ∀h,o,G1,L1,T1,T. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T → + ∀G,L,T0. ⦃G1, L1, T⦄ ⊐* ⦃G, L, T0⦄ → + ∀L0. ⦃G, L⦄ ⊢ ⬈*[h, T0] L0 → + ∀G2,L2,T2. ⦃G, L0, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ . /3 width=5 by cpxs_fqus_lfpxs_fpbs, fpbs_strap1, fpbq_ffdeq/ qed. (* Advanced inversion lemmas *************************************************) -(* Basic_2A1: uses: fpbs_inv_alt *) -lemma fpbs_inv_star: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → - ∃∃G,L,L0,T,T0. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T & ⦃G1, L1, T⦄ ⊐* ⦃G, L, T0⦄ - & ⦃G, L⦄ ⊢ ⬈*[h, T0] L0 & ⦃G, L0, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄. +lemma fpbs_inv_fstar: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → + ∃∃G,L,L0,T,T0. ⦃G1, L1⦄ ⊢ T1 ⬈*[h] T & ⦃G1, L1, T⦄ ⊐* ⦃G, L, T0⦄ + & ⦃G, L⦄ ⊢ ⬈*[h, T0] L0 & ⦃G, L0, T0⦄ ≛[h, o] ⦃G2, L2, T2⦄. #h #o #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind_dx … H) -G1 -L1 -T1 [ /2 width=9 by ex4_5_intro/ | #G1 #G0 #L1 #L0 #T1 #T0 * -G0 -L0 -T0