X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Flambda-delta%2Freduction%2Fpr_defs.ma;h=b8ddd5aceb69ae7036a6a7880dfc030d74c4eee2;hb=c7731146663b37f09dc62f4f22924993203c1ba6;hp=e1cd216c2277f253d070b1bce773618e94b101d7;hpb=13935d33cc0899b9555648a4d49586e17274c748;p=helm.git 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