X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Funfold%2Ftpss.ma;h=93916208b649a6d301c53ab4ef1e56f9d6417871;hb=9c09a0b1f8801e40612eef429b82fc6dbae01b85;hp=7631985bb1c6ea1dac81f2ff335e9fbd7eeb66f7;hpb=708aa01d44c67343f0dac0353b52c7c2457069b3;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma index 7631985bb..93916208b 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma @@ -40,28 +40,32 @@ qed-. (* Basic properties *********************************************************) -lemma tpss_strap: ∀L,T1,T,T2,d,e. - L ⊢ T1 ▶ [d, e] T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. +lemma tpss_strap1: ∀L,T1,T,T2,d,e. + L ⊢ T1 ▶* [d, e] T → L ⊢ T ▶ [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. /2 width=3/ qed. -lemma tpss_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶* [d, e] T2 → - ∀L2. L1 ≼ [d, e] L2 → L2 ⊢ T1 ▶* [d, e] T2. +lemma tpss_strap2: ∀L,T1,T,T2,d,e. + L ⊢ T1 ▶ [d, e] T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. +/2 width=3/ qed. + +lemma tpss_lsubs_trans: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶* [d, e] T2 → + ∀L2. L2 ≼ [d, e] L1 → L2 ⊢ T1 ▶* [d, e] T2. /3 width=3/ qed. lemma tpss_refl: ∀d,e,L,T. L ⊢ T ▶* [d, e] T. /2 width=1/ qed. lemma tpss_bind: ∀L,V1,V2,d,e. L ⊢ V1 ▶* [d, e] V2 → - ∀I,T1,T2. L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 → - L ⊢ ⓑ{I} V1. T1 ▶* [d, e] ⓑ{I} V2. T2. + ∀a,I,T1,T2. L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 → + L ⊢ ⓑ{a,I} V1. T1 ▶* [d, e] ⓑ{a,I} V2. T2. #L #V1 #V2 #d #e #HV12 elim HV12 -V2 -[ #V2 #HV12 #I #T1 #T2 #HT12 elim HT12 -T2 +[ #V2 #HV12 #a #I #T1 #T2 #HT12 elim HT12 -T2 [ /3 width=5/ | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) ] -| #V #V2 #_ #HV12 #IHV #I #T1 #T2 #HT12 - lapply (tpss_lsubs_conf … HT12 (L. ⓑ{I} V) ?) -HT12 /2 width=1/ #HT12 - lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) +| #V #V2 #_ #HV12 #IHV #a #I #T1 #T2 #HT12 + lapply (tpss_lsubs_trans … HT12 (L. ⓑ{I} V) ?) -HT12 /2 width=1/ #HT12 + lapply (IHV a … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) ] qed. @@ -104,6 +108,11 @@ lapply (tpss_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 lapply (tpss_weak_top … HT12) // qed. +lemma tpss_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶* [d, e] T2 → + ∀L. L @@ K ⊢ T1 ▶* [d, e] T2. +#K #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // /3 width=3/ +qed. + (* Basic inversion lemmas ***************************************************) (* Note: this can be derived from tpss_inv_atom1 *) @@ -124,15 +133,15 @@ lemma tpss_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p ▶* [d, e] T2 → T2 = §p. ] qed-. -lemma tpss_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 ▶* [d, e] U2 → +lemma tpss_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. L ⊢ ⓑ{a,I} V1. T1 ▶* [d, e] U2 → ∃∃V2,T2. L ⊢ V1 ▶* [d, e] V2 & L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 & - U2 = ⓑ{I} V2. T2. -#d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 + U2 = ⓑ{a,I} V2. T2. +#d #e #L #a #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 [ /2 width=5/ | #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct elim (tps_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H - lapply (tpss_lsubs_conf … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ + lapply (tpss_lsubs_trans … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ ] qed-. @@ -155,9 +164,19 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma tpss_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → #[T1] ≤ #[T2]. +lemma tpss_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → #{T1} ≤ #{T2}. #L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // #T #T2 #_ #HT2 #IHT1 lapply (tps_fwd_tw … HT2) -HT2 #HT2 @(transitive_le … IHT1) // qed-. + +lemma tpss_fwd_shift1: ∀L,L1,T1,T,d,e. L ⊢ L1 @@ T1 ▶*[d, e] T → + ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. +#L #L1 #T1 #T #d #e #H @(tpss_ind … H) -T +[ /2 width=4/ +| #T #X #_ #H0 * #L0 #T0 #HL10 #H destruct + elim (tps_fwd_shift1 … H0) -H0 #L2 #T2 #HL02 #H destruct /2 width=4/ +] +qed-. + \ No newline at end of file