X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Flfeq.ma;h=684e57370908a1ac8abdd71a4b1b5b0c4d225043;hp=1ad709c9e3c39ab8a9d71cef719b8ba0ca248e33;hb=f7296f9cf2ee73465a374942c46b138f35c42ccb;hpb=e8971d3db8935436e6dddc27fe1ae168c7f3b315 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfeq.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfeq.ma index 1ad709c9e..684e57370 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lfeq.ma @@ -25,15 +25,13 @@ interpretation "syntactic equivalence on referred entries (local environment)" 'LazyEqSn T L1 L2 = (lfeq T L1 L2). +(* Note: "lfeq_transitive R" is equivalent to "lfxs_transitive ceq R R" *) (* Basic_2A1: uses: lleq_transitive *) definition lfeq_transitive: predicate (relation3 lenv term term) ≝ λR. ∀L2,T1,T2. R L2 T1 T2 → ∀L1. L1 ≡[T1] L2 → R L1 T1 T2. (* Basic inversion lemmas ***************************************************) -lemma lfeq_transitive_inv_lfxs: ∀R. lfeq_transitive R → lfxs_transitive ceq R R. -/2 width=3 by/ qed-. - lemma lfeq_inv_bind: ∀p,I,L1,L2,V,T. L1 ≡[ⓑ{p,I}V.T] L2 → ∧∧ L1 ≡[V] L2 & L1.ⓑ{I}V ≡[T] L2.ⓑ{I}V. /2 width=2 by lfxs_inv_bind/ qed-. @@ -78,9 +76,6 @@ qed-. (* Basic_properties *********************************************************) -lemma lfxs_transitive_lfeq: ∀R. lfxs_transitive ceq R R → lfeq_transitive R. -/2 width=5 by/ qed. - lemma frees_lfeq_conf: ∀f,L1,T. L1 ⊢ 𝐅*⦃T⦄ ≡ f → ∀L2. L1 ≡[T] L2 → L2 ⊢ 𝐅*⦃T⦄ ≡ f. #f #L1 #T #H elim H -f -L1 -T