X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Flreq.ma;h=747e9a511c5cf5eb11e454362772d21bd41dbe5b;hb=5c186c72f508da0849058afeecc6877cd9ed6303;hp=20541fef7a65224e215d644891261795d27d02e4;hpb=ebc170efe71cf4ee842acfbe58bb6864e76ba98c;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 20541fef7..747e9a511 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/grammar/ceq.ma". +include "basic_2/notation/relations/ideqsn_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). + 'IdEqSn f L1 L2 = (lreq f L1 L2). (* Basic properties *********************************************************) @@ -33,7 +33,7 @@ lemma lreq_eq_repl_back: ∀L1,L2. eq_repl_back … (λf. L1 ≡[f] L2). lemma lreq_eq_repl_fwd: ∀L1,L2. eq_repl_fwd … (λf. L1 ≡[f] L2). /2 width=3 by lexs_eq_repl_fwd/ qed-. -lemma sle_lreq_trans: ∀L1,L2,f2. L1 ≡[f2] L2 → +lemma sle_lreq_trans: ∀f2,L1,L2. L1 ≡[f2] L2 → ∀f1. f1 ⊆ f2 → L1 ≡[f1] L2. /2 width=3 by sle_lexs_trans/ qed-. @@ -43,55 +43,59 @@ lemma lreq_refl: ∀f. reflexive … (lreq f). (* Basic_2A1: includes: lreq_sym *) lemma lreq_sym: ∀f. symmetric … (lreq f). -#f #L1 #L2 #H elim H -L1 -L2 -f -/2 width=1 by lexs_next, lexs_push/ -qed-. +/3 width=2 by lexs_sym, cext2_sym/ qed-. (* Basic inversion lemmas ***************************************************) (* Basic_2A1: includes: lreq_inv_atom1 *) -lemma lreq_inv_atom1: ∀Y,f. ⋆ ≡[f] Y → Y = ⋆. +lemma lreq_inv_atom1: ∀f,Y. ⋆ ≡[f] Y → Y = ⋆. /2 width=4 by lexs_inv_atom1/ qed-. (* Basic_2A1: includes: lreq_inv_pair1 *) -lemma lreq_inv_next1: ∀J,K1,Y,W1,g. K1.ⓑ{J}W1 ≡[⫯g] Y → - ∃∃K2. K1 ≡[g] K2 & Y = K2.ⓑ{J}W1. -#J #K1 #Y #W1 #g #H elim (lexs_inv_next1 … H) -H /2 width=3 by ex2_intro/ +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 #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 *) -lemma lreq_inv_push1: ∀J,K1,Y,W1,g. K1.ⓑ{J}W1 ≡[↑g] Y → - ∃∃K2,W2. K1 ≡[g] K2 & Y = K2.ⓑ{J}W2. -#J #K1 #Y #W1 #g #H elim (lexs_inv_push1 … H) -H /2 width=4 by ex2_2_intro/ qed-. +lemma lreq_inv_push1: ∀g,J1,K1,Y. K1.ⓘ{J1} ≡[⫯g] Y → + ∃∃J2,K2. K1 ≡[g] K2 & Y = K2.ⓘ{J2}. +#g #J1 #K1 #Y #H elim (lexs_inv_push1 … H) -H /2 width=4 by ex2_2_intro/ +qed-. (* Basic_2A1: includes: lreq_inv_atom2 *) -lemma lreq_inv_atom2: ∀X,f. X ≡[f] ⋆ → X = ⋆. +lemma lreq_inv_atom2: ∀f,X. X ≡[f] ⋆ → X = ⋆. /2 width=4 by lexs_inv_atom2/ qed-. (* Basic_2A1: includes: lreq_inv_pair2 *) -lemma lreq_inv_next2: ∀J,X,K2,W2,g. X ≡[⫯g] K2.ⓑ{J}W2 → - ∃∃K1. K1 ≡[g] K2 & X = K1.ⓑ{J}W2. -#J #X #K2 #W2 #g #H elim (lexs_inv_next2 … H) -H /2 width=3 by ex2_intro/ qed-. +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 #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 *) -lemma lreq_inv_push2: ∀J,X,K2,W2,g. X ≡[↑g] K2.ⓑ{J}W2 → - ∃∃K1,W1. K1 ≡[g] K2 & X = K1.ⓑ{J}W1. -#J #X #K2 #W2 #g #H elim (lexs_inv_push2 … H) -H /2 width=4 by ex2_2_intro/ qed-. +lemma lreq_inv_push2: ∀g,J2,X,K2. X ≡[⫯g] K2.ⓘ{J2} → + ∃∃J1,K1. K1 ≡[g] K2 & X = K1.ⓘ{J1}. +#g #J2 #X #K2 #H elim (lexs_inv_push2 … H) -H /2 width=4 by ex2_2_intro/ +qed-. (* Basic_2A1: includes: lreq_inv_pair *) -lemma lreq_inv_next: ∀I1,I2,L1,L2,V1,V2,f. - L1.ⓑ{I1}V1 ≡[⫯f] (L2.ⓑ{I2}V2) → - ∧∧ L1 ≡[f] L2 & V1 = V2 & I1 = I2. -/2 width=1 by lexs_inv_next/ qed-. +lemma lreq_inv_next: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[↑f] L2.ⓘ{I2} → + L1 ≡[f] L2 ∧ I1 = I2. +#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: ∀I1,I2,L1,L2,V1,V2,f. - L1.ⓑ{I1}V1 ≡[↑f] (L2.ⓑ{I2}V2) → - L1 ≡[f] L2 ∧ I1 = I2. -#I1 #I2 #L1 #L2 #V1 #V2 #f #H elim (lexs_inv_push … H) -H /2 width=1 by conj/ +lemma lreq_inv_push: ∀f,I1,I2,L1,L2. L1.ⓘ{I1} ≡[⫯f] L2.ⓘ{I2} → L1 ≡[f] L2. +#f #I1 #I2 #L1 #L2 #H elim (lexs_inv_push … H) -H /2 width=1 by conj/ qed-. -lemma lreq_inv_tl: ∀I,L1,L2,V,f. L1 ≡[⫱f] L2 → L1.ⓑ{I}V ≡[f] L2.ⓑ{I}V. +lemma lreq_inv_tl: ∀f,I,L1,L2. L1 ≡[⫱f] L2 → L1.ⓘ{I} ≡[f] L2.ⓘ{I}. /2 width=1 by lexs_inv_tl/ qed-. (* Basic_2A1: removed theorems 5: