From 85d772f5c3d5c86bbb474ba7ab4a259dc06687f9 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Sun, 18 Sep 2011 21:31:37 +0000 Subject: [PATCH] some improvements about the partial unfold on terms (tpss) --- .../Basic-2/grammar/term_weight.ma | 6 ++ .../lambda-delta/Basic-2/substitution/drop.ma | 11 ++- .../lambda-delta/Basic-2/substitution/ltps.ma | 4 +- .../lambda-delta/Basic-2/substitution/tps.ma | 12 +++ .../lambda-delta/Basic-2/unfold/tpss.ma | 10 +- .../lambda-delta/Basic-2/unfold/tpss_ltps.ma | 94 +++++++++++++++++++ .../lambda-delta/Basic-2/unfold/tpss_tpss.ma | 24 ++++- .../contribs/lambda-delta/Ground-2/star.ma | 14 ++- 8 files changed, 161 insertions(+), 14 deletions(-) create mode 100644 matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss_ltps.ma diff --git a/matita/matita/contribs/lambda-delta/Basic-2/grammar/term_weight.ma b/matita/matita/contribs/lambda-delta/Basic-2/grammar/term_weight.ma index 443cc2abc..d383ded4d 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/grammar/term_weight.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/grammar/term_weight.ma @@ -25,6 +25,12 @@ interpretation "weight (term)" 'Weight T = (tw T). (* Basic properties *********************************************************) +lemma tw_pos: ∀T. 1 ≤ #[T]. +#T elim T -T /2/ +qed. + +(* Basic eliminators ********************************************************) + axiom tw_wf_ind: ∀R:term→Prop. (∀T2. (∀T1. #[T1] < #[T2] → R T1) → R T2) → ∀T. R T. diff --git a/matita/matita/contribs/lambda-delta/Basic-2/substitution/drop.ma b/matita/matita/contribs/lambda-delta/Basic-2/substitution/drop.ma index 4d31115b2..729e3fa3d 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/substitution/drop.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/substitution/drop.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "Basic-2/grammar/lenv_weight.ma". include "Basic-2/grammar/leq.ma". include "Basic-2/substitution/lift.ma". @@ -177,7 +178,15 @@ lemma drop_fwd_drop2: ∀L1,I2,K2,V2,e. ↓[O, e] L1 ≡ K2. 𝕓{I2} V2 → ] qed. -lemma drop_fwd_drop2_length: ∀L1,I2,K2,V2,e. +lemma drop_fwd_lw: ∀L1,L2,d,e. ↓[d, e] L1 ≡ L2 → #[L2] ≤ #[L1]. +#L1 #L2 #d #e #H elim H -H L1 L2 d e // normalize +[ /2/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12 + >(tw_lift … HV21) -HV21 /2/ +] +qed. + +lemma drop_fwd_drop2_length: ∀L1,I2,K2,V2,e. ↓[0, e] L1 ≡ K2. 𝕓{I2} V2 → e < |L1|. #L1 elim L1 -L1 [ #I2 #K2 #V2 #e #H lapply (drop_inv_atom1 … H) -H #H destruct diff --git a/matita/matita/contribs/lambda-delta/Basic-2/substitution/ltps.ma b/matita/matita/contribs/lambda-delta/Basic-2/substitution/ltps.ma index 889704b44..3982ed2f9 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/substitution/ltps.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/substitution/ltps.ma @@ -77,7 +77,7 @@ fact ltps_inv_atom1_aux: ∀d,e,L1,L2. ] qed. -lemma drop_inv_atom1: ∀d,e,L2. ⋆ [d, e] ≫ L2 → L2 = ⋆. +lemma ltps_inv_atom1: ∀d,e,L2. ⋆ [d, e] ≫ L2 → L2 = ⋆. /2 width=5/ qed. fact ltps_inv_tps21_aux: ∀d,e,L1,L2. L1 [d, e] ≫ L2 → d = 0 → 0 < e → @@ -128,7 +128,7 @@ fact ltps_inv_atom2_aux: ∀d,e,L1,L2. ] qed. -lemma drop_inv_atom2: ∀d,e,L1. L1 [d, e] ≫ ⋆ → L1 = ⋆. +lemma ltps_inv_atom2: ∀d,e,L1. L1 [d, e] ≫ ⋆ → L1 = ⋆. /2 width=5/ qed. fact ltps_inv_tps22_aux: ∀d,e,L1,L2. L1 [d, e] ≫ L2 → d = 0 → 0 < e → 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 5a2f86aca..7197a53d9 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/substitution/tps.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "Basic-2/grammar/cl_weight.ma". include "Basic-2/substitution/drop.ma". (* PARALLEL SUBSTITUTION ON TERMS *******************************************) @@ -198,6 +199,17 @@ qed. lemma tps_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 [d, 0] ≫ T2 → T1 = T2. /2 width=6/ qed. +(* Basic forward lemmas *****************************************************) + +lemma tps_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → #[T1] ≤ #[T2]. +#L #T1 #T2 #d #e #H elim H normalize -H L T1 T2 d e +[ // +| /2/ +| /3 by monotonic_le_plus_l, le_plus/ (**) (* just /3/ is too slow *) +| /3 by monotonic_le_plus_l, le_plus/ (**) (* just /3/ is too slow *) +] +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/tpss.ma b/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss.ma index 2c90c9f86..6b63c4e1b 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss.ma @@ -22,14 +22,20 @@ definition tpss: nat → nat → lenv → relation term ≝ interpretation "partial unfold (term)" 'PSubstStar L T1 d e T2 = (tpss d e L T1 T2). -(* Basic properties *********************************************************) +(* Basic eliminators ********************************************************) lemma tpss_ind: ∀d,e,L,T1. ∀R: term → Prop. R T1 → - (∀T,T2. L ⊢ T1 [d, e] ≫* T → L ⊢ T[d, e] ≫ T2 → R T → R T2) → + (∀T,T2. L ⊢ T1 [d, e] ≫* T → L ⊢ T [d, e] ≫ T2 → R T → R T2) → ∀T2. L ⊢ T1 [d, e] ≫* T2 → R T2. #d #e #L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // 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. +/2/ qed. + lemma tpss_leq_repl_dx: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ≫* T2 → ∀L2. L1 [d, e] ≈ L2 → L2 ⊢ T1 [d, e] ≫* T2. /3/ qed. diff --git a/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss_ltps.ma b/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss_ltps.ma new file mode 100644 index 000000000..645d9d08d --- /dev/null +++ b/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss_ltps.ma @@ -0,0 +1,94 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "Basic-2/substitution/ltps_tps.ma". +include "Basic-2/unfold/tpss_tpss.ma". + +(* PARTIAL UNFOLD ON TERMS **************************************************) + +(* Properties concerning parallel substitution on local environments ********) + +lemma ltps_tpss_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. + d1 + e1 ≤ d2 → L0 [d1, e1] ≫ L1 → + L0 ⊢ T2 [d2, e2] ≫* U2 → L1 ⊢ T2 [d2, e2] ≫* U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HL01 #H @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU +lapply (ltps_tps_conf_ge … HU2 … HL01 ?) -HU2 HL01 /2/ +qed. + +lemma ltps_tpss_conf: ∀L0,L1,T2,U2,d1,e1,d2,e2. + L0 [d1, e1] ≫ L1 → L0 ⊢ T2 [d2, e2] ≫* U2 → + ∃∃T. L1 ⊢ T2 [d2, e2] ≫* T & L1 ⊢ U2 [d1, e1] ≫* T. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #HL01 #H @(tpss_ind … H) -U2 +[ /3/ +| #U #U2 #_ #HU2 * #T #HT2 #HUT + elim (ltps_tps_conf … HU2 … HL01) -HU2 HL01 #W #HUW #HU2W + elim (tpss_strip_eq … HUT … HUW) -U + /3 width=5 by ex2_1_intro, step, tpss_strap/ (**) (* just /3 width=5/ is too slow *) +] +qed. + +lemma ltps_tpss_trans_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. + d1 + e1 ≤ d2 → L1 [d1, e1] ≫ L0 → + L0 ⊢ T2 [d2, e2] ≫* U2 → L1 ⊢ T2 [d2, e2] ≫* U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HL10 #H @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU +lapply (ltps_tps_trans_ge … HU2 … HL10 ?) -HU2 HL10 /2/ +qed. + +lemma ltps_tpss_trans_down: ∀L0,L1,T2,U2,d1,e1,d2,e2. d2 + e2 ≤ d1 → + L1 [d1, e1] ≫ L0 → L0 ⊢ T2 [d2, e2] ≫* U2 → + ∃∃T. L1 ⊢ T2 [d2, e2] ≫* T & L0 ⊢ T [d1, e1] ≫* U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde2d1 #HL10 #H @(tpss_ind … H) -U2 +[ /3/ +| #U #U2 #_ #HU2 * #T #HT2 #HTU + elim (tpss_strap1_down … HTU … HU2 ?) -U // #U #HTU #HU2 + elim (ltps_tps_trans … HTU … HL10) -HTU HL10 #W #HTW #HWU + @(ex2_1_intro … W) /2/ (**) (* /3 width=5/ does not work as in ltps_tpss_conf *) +] +qed. + +fact ltps_tps_trans_eq_aux: ∀Y1,X2,L1,T2,U2,d,e. + L1 ⊢ T2 [d, e] ≫ U2 → ∀L0. L0 [d, e] ≫ L1 → + Y1 = L1 → X2 = T2 → L0 ⊢ T2 [d, e] ≫* U2. +#Y1 #X2 @(cw_wf_ind … Y1 X2) -Y1 X2 #Y1 #X2 #IH +#L1 #T2 #U2 #d #e * -L1 T2 U2 d e +[ // +| #L1 #K1 #V1 #W1 #i #d #e #Hdi #Hide #HLK1 #HVW1 #L0 #HL10 #H1 #H2 destruct -Y1 X2; + lapply (drop_fwd_lw … HLK1) normalize #H1 + elim (ltps_drop_trans_be … HL10 … HLK1 ? ?) -HL10 HLK1 [2,3: /2/ ] #X #H #HLK0 + elim (ltps_inv_tps22 … H ?) -H [2: /2/ ] #K0 #V0 #HK01 #HV01 #H destruct -X; + lapply (tps_fwd_tw … HV01) #H2 + lapply (transitive_le (#[K1] + #[V0]) … H1) -H1 [ /2/ ] -H2 #H + lapply (IH … HV01 … HK01 ? ?) -IH HV01 HK01 [1,3: // |2,4: skip | /2/ | /3 width=6/ ] +| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #L0 #HL0 #H1 #H2 destruct -Y1 X2; + lapply (tps_leq_repl_dx … HT12 (L. 𝕓{I} V1) ?) -HT12 /2/ #HT12 + lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3,5: // |2,4: skip ] #HV12 + lapply (IH … HT12 (L0. 𝕓{I} V1) ? ? ?) -IH HT12 [1,3,5: /2/ |2,4: skip | normalize // ] -HL0 #HT12 + lapply (tpss_leq_repl_dx … HT12 (L0. 𝕓{I} V2) ?) -HT12 /2/ +| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #L0 #HL0 #H1 #H2 destruct -Y1 X2; + lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3,5: // |2,4: skip ] + lapply (IH … HT12 … HL0 ? ?) -IH HT12 [1,3,5: // |2,4: skip ] -HL0 /2/ +] +qed. + +lemma ltps_tps_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 [d, e] ≫ U2 → + ∀L0. L0 [d, e] ≫ L1 → L0 ⊢ T2 [d, e] ≫* U2. +/2 width=5/ qed. + +lemma ltps_tpss_trans_eq: ∀L0,L1,T2,U2,d,e. L0 [d, e] ≫ L1 → + L1 ⊢ T2 [d, e] ≫* U2 → L0 ⊢ T2 [d, e] ≫* U2. +#L0 #L1 #T2 #U2 #d #e #HL01 #H @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU @(tpss_trans_eq … IHU) /2/ +qed. diff --git a/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss_tpss.ma b/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss_tpss.ma index 28248cf54..5cc6b38b7 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss_tpss.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/unfold/tpss_tpss.ma @@ -19,9 +19,25 @@ include "Basic-2/unfold/tpss_lift.ma". (* Advanced properties ******************************************************) -lemma tpss_trans_down_strap1: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ≫* T0 → - ∀T2,d2,e2. L ⊢ T0 [d2, e2] ≫ T2 → d2 + e2 ≤ d1 → - ∃∃T. L ⊢ T1 [d2, e2] ≫ T & L ⊢ T [d1, e1] ≫* T2. +lemma tpss_strip_eq: ∀L,T0,T1,d1,e1. L ⊢ T0 [d1, e1] ≫* T1 → + ∀T2,d2,e2. L ⊢ T0 [d2, e2] ≫ T2 → + ∃∃T. L ⊢ T1 [d2, e2] ≫ T & L ⊢ T2 [d1, e1] ≫* T. +/3/ qed. + +lemma tpss_strip_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 [d1, e1] ≫* T1 → + ∀L2,T2,d2,e2. L2 ⊢ T0 [d2, e2] ≫ T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. L2 ⊢ T1 [d2, e2] ≫ T & L1 ⊢ T2 [d1, e1] ≫* T. +/3/ qed. + +lemma tpss_strap1_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ≫* T0 → + ∀T2,d2,e2. L ⊢ T0 [d2, e2] ≫ T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 [d2, e2] ≫ T & L ⊢ T [d1, e1] ≫* T2. +/3/ qed. + +lemma tpss_strap2_down: ∀L,T1,T0,d1,e1. L ⊢ T1 [d1, e1] ≫ T0 → + ∀T2,d2,e2. L ⊢ T0 [d2, e2] ≫* T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 [d2, e2] ≫* T & L ⊢ T [d1, e1] ≫ T2. /3/ qed. lemma tpss_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫* T2 → @@ -31,7 +47,7 @@ lemma tpss_split_up: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫* T2 → [ /2/ | #T #T2 #_ #HT12 * #T3 #HT13 #HT3 elim (tps_split_up … HT12 … Hdi Hide) -HT12 Hide #T0 #HT0 #HT02 - elim (tpss_trans_down_strap1 … HT3 … HT0 ?) -T [2: