From: Ferruccio Guidi Date: Wed, 5 Sep 2012 16:09:37 +0000 (+0000) Subject: the partial commit continues ... X-Git-Tag: make_still_working~1509 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=039f4f6db3a3c128959cd471eb78f575906e07b6;p=helm.git the partial commit continues ... --- diff --git a/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_append.ma b/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_append.ma index 83b65a335..2f6932188 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_append.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_append.ma @@ -29,6 +29,10 @@ lemma append_atom_sn: ∀L. ⋆ @@ L = L. #L elim L -L normalize // qed. +lemma append_assoc: associative … append. +#L1 #L2 #L3 elim L3 -L3 normalize // +qed. + lemma append_length: ∀L1,L2. |L1 @@ L2| = |L1| + |L2|. #L1 #L2 elim L2 -L2 normalize // qed. diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf_cif.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf_cif.ma index 33e52e1b7..51ce95aa0 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf_cif.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cnf_cif.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -include "basic_2/substitution/tps_lift.ma". -include "basic_2/unfold/tpss.ma". include "basic_2/reducibility/cif.ma". include "basic_2/reducibility/cnf_lift.ma". diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/crf.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/crf.ma index 3823b4c29..5015f033c 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/reducibility/crf.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/crf.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "basic_2/grammar/term_simple.ma". include "basic_2/substitution/ldrop.ma". (* CONTEXT-SENSITIVE REDUCIBLE TERMS ****************************************) diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_append.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_append.ma index 3f6b2c538..359d39c80 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_append.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/ldrop_append.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "basic_2/grammar/lenv_append.ma". include "basic_2/substitution/ldrop.ma". (* DROPPING *****************************************************************) @@ -31,6 +30,8 @@ lemma ldrop_O1_append_sn_le: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → e ≤ |L1| → ∀L. ⇩[0, e] L @@ L1 ≡ L @@ L2. /2 width=3 by ldrop_O1_append_sn_le_aux/ qed. +(* Inversion lemmas on append for local environments ************************) + lemma ldrop_O1_inv_append1_ge: ∀K,L1,L2,e. ⇩[0, e] L1 @@ L2 ≡ K → |L2| ≤ e → ⇩[0, e - |L2|] L1 ≡ K. #K #L1 #L2 elim L2 -L2 normalize // 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 7e35dca7c..0bde24417 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma @@ -12,8 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/grammar/lenv_top.ma". -include "basic_2/substitution/ldrop.ma". +include "basic_2/substitution/ldrop_append.ma". (* PARALLEL SUBSTITUTION ON TERMS *******************************************) @@ -145,6 +144,15 @@ lemma tps_split_down: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ] qed. +lemma tps_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶ [d, e] T2 → + ∀L. L @@ K ⊢ T1 ▶ [d, e] T2. +#K #T1 #T2 #d #e #H elim H -K -T1 -T2 -d -e // /2 width=1/ +#K #K0 #V #W #i #d #e #Hdi #Hide #HK0 #HVW #L +lapply (ldrop_fwd_ldrop2_length … HK0) #H +@(tps_subst … (L@@K0) … HVW) // (**) (* /3/ does not work *) +@(ldrop_O1_append_sn_le … HK0) /2 width=2/ +qed. + (* Basic inversion lemmas ***************************************************) fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ∀I. T1 = ⓪{I} → @@ -251,17 +259,6 @@ lemma tps_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → #{T1} ≤ #{T2}. /3 by monotonic_le_plus_l, le_plus/ (**) (* just /3 width=1/ is too slow *) qed-. -lemma tps_fwd_shift1: ∀L1,L,T1,T,d,e. L ⊢ L1 @@ T1 ▶ [d, e] T → - ∃∃L2,T2. L1 𝟙 L2 & T = L2 @@ T2. -#L1 @(lenv_ind_dx … L1) -L1 -[ #L #T1 #T #d #e #_ @ex2_2_intro [3: // |4: // |1,2: skip ] (**) (* /2 width=4/ does not work *) -| #I #L1 #V1 #IH #L #T1 #T #d #e >shift_append_assoc #H - elim (tps_inv_bind1 … H) -H #V2 #T2 #_ #HT12 #H destruct - elim (IH … HT12) -IH -L -T1 -d -e #L2 #T #HL12 #H destruct - @(ex2_2_intro … (⋆.ⓑ{I}V2@@L2)) /2 width=4/ /3 width=2/ -] -qed-. - (* Basic_1: removed theorems 25: subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss.ma index 89b19ea45..dfd1a19f6 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss.ma @@ -31,79 +31,6 @@ inductive ltpss: nat → nat → relation lenv ≝ interpretation "parallel unfold (local environment)" 'PSubstStar L1 d e L2 = (ltpss d e L1 L2). -(* Basic properties *********************************************************) - -lemma ltpss_tps2: ∀L1,L2,I,V1,V2,e. - L1 ▶* [0, e] L2 → L2 ⊢ V1 ▶ [0, e] V2 → - L1. ⓑ{I} V1 ▶* [0, e + 1] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_tps1: ∀L1,L2,I,V1,V2,d,e. - L1 ▶* [d, e] L2 → L2 ⊢ V1 ▶ [d, e] V2 → - L1. ⓑ{I} V1 ▶* [d + 1, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_tpss2_lt: ∀L1,L2,I,V1,V2,e. - L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶* [0, e - 1] V2 → - 0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He ->(plus_minus_m_m e 1) /2 width=1/ -qed. - -lemma ltpss_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶* [d - 1, e] V2 → - 0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd ->(plus_minus_m_m d 1) /2 width=1/ -qed. - -lemma ltpss_tps2_lt: ∀L1,L2,I,V1,V2,e. - L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶ [0, e - 1] V2 → - 0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_tps1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶ [d - 1, e] V2 → - 0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -(* Basic_1: was by definition: csubst1_refl *) -lemma ltpss_refl: ∀L,d,e. L ▶* [d, e] L. -#L elim L -L // -#L #I #V #IHL * /2 width=1/ * /2 width=1/ -qed. - -lemma ltpss_weak: ∀L1,L2,d1,e1. L1 ▶* [d1, e1] L2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → L1 ▶* [d2, e2] L2. -#L1 #L2 #d1 #e1 #H elim H -L1 -L2 -d1 -e1 // -[ #L1 #L2 #I #V1 #V2 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd2 #Hde2 - lapply (le_n_O_to_eq … Hd2) #H destruct normalize in Hde2; - lapply (lt_to_le_to_lt 0 … Hde2) // #He2 - lapply (le_plus_to_minus_r … Hde2) -Hde2 /3 width=5/ -| #L1 #L2 #I #V1 #V2 #d1 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd21 #Hde12 - >plus_plus_comm_23 in Hde12; #Hde12 - elim (le_to_or_lt_eq 0 d2 ?) // #H destruct - [ lapply (le_plus_to_minus_r … Hde12) -Hde12 (plus_minus_m_m e 1) /2 width=1/ +qed. + +lemma ltpss_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶* [d - 1, e] V2 → + 0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd +>(plus_minus_m_m d 1) /2 width=1/ +qed. + +lemma ltpss_tps2_lt: ∀L1,L2,I,V1,V2,e. + L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶ [0, e - 1] V2 → + 0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2. +/3 width=1/ qed. + +lemma ltpss_tps1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶ [d - 1, e] V2 → + 0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2. +/3 width=1/ qed. + +(* Basic_1: was by definition: csubst1_refl *) +lemma ltpss_refl: ∀L,d,e. L ▶* [d, e] L. +#L elim L -L // +#L #I #V #IHL * /2 width=1/ * /2 width=1/ +qed. + +lemma ltpss_weak: ∀L1,L2,d1,e1. L1 ▶* [d1, e1] L2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → L1 ▶* [d2, e2] L2. +#L1 #L2 #d1 #e1 #H elim H -L1 -L2 -d1 -e1 // +[ #L1 #L2 #I #V1 #V2 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd2 #Hde2 + lapply (le_n_O_to_eq … Hd2) #H destruct normalize in Hde2; + lapply (lt_to_le_to_lt 0 … Hde2) // #He2 + lapply (le_plus_to_minus_r … Hde2) -Hde2 /3 width=5/ +| #L1 #L2 #I #V1 #V2 #d1 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd21 #Hde12 + >plus_plus_comm_23 in Hde12; #Hde12 + elim (le_to_or_lt_eq 0 d2 ?) // #H destruct + [ lapply (le_plus_to_minus_r … Hde12) -Hde12 plus_plus_comm_23 + /4 width=5 by ltpss_tpss2, tpss_append, tpss_weak, monotonic_le_plus_r/ (**) (* too slow without trace *) +| #K1 #K2 #I #V1 #V2 #d #x #_ #HV12 #IHK12 normalize shift_append_assoc #H + elim (tps_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + elim (IH … HT12) -IH -T1 #L2 #T #HL12 #H destruct + append_length