X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Flfpxs.ma;h=507f4d9794f0af0b2bdcd3b7ec4f6e87adffb83a;hb=58ea181757dce19b875b2f5a224fe193b2263004;hp=bfee0b2865d26f5de671c41117ee0dde64de2106;hpb=670ad7822d59e598a38d9037d482d3de188b170c;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs.ma index bfee0b286..507f4d979 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs.ma @@ -13,12 +13,13 @@ (**************************************************************************) include "basic_2/notation/relations/predtysnstar_5.ma". +include "basic_2/i_static/tc_lfxs.ma". include "basic_2/rt_transition/lfpx.ma". (* UNCOUNTED PARALLEL RT-COMPUTATION FOR LOCAL ENV.S ON REFERRED ENTRIES ****) definition lfpxs: ∀h. relation4 genv term lenv lenv ≝ - λh,G,T. TC … (lfpx h G T). + λh,G. tc_lfxs (cpx h G). interpretation "uncounted parallel rt-computation on referred entries (local environment)" @@ -31,28 +32,27 @@ lemma lfpx_lfpxs: ∀h,G,T,L1,L2. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → ⦃G, L1⦄ /2 width=1 by inj/ qed. (* Basic_2A1: was just: lpxs_strap1 *) -lemma lfpxs_strap1: ∀h,G,T,L1,L,L2. ⦃G, L1⦄ ⊢ ⬈*[h, T] L → ⦃G, L⦄ ⊢ ⬈[h, T] L2 → ⦃G, L1⦄ ⊢ ⬈*[h, T] L2. -/2 width=3 by step/ qed. +lemma lfpxs_step_dx: ∀h,G,T,L1,L,L2. ⦃G, L1⦄ ⊢ ⬈*[h, T] L → ⦃G, L⦄ ⊢ ⬈[h, T] L2 → ⦃G, L1⦄ ⊢ ⬈*[h, T] L2. +/2 width=3 by tc_lfxs_step_dx/ qed. (* Basic_2A1: was just: lpxs_strap2 *) -lemma lfpxs_strap2: ∀h,G,T,L1,L,L2. ⦃G, L1⦄ ⊢ ⬈[h, T] L → ⦃G, L⦄ ⊢ ⬈*[h, T] L2 → ⦃G, L1⦄ ⊢ ⬈*[h, T] L2. -/2 width=3 by TC_strap/ qed. -(* -(* Basic_2A1: was just: lpxs_pair_refl *) -lemma lfpxs_pair_refl: ∀h,G,T,L1,L2. ⦃G, L1⦄ ⊢ ⬈*[h, T] L2 → ∀I,V. ⦃G, L1.ⓑ{I}V⦄ ⊢ ⬈*[h, T] L2.ⓑ{I}V. -/2 width=1 by TC_lpx_sn_pair_refl/ qed. +lemma lfpxs_step_sn: ∀h,G,T,L1,L,L2. ⦃G, L1⦄ ⊢ ⬈[h, T] L → ⦃G, L⦄ ⊢ ⬈*[h, T] L2 → ⦃G, L1⦄ ⊢ ⬈*[h, T] L2. +/2 width=3 by tc_lfxs_step_sn/ qed. (* Basic inversion lemmas ***************************************************) -(* Basic_2A1: was just: lpxs_inv_atom1 *) -lemma lfpxs_inv_atom1: ∀h,G,L2.T. ⦃G, ⋆⦄ ⊢ ⬈*[h, T] L2 → L2 = ⋆. -/2 width=2 by TC_lpx_sn_inv_atom1/ qed-. +(* Basic_2A1: uses: lpxs_inv_atom1 *) +lemma lfpxs_inv_atom1: ∀h,I,G,L2. ⦃G, ⋆⦄ ⊢ ⬈*[h, ⓪{I}] L2 → L2 = ⋆. +/2 width=3 by tc_lfxs_inv_atom_sn/ qed-. -(* Basic_2A1: was just: lpxs_inv_atom2 *) -lemma lfpxs_inv_atom2: ∀h,G,L1,T. ⦃G, L1⦄ ⊢ ⬈*[h, T] ⋆ → L1 = ⋆. -/2 width=2 by TC_lpx_sn_inv_atom2/ qed-. -*) +(* Basic_2A1: uses: lpxs_inv_atom2 *) +lemma lfpxs_inv_atom2: ∀h,I,G,L1. ⦃G, L1⦄ ⊢ ⬈*[h, ⓪{I}] ⋆ → L1 = ⋆. +/2 width=3 by tc_lfxs_inv_atom_dx/ qed-. -(* Basic_2A1: removed theorems 1: - lpxs_pair -*) +(* Basic forward lemmas *****************************************************) + +lemma lfpxs_fwd_pair_sn: ∀h,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈*[h, ②{I}V.T] L2 → ⦃G, L1⦄ ⊢ ⬈*[h, V] L2. +/2 width=3 by tc_lfxs_fwd_pair_sn/ qed-. + +lemma lfpxs_flat_dx: ∀h,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈*[h, ⓕ{I}V.T] L2 → ⦃G, L1⦄ ⊢ ⬈*[h, T] L2. +/2 width=3 by tc_lfxs_fwd_flat_dx/ qed-.