]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/lfdeq_lfdeq.ma
advances on lfsx ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / lfdeq_lfdeq.ma
index 71b1c0d4f668c84bad3295e3ae5f6ae7d1afa745..9f92ecd272e3c8aa1c5e13eedeb557b1a5c79a19 100644 (file)
@@ -18,6 +18,11 @@ include "basic_2/static/lfdeq.ma".
 
 (* DEGREE-BASED EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ******)
 
+(* Advanced properties ******************************************************)
+
+lemma lfdeq_dec: ∀h,o,L1,L2. ∀T:term. Decidable (L1 ≡[h, o, T] L2).
+/3 width=1 by lfxs_dec, tdeq_dec/ qed-.
+
 (* Main properties **********************************************************)
 
 theorem lfdeq_bind: ∀h,o,p,I,L1,L2,V1,V2,T.
@@ -42,6 +47,10 @@ theorem lfdeq_canc_sn: ∀h,o,T. left_cancellable … (lfdeq h o T).
 theorem lfdeq_canc_dx: ∀h,o,T. right_cancellable … (lfdeq h o T).
 /3 width=3 by lfdeq_trans, lfdeq_sym/ qed-.
 
+theorem lfdeq_repl: ∀h,o,L1,L2. ∀T:term. L1 ≡[h, o, T] L2 →
+                    ∀K1. L1 ≡[h, o, T] K1 → ∀K2. L2 ≡[h, o, T] K2 → K1 ≡[h, o, T] K2.
+/3 width=3 by lfdeq_canc_sn, lfdeq_trans/ qed-.
+
 (* Advanced properies on negated lazy equivalence *****************************)
 
 (* Note: auto works with /4 width=8/ so lfdeq_canc_sn is preferred ************)