X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=inline;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Flreq.ma;h=cc3e75b0a5107964bdd5b503c9e8904aca6fa0d4;hb=42705ef31dd3513a998533e02b5f20fb38dd4fb2;hp=6f8e3bc74cc3ca54b622c01d8b6f1a3b3323666e;hpb=98fbba1b68d457807c73ebf70eb2a48696381da4;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq.ma index 6f8e3bc74..cc3e75b0a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lreq.ma @@ -12,18 +12,18 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/lazyeq_3.ma". -include "basic_2/syntax/ext2.ma". +include "basic_2/notation/relations/lazyeqsn_3.ma". +include "basic_2/syntax/ceq_ext.ma". include "basic_2/relocation/lexs.ma". (* RANGED EQUIVALENCE FOR LOCAL ENVIRONMENTS ********************************) (* Basic_2A1: includes: lreq_atom lreq_zero lreq_pair lreq_succ *) -definition lreq: relation3 rtmap lenv lenv ≝ lexs ceq cfull. +definition lreq: relation3 rtmap lenv lenv ≝ lexs ceq_ext cfull. interpretation "ranged equivalence (local environment)" - 'LazyEq f L1 L2 = (lreq f L1 L2). + 'LazyEqSn f L1 L2 = (lreq f L1 L2). (* Basic properties *********************************************************) @@ -39,11 +39,11 @@ lemma sle_lreq_trans: ∀f2,L1,L2. L1 ≡[f2] L2 → (* Basic_2A1: includes: lreq_refl *) lemma lreq_refl: ∀f. reflexive … (lreq f). -/3 width=1 by lexs_refl, ext2_refl/ qed. +/2 width=1 by lexs_refl/ qed. (* Basic_2A1: includes: lreq_sym *) lemma lreq_sym: ∀f. symmetric … (lreq f). -/3 width=1 by lexs_sym, ext2_sym/ qed-. +/3 width=2 by lexs_sym, cext2_sym/ qed-. (* Basic inversion lemmas ***************************************************) @@ -54,7 +54,9 @@ lemma lreq_inv_atom1: ∀f,Y. ⋆ ≡[f] Y → Y = ⋆. (* Basic_2A1: includes: lreq_inv_pair1 *) lemma lreq_inv_next1: ∀g,J,K1,Y. K1.ⓘ{J} ≡[⫯g] Y → ∃∃K2. K1 ≡[g] K2 & Y = K2.ⓘ{J}. -#g #J #K1 #Y #H elim (lexs_inv_next1 … H) -H /2 width=3 by ex2_intro/ +#g #J #K1 #Y #H +elim (lexs_inv_next1 … H) -H #Z #K2 #HK12 #H1 #H2 destruct +<(ceq_ext_inv_eq … H1) -Z /2 width=3 by ex2_intro/ qed-. (* Basic_2A1: includes: lreq_inv_zero1 lreq_inv_succ1 *) @@ -70,7 +72,9 @@ lemma lreq_inv_atom2: ∀f,X. X ≡[f] ⋆ → X = ⋆. (* Basic_2A1: includes: lreq_inv_pair2 *) lemma lreq_inv_next2: ∀g,J,X,K2. X ≡[⫯g] K2.ⓘ{J} → ∃∃K1. K1 ≡[g] K2 & X = K1.ⓘ{J}. -#g #J #X #K2 #H elim (lexs_inv_next2 … H) -H /2 width=3 by ex2_intro/ +#g #J #X #K2 #H +elim (lexs_inv_next2 … H) -H #Z #K1 #HK12 #H1 #H2 destruct +<(ceq_ext_inv_eq … H1) -J /2 width=3 by ex2_intro/ qed-. (* Basic_2A1: includes: lreq_inv_zero2 lreq_inv_succ2 *) @@ -82,7 +86,9 @@ qed-. (* Basic_2A1: includes: lreq_inv_pair *) lemma lreq_inv_next: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[⫯f] L2.ⓘ{I2} → L1 ≡[f] L2 ∧ I1 = I2. -/2 width=1 by lexs_inv_next/ qed-. +#f #I1 #I2 #L1 #L2 #H elim (lexs_inv_next … H) -H +/3 width=3 by ceq_ext_inv_eq, conj/ +qed-. (* Basic_2A1: includes: lreq_inv_succ *) lemma lreq_inv_push: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[↑f] L2.ⓘ{I2} → L1 ≡[f] L2.