X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fcprs_lift.ma;h=36ce0ef83b4a9d9d9f6d80d78dfa0324089d693f;hb=eae50cc815292d335df1c488a00b39ef98fa5870;hp=ccff2f1feafeea00e1ad0534bd42a2927f2ad2e1;hpb=eb918fc784eacd2094e3986ba321ef47690d9983;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lift.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lift.ma index ccff2f1fe..36ce0ef83 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lift.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lift.ma @@ -12,11 +12,47 @@ (* *) (**************************************************************************) -include "Basic_2/reducibility/cpr_lift.ma". -include "Basic_2/computation/cprs.ma". +include "basic_2/reducibility/cpr_lift.ma". +include "basic_2/computation/cprs.ma". (* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) +(* Advanced inversion lemmas ************************************************) + +(* Basic_1: was: pr3_gen_lref *) +lemma cprs_inv_lref1: ∀L,T2,i. 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/ +#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 + lapply (ldrop_fwd_ldrop2 … HLK) #H0LK + elim (cpr_inv_lift1 … H0LK … HT1 … HT2) -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 *) @@ -33,10 +69,10 @@ lemma cprs_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → ∀T1,U1. ⇧[d, e] T1 ≡ U1 qed. (* Basic_1: was: pr3_gen_lift *) -lemma cprs_inv_lift: ∀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. +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_lift … HLK … HTU … HU2) -U -HLK /3 width=5/ +elim (cpr_inv_lift1 … HLK … HTU … HU2) -U -HLK /3 width=5/ qed.