X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda-delta%2FBasic-2%2Freduction%2Ftpr_tps.ma;h=3a00c62c6b98a298ffb6ed1952f40bbdae09d769;hb=b264ad188cb0023a16dae105b037357fa75c5c1a;hp=96bc9b76ebaad9523456f06e711b5ce8e297ed8f;hpb=e4f11cddf44dd9bba21f689d4f56e2d00d8d7bb5;p=helm.git diff --git a/matita/matita/contribs/lambda-delta/Basic-2/reduction/tpr_tps.ma b/matita/matita/contribs/lambda-delta/Basic-2/reduction/tpr_tps.ma index 96bc9b76e..3a00c62c6 100644 --- a/matita/matita/contribs/lambda-delta/Basic-2/reduction/tpr_tps.ma +++ b/matita/matita/contribs/lambda-delta/Basic-2/reduction/tpr_tps.ma @@ -12,13 +12,13 @@ (* *) (**************************************************************************) -include "lambda-delta/reduction/lpr.ma". +include "Basic-2/reduction/ltpr.ma". (* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) -axiom tpr_tps_lpr: ∀L1,L2. L1 ⇒ L2 → ∀T1,T2. T1 ⇒ T2 → - ∀d,e,U1. L1 ⊢ T1 [d, e] ≫ U1 → - ∃∃U2. U1 ⇒ U2 & L2 ⊢ T2 [d, e] ≫ U2. +axiom tpr_tps_ltpr: ∀L1,L2. L1 ⇒ L2 → ∀T1,T2. T1 ⇒ T2 → + ∀d,e,U1. L1 ⊢ T1 [d, e] ≫ U1 → + ∃∃U2. U1 ⇒ U2 & L2 ⊢ T2 [d, e] ≫ U2. lemma tpr_tps_bind: ∀I,V1,V2,T1,T2,U1. V1 ⇒ V2 → T1 ⇒ T2 → ⋆. 𝕓{I} V1 ⊢ T1 [0, 1] ≫ U1 →