X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fsubstitution%2Ftps.ma;h=db116f77e63b526eb2dc408ec50dfb645a0b9327;hb=badd398f31309584efc39254e26b056683157a65;hp=f1cc2d9eedb2726d757008ce095be2eb3d410ffa;hpb=708aa01d44c67343f0dac0353b52c7c2457069b3;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma index f1cc2d9ee..db116f77e 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "basic_2/grammar/cl_weight.ma". include "basic_2/substitution/ldrop.ma". (* PARALLEL SUBSTITUTION ON TERMS *******************************************) @@ -21,9 +20,9 @@ inductive tps: nat → nat → lenv → relation term ≝ | tps_atom : ∀L,I,d,e. tps d e L (⓪{I}) (⓪{I}) | tps_subst: ∀L,K,V,W,i,d,e. d ≤ i → i < d + e → ⇩[0, i] L ≡ K. ⓓV → ⇧[0, i + 1] V ≡ W → tps d e L (#i) W -| tps_bind : ∀L,I,V1,V2,T1,T2,d,e. +| tps_bind : ∀L,a,I,V1,V2,T1,T2,d,e. tps d e L V1 V2 → tps (d + 1) e (L. ⓑ{I} V2) T1 T2 → - tps d e L (ⓑ{I} V1. T1) (ⓑ{I} V2. T2) + tps d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) | tps_flat : ∀L,I,V1,V2,T1,T2,d,e. tps d e L V1 V2 → tps d e L T1 T2 → tps d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) @@ -34,12 +33,12 @@ interpretation "parallel substritution (term)" (* Basic properties *********************************************************) -lemma tps_lsubs_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶ [d, e] T2 → - ∀L2. L1 ≼ [d, e] L2 → L2 ⊢ T1 ▶ [d, e] T2. +lemma tps_lsubs_trans: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶ [d, e] T2 → + ∀L2. L2 ≼ [d, e] L1 → L2 ⊢ T1 ▶ [d, e] T2. #L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e [ // | #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12 - elim (ldrop_lsubs_ldrop1_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /2 width=4/ + elim (ldrop_lsubs_ldrop2_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /2 width=4/ | /4 width=1/ | /3 width=1/ ] @@ -59,9 +58,9 @@ lemma tps_full: ∀K,V,T1,L,d. ⇩[0, d] L ≡ (K. ⓓV) → destruct elim (lift_total V 0 (i+1)) #W #HVW elim (lift_split … HVW i i ? ? ?) // /3 width=4/ -| * #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK +| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 - [ elim (IHU1 (L. ⓑ{I} W2) (d+1) ?) -IHU1 /2 width=1/ -HLK /3 width=8/ + [ elim (IHU1 (L. ⓑ{I} W2) (d+1) ?) -IHU1 /2 width=1/ -HLK /3 width=9/ | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8/ ] ] @@ -111,11 +110,11 @@ lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ∀i. d ≤ i → generalize in match Hide; -Hide (**) (* rewriting in the premises, rewrites in the goal too *) >(plus_minus_m_m … Hjde) in ⊢ (% → ?); -Hjde /4 width=4/ ] -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide +| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ -Hdi -Hide >arith_c1x #T #HT1 #HT2 - lapply (tps_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ + lapply (tps_lsubs_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // -Hdi -Hide /3 width=5/ @@ -134,11 +133,11 @@ lemma tps_split_down: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → | -Hdi -Hdj >(plus_minus_m_m (d+e) j) in Hide; // -Hjde /3 width=4/ ] -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide +| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ -Hdi -Hide >arith_c1x #T #HT1 #HT2 - lapply (tps_lsubs_conf … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ + lapply (tps_lsubs_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // -Hdi -Hide /3 width=5/ @@ -156,7 +155,7 @@ fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ∀I. T1 = #L #T1 #T2 #d #e * -L -T1 -T2 -d -e [ #L #I #d #e #J #H destruct /2 width=1/ | #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #I #H destruct /3 width=8/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct | #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct ] qed. @@ -195,22 +194,22 @@ elim (tps_inv_atom1 … H) -H // qed-. fact tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 → - ∀I,V1,T1. U1 = ⓑ{I} V1. T1 → + ∀a,I,V1,T1. U1 = ⓑ{a,I} V1. T1 → ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & L. ⓑ{I} V2 ⊢ T1 ▶ [d + 1, e] T2 & - U2 = ⓑ{I} V2. T2. + U2 = ⓑ{a,I} V2. T2. #d #e #L #U1 #U2 * -d -e -L -U1 -U2 -[ #L #k #d #e #I #V1 #T1 #H destruct -| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct -| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ -| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct +[ #L #k #d #e #a #I #V1 #T1 #H destruct +| #L #K #V #W #i #d #e #_ #_ #_ #_ #a #I #V1 #T1 #H destruct +| #L #b #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #a #I #V #T #H destruct /2 width=5/ +| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #a #I #V #T #H destruct ] qed. -lemma tps_inv_bind1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓑ{I} V1. T1 ▶ [d, e] U2 → +lemma tps_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. + U2 = ⓑ{a,I} V2. T2. /2 width=3/ qed-. fact tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 → @@ -220,7 +219,7 @@ fact tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 → #d #e #L #U1 #U2 * -d -e -L -U1 -U2 [ #L #k #d #e #I #V1 #T1 #H destruct | #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct -| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct +| #L #a #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct | #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ ] qed. @@ -246,7 +245,7 @@ lemma tps_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 ▶ [d, 0] T2 → T1 = T2. (* Basic forward lemmas *****************************************************) -lemma tps_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → #[T1] ≤ #[T2]. +lemma tps_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → #{T1} ≤ #{T2}. #L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e normalize /3 by monotonic_le_plus_l, le_plus/ (**) (* just /3 width=1/ is too slow *) qed-.