X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Fcprs_lift.ma;h=e570307acf0f7f875af1dd6044bb2e7e163a92ff;hb=d7ccf1bd91637d3c59a285df6f215ecfde2a2450;hp=36ce0ef83b4a9d9d9f6d80d78dfa0324089d693f;hpb=e8998d29ab83e7b6aa495a079193705b2f6743d3;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma index 36ce0ef83..e570307ac 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma @@ -12,67 +12,48 @@ (* *) (**************************************************************************) -include "basic_2/reducibility/cpr_lift.ma". +include "basic_2/reduction/cpr_lift.ma". include "basic_2/computation/cprs.ma". (* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) +(* Advanced properties ******************************************************) + +(* Note: apparently this was missing in basic_1 *) +lemma cprs_delta: ∀G,L,K,V,V2,i. + ⇩[0, i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V ➡* V2 → + ∀W2. ⇧[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡* W2. +#G #L #K #V #V2 #i #HLK #H elim H -V2 [ /3 width=6/ ] +#V1 #V2 #_ #HV12 #IHV1 #W2 #HVW2 +lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK +elim (lift_total V1 0 (i+1)) /4 width=11 by cpr_lift, cprs_strap1/ +qed. + (* Advanced inversion lemmas ************************************************) (* Basic_1: was: pr3_gen_lref *) -lemma cprs_inv_lref1: ∀L,T2,i. L ⊢ #i ➡* T2 → +lemma cprs_inv_lref1: ∀G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡* T2 → T2 = #i ∨ - ∃∃K,V1,T1. ⇩[0, i] L ≡ K. ⓓV1 & - K ⊢ V1 ➡* T1 & - ⇧[0, i + 1] T1 ≡ T2 & - i < |L|. -#L #T2 #i #H @(cprs_ind … H) -T2 /2 width=1/ + ∃∃K,V1,T1. ⇩[0, i] L ≡ K. ⓓV1 & ⦃G, K⦄ ⊢ V1 ➡* T1 & + ⇧[0, i + 1] T1 ≡ T2. +#G #L #T2 #i #H @(cprs_ind … H) -T2 /2 width=1/ #T #T2 #_ #HT2 * [ #H destruct elim (cpr_inv_lref1 … HT2) -HT2 /2 width=1/ - * #K #V1 #T1 #HLK #HVT1 #HT12 #Hi - @or_intror @(ex4_3_intro … HLK … HT12) // /3 width=3/ (**) (* explicit constructors *) -| * #K #V1 #T1 #HLK #HVT1 #HT1 #Hi + * /4 width=6/ +| * #K #V1 #T1 #HLK #HVT1 #HT1 lapply (ldrop_fwd_ldrop2 … HLK) #H0LK - elim (cpr_inv_lift1 … H0LK … HT1 … HT2) -H0LK -T /4 width=6/ + elim (cpr_inv_lift1 … HT2 … H0LK … HT1) -H0LK -T /4 width=6/ ] qed-. -(* Basic_1: was: pr3_gen_abst *) -lemma cprs_inv_abst1: ∀I,W,a,L,V1,T1,U2. L ⊢ ⓛ{a}V1. T1 ➡* U2 → - ∃∃V2,T2. L ⊢ V1 ➡* V2 & L. ⓑ{I} W ⊢ T1 ➡* T2 & - U2 = ⓛ{a}V2. T2. -#I #W #a #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5/ -#U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct -elim (cpr_inv_abst1 … HU2 I W) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /3 width=5/ -qed-. - -lemma cprs_inv_abst: ∀a,L,V1,V2,T1,T2. L ⊢ ⓛ{a}V1. T1 ➡* ⓛ{a}V2. T2 → ∀I,W. - L ⊢ V1 ➡* V2 ∧ L. ⓑ{I} W ⊢ T1 ➡* T2. -#a #L #V1 #V2 #T1 #T2 #H #I #W -elim (cprs_inv_abst1 I W … H) -H #V #T #HV1 #HT1 #H destruct /2 width=1/ -qed-. - (* Relocation properties ****************************************************) (* Basic_1: was: pr3_lift *) -lemma cprs_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → ∀T1,U1. ⇧[d, e] T1 ≡ U1 → - ∀T2. K ⊢ T1 ➡* T2 → ∀U2. ⇧[d, e] T2 ≡ U2 → - L ⊢ U1 ➡* U2. -#L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #HT12 @(cprs_ind … HT12) -T2 -[ -HLK #T2 #HT12 - <(lift_mono … HTU1 … HT12) -T1 // -| -HTU1 #T #T2 #_ #HT2 #IHT2 #U2 #HTU2 - elim (lift_total T d e) #U #HTU - lapply (cpr_lift … HLK … HTU … HTU2 … HT2) -T2 -HLK /3 width=3/ -] -qed. +lemma cprs_lift: ∀G. l_liftable (cprs G). +/3 width=9/ qed. (* Basic_1: was: pr3_gen_lift *) -lemma cprs_inv_lift1: ∀L,K,d,e. ⇩[d, e] L ≡ K → - ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀U2. L ⊢ U1 ➡* U2 → - ∃∃T2. ⇧[d, e] T2 ≡ U2 & K ⊢ T1 ➡* T2. -#L #K #d #e #HLK #T1 #U1 #HTU1 #U2 #HU12 @(cprs_ind … HU12) -U2 /2 width=3/ --HTU1 #U #U2 #_ #HU2 * #T #HTU #HT1 -elim (cpr_inv_lift1 … HLK … HTU … HU2) -U -HLK /3 width=5/ -qed. +lemma cprs_inv_lift1: ∀G. l_deliftable_sn (cprs G). +/3 width=5 by l_deliftable_sn_LTC, cpr_inv_lift1/ +qed-.