X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Freducibility%2Fltpr.ma;h=a910ea7df3f835c0520b6d1b5a6420915bf2f8f1;hb=5613a25cee29ef32a597cb4b44e8f2f4d71c4df0;hp=8e834b2dadaae3d79d6c4ce3961ddcae716fcb85;hpb=eb918fc784eacd2094e3986ba321ef47690d9983;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma index 8e834b2da..a910ea7df 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/ltpr.ma @@ -12,15 +12,12 @@ (* *) (**************************************************************************) -include "Basic_2/reducibility/tpr.ma". +include "basic_2/grammar/lenv_px.ma". +include "basic_2/reducibility/tpr.ma". (* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) -inductive ltpr: relation lenv ≝ -| ltpr_stom: ltpr (⋆) (⋆) -| ltpr_pair: ∀K1,K2,I,V1,V2. - ltpr K1 K2 → V1 ➡ V2 → ltpr (K1. ⓑ{I} V1) (K2. ⓑ{I} V2) (*ⓑ*) -. +definition ltpr: relation lenv ≝ lpx tpr. interpretation "context-free parallel reduction (environment)" @@ -28,62 +25,43 @@ interpretation (* Basic properties *********************************************************) -lemma ltpr_refl: ∀L:lenv. L ➡ L. -#L elim L -L // /2 width=1/ -qed. +lemma ltpr_refl: reflexive … ltpr. +/2 width=1/ qed. -(* Basic inversion lemmas ***************************************************) +lemma ltpr_append: ∀K1,K2. K1 ➡ K2 → ∀L1,L2:lenv. L1 ➡ L2 → K1 @@ L1 ➡ K2 @@ L2. +/2 width=1/ qed. -fact ltpr_inv_atom1_aux: ∀L1,L2. L1 ➡ L2 → L1 = ⋆ → L2 = ⋆. -#L1 #L2 * -L1 -L2 -[ // -| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct -] -qed. +(* Basic inversion lemmas ***************************************************) (* Basic_1: was: wcpr0_gen_sort *) lemma ltpr_inv_atom1: ∀L2. ⋆ ➡ L2 → L2 = ⋆. -/2 width=3/ qed-. - -fact ltpr_inv_pair1_aux: ∀L1,L2. L1 ➡ L2 → ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2. -#L1 #L2 * -L1 -L2 -[ #K1 #I #V1 #H destruct -| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct /2 width=5/ -] -qed. +/2 width=2 by lpx_inv_atom1/ qed-. (* Basic_1: was: wcpr0_gen_head *) lemma ltpr_inv_pair1: ∀K1,I,V1,L2. K1. ⓑ{I} V1 ➡ L2 → ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2. -/2 width=3/ qed-. - -fact ltpr_inv_atom2_aux: ∀L1,L2. L1 ➡ L2 → L2 = ⋆ → L1 = ⋆. -#L1 #L2 * -L1 -L2 -[ // -| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct -] -qed. +/2 width=1 by lpx_inv_pair1/ qed-. lemma ltpr_inv_atom2: ∀L1. L1 ➡ ⋆ → L1 = ⋆. -/2 width=3/ qed-. - -fact ltpr_inv_pair2_aux: ∀L1,L2. L1 ➡ L2 → ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1. -#L1 #L2 * -L1 -L2 -[ #K2 #I #V2 #H destruct -| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #K #J #W #H destruct /2 width=5/ -] -qed. +/2 width=2 by lpx_inv_atom2/ qed-. lemma ltpr_inv_pair2: ∀L1,K2,I,V2. L1 ➡ K2. ⓑ{I} V2 → ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1. -/2 width=3/ qed-. +/2 width=1 by lpx_inv_pair2/ qed-. (* Basic forward lemmas *****************************************************) lemma ltpr_fwd_length: ∀L1,L2. L1 ➡ L2 → |L1| = |L2|. -#L1 #L2 #H elim H -L1 -L2 normalize // -qed-. +/2 width=2 by lpx_fwd_length/ qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma ltpr_inv_append1: ∀K1,L1. ∀L:lenv. K1 @@ L1 ➡ L → + ∃∃K2,L2. K1 ➡ K2 & L1 ➡ L2 & L = K2 @@ L2. +/2 width=1 by lpx_inv_append1/ qed-. + +lemma ltpr_inv_append2: ∀L:lenv. ∀K2,L2. L ➡ K2 @@ L2 → + ∃∃K1,L1. K1 ➡ K2 & L1 ➡ L2 & L = K1 @@ L1. +/2 width=1 by lpx_inv_append2/ qed-. (* Basic_1: removed theorems 2: wcpr0_getl wcpr0_getl_back *)