From 2d7c5c2bc32162612138e329a48c8108e010334d Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Tue, 14 Jun 2011 18:24:05 +0000 Subject: [PATCH] more lemmas to prove and a correction in subst --- .../matita/lib/lambda-delta/reduction/pr.ma | 79 +++++-------------- .../lib/lambda-delta/substitution/lift.ma | 2 + .../lib/lambda-delta/substitution/subst.ma | 3 +- .../lib/lambda-delta/substitution/thin.ma | 27 +++++++ 4 files changed, 49 insertions(+), 62 deletions(-) diff --git a/matita/matita/lib/lambda-delta/reduction/pr.ma b/matita/matita/lib/lambda-delta/reduction/pr.ma index 3abebc848..3c5f4ca96 100644 --- a/matita/matita/lib/lambda-delta/reduction/pr.ma +++ b/matita/matita/lib/lambda-delta/reduction/pr.ma @@ -11,7 +11,7 @@ include "lambda-delta/substitution/thin.ma". -(* SINGLE STEP PARALLEL REDUCTION *******************************************) +(* SINGLE STEP PARALLEL REDUCTION ON TERMS **********************************) inductive pr: lenv → term → term → Prop ≝ | pr_sort : ∀L,k. pr L (⋆k) (⋆k) @@ -34,69 +34,26 @@ inductive pr: lenv → term → term → Prop ≝ | pr_tau : ∀L,V,T1,T2. pr L T1 T2 → pr L (𝕚{Cast} V. T1) T2 . -interpretation "single step parallel reduction" 'PR L T1 T2 = (pr L T1 T2). +interpretation + "single step parallel reduction (term)" + 'PR L T1 T2 = (pr L T1 T2). -(* The three main lemmas on reduction ***************************************) +(* The basic properties *****************************************************) -lemma pr_inv_lift: ∀L,T1,T2. L ⊢ T1 ⇒ T2 → - ∀d,e,K. ↓[d,e] L ≡ K → ∀U1. ↑[d,e] U1 ≡ T1 → - ∃∃U2. ↑[d,e] U2 ≡ T2 & K ⊢ U1 ⇒ U2. -#L #T1 #T2 #H elim H -H L T1 T2 -[ #L #k #d #e #K #_ #U1 #HU1 - lapply (lift_inv_sort2 … HU1) -HU1 #H destruct -U1 /2/ -| #L #i #d #e #K #_ #U1 #HU1 - lapply (lift_inv_lref2 … HU1) -HU1 * * #Hid #H destruct -U1 /3/ -| #L #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HLK #X #HX - lapply (lift_inv_bind2 … HX) -HX * #V0 #T0 #HV01 #HT01 #HX destruct -X; - elim (IHV12 … HLK … HV01) -IHV12 #V3 #HV32 #HV03 - elim (IHT12 … HT01) -IHT12 HT01 [2,3: -HV32 HV03 /3/] -HLK HV01 /3 width=5/ -| #L #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct -X; - elim (IHV12 … HLK … HV01) -IHV12 HV01 #V3 #HV32 #HV03 - elim (IHT12 … HLK … HT01) -IHT12 HT01 HLK /3 width=5/ -| #L #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct -X; - elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct -Y; - elim (IHV12 … HLK … HV01) -IHV12 HV01 #V3 #HV32 #HV03 - elim (IHT12 … HT01) -IHT12 HT01 - [3: -HV32 HV03 @(thin_skip … HLK) /2/ |2: skip ] (**) (* /3 width=5/ is too slow *) - -HLK HW01 - /3 width=5/ -| #L #K0 #V1 #V2 #V0 #i #HLK0 #HV12 #HV20 #IHV12 #d #e #K #HLK #X #HX - lapply (lift_inv_lref2 … HX) -HX * * #Hid #HX destruct -X; - [ -HV12; - elim (thin_conf_lt … HLK … HLK0 Hid) -HLK HLK0 L #L #V3 #HKL #HK0L #HV31 - elim (IHV12 … HK0L … HV31) -IHV12 HK0L HV31 #V4 #HV42 #HV34 - elim (lift_trans_le … HV42 … HV20 ?) -HV42 HV20 V2 // #V2 #HV42 - >arith5 // -Hid #HV20 - @(ex2_1_intro … V2) /2 width=6/ (**) (* /3 width=8/ is slow *) - | -IHV12; - lapply (thin_conf_ge … HLK … HLK0 Hid) -HLK HLK0 L #HK - elim (lift_free … HV20 d (i - e + 1) ? ? ?) -HV20 /2/ - >arith3 /2/ -Hid /3 width=8/ (**) (* just /3 width=8/ is a bit slow *) - ] -| #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct -X; - elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct -Y; - elim (IHV12 ? ? ? HLK ? HV01) -IHV12 HV01 #V3 #HV32 #HV03 - elim (IHW12 ? ? ? HLK ? HW01) -IHW12 #W3 #HW32 #HW03 - elim (IHT12 … HT01) -IHT12 HT01 - [3: -HV2 HV32 HV03 HW32 HW03 @(thin_skip … HLK) /2/ |2: skip ] (**) (* /3/ is too slow *) - -HLK HW01 #T3 #HT32 #HT03 - elim (lift_trans_le … HV32 … HV2 ?) -HV32 HV2 V2 // #V2 #HV32 #HV2 - @(ex2_1_intro … (𝕓{Abbr}W3.𝕗{Appl}V2.T3)) /3/ (**) (* /4/ loops *) -| #L #V #T #T1 #T2 #HT1 #_ #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_bind2 … HX) -HX #V0 #T0 #_ #HT0 #H destruct -X; - elim (lift_conf_rev … HT1 … HT0 ?) -HT1 HT0 T // #T #HT0 #HT1 - elim (IHT12 … HLK … HT1) -IHT12 HLK HT1 /3 width=5/ -| #L #V #T1 #T2 #_ #IHT12 #d #e #K #HLK #X #HX - elim (lift_inv_flat2 … HX) -HX #V0 #T0 #_ #HT01 #H destruct -X; - elim (IHT12 … HLK … HT01) -IHT12 HLK HT01 /3/ -] -qed. - -(* this may be moved *) 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/lift.ma b/matita/matita/lib/lambda-delta/substitution/lift.ma index acd84df08..60b1b62e7 100644 --- a/matita/matita/lib/lambda-delta/substitution/lift.ma +++ b/matita/matita/lib/lambda-delta/substitution/lift.ma @@ -169,6 +169,8 @@ qed. (* the main properies *******************************************************) +axiom lift_total: ∀d,e,T1. ∃T2. ↑[d,e] T1 ≡ T2. + axiom lift_mono: ∀d,e,T,U1. ↑[d,e] T ≡ U1 → ∀U2. ↑[d,e] T ≡ U2 → U1 = U2. theorem lift_conf_rev: ∀d1,e1,T1,T. ↑[d1,e1] T1 ≡ T → diff --git a/matita/matita/lib/lambda-delta/substitution/subst.ma b/matita/matita/lib/lambda-delta/substitution/subst.ma index a9732529d..18eb6c526 100644 --- a/matita/matita/lib/lambda-delta/substitution/subst.ma +++ b/matita/matita/lib/lambda-delta/substitution/subst.ma @@ -17,7 +17,8 @@ include "lambda-delta/substitution/lift.ma". 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,V,e. 0 < e → subst (L. 𝕓{Abbr} V) #0 0 e V +| 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 diff --git a/matita/matita/lib/lambda-delta/substitution/thin.ma b/matita/matita/lib/lambda-delta/substitution/thin.ma index 96432922d..be7957422 100644 --- a/matita/matita/lib/lambda-delta/substitution/thin.ma +++ b/matita/matita/lib/lambda-delta/substitution/thin.ma @@ -23,6 +23,26 @@ inductive thin: lenv → nat → nat → lenv → Prop ≝ interpretation "thinning" 'RSubst L1 d e L2 = (thin L1 d e L2). +(* the basic inversion lemmas ***********************************************) + +lemma thin_inv_skip2_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → 0 < d → + ∀I,K2,V2. L2 = K2. 𝕓{I} V2 → + ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & + K1 ⊢ ↓[d - 1, e] V1 ≡ V2 & + 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 thin_inv_skip2: ∀d,e,I,L1,K2,V2. ↓[d, e] L1 ≡ K2. 𝕓{I} V2 → 0 < d → + ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & K1 ⊢ ↓[d - 1, e] V1 ≡ V2 & + L1 = K1. 𝕓{I} V1. +/2/ qed. + (* the main properties ******************************************************) axiom thin_conf_ge: ∀d1,e1,L,L1. ↓[d1,e1] L ≡ L1 → @@ -32,3 +52,10 @@ axiom thin_conf_lt: ∀d1,e1,L,L1. ↓[d1,e1] L ≡ L1 → ∀e2,K2,I,V2. ↓[0,e2] L ≡ K2. 𝕓{I} V2 → e2 < d1 → let d ≝ d1 - e2 - 1 in ∃∃K1,V1. ↓[0,e2] L1 ≡ K1. 𝕓{I} V1 & ↓[d,e1] K2 ≡ K1 & ↑[d,e1] V1 ≡ V2. + +axiom thin_trans_le: ∀d1,e1,L1,L. ↓[d1, e1] L1 ≡ L → + ∀e2,L2. ↓[0, e2] L ≡ L2 → e2 ≤ d1 → + ∃∃L0. ↓[0, e2] L1 ≡ L0 & ↓[d1 - e2, e1] L0 ≡ L2. + +axiom thin_trans_ge: ∀d1,e1,L1,L. ↓[d1, e1] L1 ≡ L → + ∀e2,L2. ↓[0, e2] L ≡ L2 → d1 ≤ e2 → ↓[0, e1 + e2] L1 ≡ L2. -- 2.39.2