X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Freducibility%2Fltpr.ma;h=8e834b2dadaae3d79d6c4ce3961ddcae716fcb85;hb=44c1079dabf1d3c0b69d0155ddbaea8627ec901c;hp=ecb0ad08ceb04959634bfd3b7d9d08ef76fa736a;hpb=f75be90562ddd964ef7ed43b956eb908f3133e3a;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 ecb0ad08c..8e834b2da 100644 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr.ma +++ b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/ltpr.ma @@ -19,7 +19,7 @@ include "Basic_2/reducibility/tpr.ma". 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) (*𝕓*) + ltpr K1 K2 → V1 ➡ V2 → ltpr (K1. ⓑ{I} V1) (K2. ⓑ{I} V2) (*ⓑ*) . interpretation @@ -28,62 +28,62 @@ interpretation (* Basic properties *********************************************************) -lemma ltpr_refl: ∀L:lenv. L ⇒ L. -#L elim L -L /2/ +lemma ltpr_refl: ∀L:lenv. L ➡ L. +#L elim L -L // /2 width=1/ qed. (* Basic inversion lemmas ***************************************************) -fact ltpr_inv_atom1_aux: ∀L1,L2. L1 ⇒ L2 → L1 = ⋆ → L2 = ⋆. -#L1 #L2 * -L1 L2 +fact ltpr_inv_atom1_aux: ∀L1,L2. L1 ➡ L2 → L1 = ⋆ → L2 = ⋆. +#L1 #L2 * -L1 -L2 [ // | #K1 #K2 #I #V1 #V2 #_ #_ #H destruct ] qed. (* Basic_1: was: wcpr0_gen_sort *) -lemma ltpr_inv_atom1: ∀L2. ⋆ ⇒ L2 → L2 = ⋆. -/2/ qed. +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 +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 - K1 I V1 /2 width=5/ +| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #L #J #W #H destruct /2 width=5/ ] 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/ qed. +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 +fact ltpr_inv_atom2_aux: ∀L1,L2. L1 ➡ L2 → L2 = ⋆ → L1 = ⋆. +#L1 #L2 * -L1 -L2 [ // | #K1 #K2 #I #V1 #V2 #_ #_ #H destruct ] qed. -lemma ltpr_inv_atom2: ∀L1. L1 ⇒ ⋆ → L1 = ⋆. -/2/ 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 +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 -K2 I V2 /2 width=5/ +| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #K #J #W #H destruct /2 width=5/ ] qed. -lemma ltpr_inv_pair2: ∀L1,K2,I,V2. L1 ⇒ K2. 𝕓{I} V2 → - ∃∃K1,V1. K1 ⇒ K2 & V1 ⇒ V2 & L1 = K1. 𝕓{I} V1. -/2/ 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-. (* Basic forward lemmas *****************************************************) -lemma ltpr_fwd_length: ∀L1,L2. L1 ⇒ L2 → |L1| = |L2|. -#L1 #L2 #H elim H -H L1 L2; normalize // -qed. +lemma ltpr_fwd_length: ∀L1,L2. L1 ➡ L2 → |L1| = |L2|. +#L1 #L2 #H elim H -L1 -L2 normalize // +qed-. (* Basic_1: removed theorems 2: wcpr0_getl wcpr0_getl_back *)