From c7731146663b37f09dc62f4f22924993203c1ba6 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Wed, 13 Jul 2011 19:54:24 +0000 Subject: [PATCH] - new definition of subst based on drop - pr is now based on drop --- .../lib/lambda-delta/reduction/pr_defs.ma | 17 +------ .../lambda-delta/substitution/drop_defs.ma | 45 ++++++++++++++++++ .../lambda-delta/substitution/drop_main.ma | 36 ++++++++++++++ .../lambda-delta/substitution/subst_defs.ma | 47 +++++++++++++++---- 4 files changed, 122 insertions(+), 23 deletions(-) create mode 100644 matita/matita/lib/lambda-delta/substitution/drop_defs.ma create mode 100644 matita/matita/lib/lambda-delta/substitution/drop_main.ma diff --git a/matita/matita/lib/lambda-delta/reduction/pr_defs.ma b/matita/matita/lib/lambda-delta/reduction/pr_defs.ma index e1cd216c2..b8ddd5ace 100644 --- a/matita/matita/lib/lambda-delta/reduction/pr_defs.ma +++ b/matita/matita/lib/lambda-delta/reduction/pr_defs.ma @@ -9,7 +9,7 @@ \ / V_______________________________________________________________ *) -include "lambda-delta/substitution/thin_defs.ma". +include "lambda-delta/substitution/drop_defs.ma". (* SINGLE STEP PARALLEL REDUCTION ON TERMS **********************************) @@ -24,7 +24,7 @@ inductive pr: lenv → term → term → Prop ≝ pr L V1 V2 → pr (L. 𝕓{Abst} W) T1 T2 → (*𝕓*) pr L (𝕚{Appl} V1. 𝕚{Abst} W. T1) (𝕚{Abbr} V2. T2) | pr_delta: ∀L,K,V1,V2,V,i. - ↓[0,i] L ≡ K. 𝕓{Abbr} V1 → pr K V1 V2 → ↑[0,i+1] V2 ≡ V → + ↑[0,i] K. 𝕓{Abbr} V1 ≡ L → pr K V1 V2 → ↑[0,i+1] V2 ≡ V → pr L (#i) V | pr_theta: ∀L,V,V1,V2,W1,W2,T1,T2. pr L V1 V2 → ↑[0,1] V2 ≡ V → pr L W1 W2 → pr (L. 𝕓{Abbr} W1) T1 T2 → (*𝕓*) @@ -44,16 +44,3 @@ lemma pr_refl: ∀T,L. L ⊢ T ⇒ T. #T elim T -T // #I elim I -I /2/ qed. -(* -lemma subst_pr: ∀d,e,L,T1,U2. L ⊢ ↓[d,e] T1 ≡ U2 → ∀T2. ↑[d,e] U2 ≡ T2 → - L ⊢ T1 ⇒ T2. -#d #e #L #T1 #U2 #H elim H -H d e L T1 U2 -[ #L #k #d #e #X #HX lapply (lift_inv_sort1 … HX) -HX #HX destruct -X // -| #L #i #d #e #Hid #X #HX lapply (lift_inv_sort1 … HX) -HX #HX destruct -X // -| #L #V1 #V2 #e #HV12 * #V #HV2 #HV1 - elim (lift_total 0 1 V1) #W1 #HVW1 - @(ex2_1_intro … W1) - [ - | /2 width=6/ - -*) \ No newline at end of file diff --git a/matita/matita/lib/lambda-delta/substitution/drop_defs.ma b/matita/matita/lib/lambda-delta/substitution/drop_defs.ma new file mode 100644 index 000000000..0644efdbe --- /dev/null +++ b/matita/matita/lib/lambda-delta/substitution/drop_defs.ma @@ -0,0 +1,45 @@ +(* + ||M|| This file is part of HELM, an Hypertextual, Electronic + ||A|| Library of Mathematics, developed at the Computer Science + ||T|| Department of the University of Bologna, Italy. + ||I|| + ||T|| + ||A|| This file is distributed under the terms of the + \ / GNU General Public License Version 2 + \ / + V_______________________________________________________________ *) + +include "lambda-delta/syntax/lenv.ma". +include "lambda-delta/substitution/lift_defs.ma". + +(* DROPPING *****************************************************************) + +inductive drop: lenv → nat → nat → lenv → Prop ≝ +| drop_refl: ∀L. drop L 0 0 L +| drop_drop: ∀L1,L2,I,V,e. drop L1 0 e L2 → drop (L1. 𝕓{I} V) 0 (e + 1) L2 +| drop_skip: ∀L1,L2,I,V1,V2,d,e. + drop L1 d e L2 → ↑[d,e] V2 ≡ V1 → + drop (L1. 𝕓{I} V1) (d + 1) e (L2. 𝕓{I} V2) +. + +interpretation "dropping" 'RLift L2 d e L1 = (drop L1 d e L2). + +(* the basic inversion lemmas ***********************************************) + +lemma drop_inv_skip2_aux: ∀d,e,L1,L2. ↑[d, e] L2 ≡ L1 → 0 < d → + ∀I,K2,V2. L2 = K2. 𝕓{I} V2 → + ∃∃K1,V1. ↑[d - 1, e] K2 ≡ K1 & + ↑[d - 1, e] V2 ≡ V1 & + L1 = K1. 𝕓{I} V1. +#d #e #L1 #L2 #H elim H -H d e L1 L2 +[ #L #H elim (lt_false … H) +| #L1 #L2 #I #V #e #_ #_ #H elim (lt_false … H) +| #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #_ #I #L2 #V2 #H destruct -X Y Z; + /2 width=5/ +] +qed. + +lemma drop_inv_skip2: ∀d,e,I,L1,K2,V2. ↑[d, e] K2. 𝕓{I} V2 ≡ L1 → 0 < d → + ∃∃K1,V1. ↑[d - 1, e] K2 ≡ K1 & ↑[d - 1, e] V2 ≡ V1 & + L1 = K1. 𝕓{I} V1. +/2/ qed. diff --git a/matita/matita/lib/lambda-delta/substitution/drop_main.ma b/matita/matita/lib/lambda-delta/substitution/drop_main.ma new file mode 100644 index 000000000..f456b9e07 --- /dev/null +++ b/matita/matita/lib/lambda-delta/substitution/drop_main.ma @@ -0,0 +1,36 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "lambda-delta/substitution/drop_defs.ma". + +(* DROPPING *****************************************************************) + +(* the main properties ******************************************************) + +axiom drop_conf_ge: ∀d1,e1,L,L1. ↑[d1, e1] L1 ≡ L → + ∀e2,L2. ↑[0, e2] L2 ≡ L → d1 + e1 ≤ e2 → + ↑[0, e2 - e1] L2 ≡ L1. + +axiom drop_conf_lt: ∀d1,e1,L,L1. ↑[d1, e1] L1 ≡ L → + ∀e2,K2,I,V2. ↑[0, e2] K2. 𝕓{I} V2 ≡ L → + e2 < d1 → let d ≝ d1 - e2 - 1 in + ∃∃K1,V1. ↑[0, e2] K1. 𝕓{I} V1 ≡ L1 & + ↑[d, e1] K2 ≡ K1 & ↑[d,e1] V1 ≡ V2. + +axiom drop_trans_le: ∀d1,e1,L1. ∀L:lenv. ↑[d1, e1] L ≡ L1 → + ∀e2,L2. ↑[0, e2] L2 ≡ L → e2 ≤ d1 → + ∃∃L0. ↑[0, e2] L0 ≡ L1 & ↑[d1 - e2, e1] L2 ≡ L0. + +axiom drop_trans_ge: ∀d1,e1,L1,L. ↑[d1, e1] L ≡ L1 → + ∀e2,L2. ↑[0, e2] L2 ≡ L → d1 ≤ e2 → ↑[0, e1 + e2] L2 ≡ L1. diff --git a/matita/matita/lib/lambda-delta/substitution/subst_defs.ma b/matita/matita/lib/lambda-delta/substitution/subst_defs.ma index 5ef343857..27de6a5ce 100644 --- a/matita/matita/lib/lambda-delta/substitution/subst_defs.ma +++ b/matita/matita/lib/lambda-delta/substitution/subst_defs.ma @@ -9,19 +9,17 @@ \ / V_______________________________________________________________ *) -include "lambda-delta/syntax/lenv.ma". -include "lambda-delta/substitution/lift_defs.ma". +include "lambda-delta/substitution/drop_defs.ma". (* TELESCOPIC SUBSTITUTION **************************************************) inductive subst: lenv → term → nat → nat → term → Prop ≝ | subst_sort : ∀L,k,d,e. subst L (⋆k) d e (⋆k) | subst_lref_lt: ∀L,i,d,e. i < d → subst L (#i) d e (#i) -| subst_lref_O : ∀L,V1,V2,e. subst L V1 0 e V2 → - subst (L. 𝕓{Abbr} V1) #0 0 (e + 1) V2 -| subst_lref_S : ∀L,I,V,i,T1,T2,d,e. - d ≤ i → i < d + e → subst L #i d e T1 → ↑[d,1] T1 ≡ T2 → - subst (L. 𝕓{I} V) #(i + 1) (d + 1) e T2 +| subst_lref_be: ∀L,K,V,U,i,d,e. + d ≤ i → i < d + e → + ↑[0, i] K. 𝕓{Abbr} V ≡ L → subst K V d (d + e - i - 1) U → + subst L (#i) d e U | subst_lref_ge: ∀L,i,d,e. d + e ≤ i → subst L (#i) d e (#(i - e)) | subst_bind : ∀L,I,V1,V2,T1,T2,d,e. subst L V1 d e V2 → subst (L. 𝕓{I} V1) T1 (d + 1) e T2 → @@ -33,9 +31,42 @@ inductive subst: lenv → term → nat → nat → term → Prop ≝ interpretation "telescopic substritution" 'RSubst L T1 d e T2 = (subst L T1 d e T2). +(* The basic inversion lemmas ***********************************************) + +lemma subst_inv_lref1_be_aux: ∀d,e,L,T,U. L ⊢ ↓[d, e] T ≡ U → + ∀i. d ≤ i → i < d + e → T = #i → + ∃∃K,V. ↑[0, i] K. 𝕓{Abbr} V ≡ L & + K ⊢ ↓[d, d + e - i - 1] V ≡ U. +#d #e #L #T #U #H elim H -H d e L T U +[ #L #k #d #e #i #_ #_ #H destruct +| #L #j #d #e #Hid #i #Hdi #_ #H destruct -j; + lapply (le_to_lt_to_lt … Hdi … Hid) -Hdi Hid #Hdd + elim (lt_false … Hdd) +| #L #K #V #U #j #d #e #_ #_ #HLK #HVU #_ #i #Hdi #Hide #H destruct -j /2/ +| #L #j #d #e #Hdei #i #_ #Hide #H destruct -j; + lapply (le_to_lt_to_lt … Hdei … Hide) -Hdei Hide #Hdede + elim (lt_false … Hdede) +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #_ #_ #i #_ #_ #H destruct +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #_ #_ #i #_ #_ #H destruct +] +qed. + +lemma subst_inv_lref1_be: ∀d,e,i,L,U. L ⊢ ↓[d, e] #i ≡ U → + d ≤ i → i < d + e → + ∃∃K,V. ↑[0, i] K. 𝕓{Abbr} V ≡ L & + K ⊢ ↓[d, d + e - i - 1] V ≡ U. +/2/ qed. + (* The basic properties *****************************************************) lemma subst_lift_inv: ∀d,e,T1,T2. ↑[d,e] T1 ≡ T2 → ∀L. L ⊢ ↓[d,e] T2 ≡ T1. #d #e #T1 #T2 #H elim H -H d e T1 T2 /2/ -#i #d #e #Hdi #L >(minus_plus_m_m i e) in ⊢ (? ? ? ? ? %) /3/ (**) (* use \ldots *) +#i #d #e #Hdi #L >(minus_plus_m_m i e) in ⊢ (? ? ? ? ? %) /3/ qed. +(* +| subst_lref_O : ∀L,V1,V2,e. subst L V1 0 e V2 → + subst (L. 𝕓{Abbr} V1) #0 0 (e + 1) V2 +| subst_lref_S : ∀L,I,V,i,T1,T2,d,e. + d ≤ i → i < d + e → subst L #i d e T1 → ↑[d,1] T2 ≡ T1 → + subst (L. 𝕓{I} V) #(i + 1) (d + 1) e T2 +*) -- 2.39.2