X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fmultiple%2Ffleq.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fmultiple%2Ffleq.ma;h=cea4b650ed061aac096444b40da1633a8bf2eafb;hb=d2545ffd201b1aa49887313791386add78fa8603;hp=0000000000000000000000000000000000000000;hpb=57ae1762497a5f3ea75740e2908e04adb8642cc2;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/multiple/fleq.ma b/matita/matita/contribs/lambdadelta/basic_2A/multiple/fleq.ma new file mode 100644 index 000000000..cea4b650e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2A/multiple/fleq.ma @@ -0,0 +1,43 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2A/notation/relations/lazyeq_7.ma". +include "basic_2A/grammar/genv.ma". +include "basic_2A/multiple/lleq.ma". + +(* LAZY EQUIVALENCE FOR CLOSURES ********************************************) + +inductive fleq (l) (G) (L1) (T): relation3 genv lenv term ≝ +| fleq_intro: ∀L2. L1 ≡[T, l] L2 → fleq l G L1 T G L2 T +. + +interpretation + "lazy equivalence (closure)" + 'LazyEq l G1 L1 T1 G2 L2 T2 = (fleq l G1 L1 T1 G2 L2 T2). + +(* Basic properties *********************************************************) + +lemma fleq_refl: ∀l. tri_reflexive … (fleq l). +/2 width=1 by fleq_intro/ qed. + +lemma fleq_sym: ∀l. tri_symmetric … (fleq l). +#l #G1 #L1 #T1 #G2 #L2 #T2 * /3 width=1 by fleq_intro, lleq_sym/ +qed-. + +(* Basic inversion lemmas ***************************************************) + +lemma fleq_inv_gen: ∀G1,G2,L1,L2,T1,T2,l. ⦃G1, L1, T1⦄ ≡[l] ⦃G2, L2, T2⦄ → + ∧∧ G1 = G2 & L1 ≡[T1, l] L2 & T1 = T2. +#G1 #G2 #L1 #L2 #T1 #T2 #l * -G2 -L2 -T2 /2 width=1 by and3_intro/ +qed-.