X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Ffpns.ma;h=beb022541e83e55ffc3570c1712a212089cdf866;hb=e76eade57c0454a58b0d58e5484efe9af417847e;hp=29b81d043e8f2abbc2a1a2fec6dd2246b46d39cf;hpb=2ce98dc56948742e1d27ca4a8b96a3501962d968;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/fpns.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/fpns.ma index 29b81d043..beb022541 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/fpns.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fpns.ma @@ -12,45 +12,28 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/predsnstar_8.ma". -include "basic_2/reduction/fpn.ma". +include "basic_2/notation/relations/btpredsnstar_8.ma". +include "basic_2/relocation/lleq.ma". include "basic_2/computation/lpxs.ma". -(* ORDERED "BIG TREE" NORMAL FORMS ******************************************) +(* PARALLEL COMPUTATION FOR "BIG TREE" NORMAL FORMS *************************) -definition fpns: ∀h. sd h → tri_relation genv lenv term ≝ - λh,g,G1,L1,T1,G2,L2,T2. - ∧∧ G1 = G2 & ⦃G1, L1⦄ ⊢ ➡*[h, g] L2 & T1 = T2. +inductive fpns (h) (g) (G) (L1) (T): relation3 genv lenv term ≝ +| fpns_intro: ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → L1 ⋕[0, T] L2 → fpns h g G L1 T G L2 T +. interpretation - "ordered 'big tree' normal forms (closure)" - 'PRedSnStar h g G1 L1 T1 G2 L2 T2 = (fpns h g G1 L1 T1 G2 L2 T2). + "computation for 'big tree' normal forms (closure)" + 'BTPRedSnStar h g G1 L1 T1 G2 L2 T2 = (fpns h g G1 L1 T1 G2 L2 T2). (* Basic_properties *********************************************************) lemma fpns_refl: ∀h,g. tri_reflexive … (fpns h g). -/2 width=1 by and3_intro/ qed. +/2 width=1 by fpns_intro/ qed. -lemma fpn_fpns: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊢ ➡[h, g] ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 * /3 width=1 by lpx_lpxs, and3_intro/ -qed. +(* Basic inversion lemmas ***************************************************) -lemma fpns_strap1: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G, L, T⦄ → - ⦃G, L, T⦄ ⊢ ➡[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄. -#h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 * #H1G #H1L #G1T * -/3 width=3 by lpxs_strap1, and3_intro/ -qed-. - -lemma fpns_strap2: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ ⊢ ➡[h, g] ⦃G, L, T⦄ → - ⦃G, L, T⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄. -#h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 * #H1G #H1L #G1T * -/3 width=3 by lpxs_strap2, and3_intro/ -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma fpns_fwd_bteq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊢ ➡*[h, g] ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ⋕ ⦃G2, L2, T2⦄. -#h #g #G1 #G2 #L1 #L2 #T1 #T2 * /3 width=4 by lpxs_fwd_length, and3_intro/ +lemma fpns_inv_gen: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊢ ⋕➡*[h, g] ⦃G2, L2, T2⦄ → + ∧∧ G1 = G2 & ⦃G1, L1⦄ ⊢ ➡*[h, g] L2 & L1 ⋕[0, T1] L2 & T1 = T2. +#h #g #G1 #G2 #L1 #L2 #T1 #T2 * -G2 -L2 -T2 /2 width=1 by and4_intro/ qed-.