]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma
- exclusion binder in local environments:
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / syntax / tdeq.ma
index 52784b53fa8dca6d6a61718bc8681dc2e51f8017..43f538794df2fb16772bdf15fecebe5f68c1ad2f 100644 (file)
@@ -14,7 +14,7 @@
 
 include "basic_2/notation/relations/lazyeq_4.ma".
 include "basic_2/syntax/item_sd.ma".
-include "basic_2/syntax/lenv.ma".
+include "basic_2/syntax/term.ma".
 
 (* DEGREE-BASED EQUIVALENCE ON TERMS ****************************************)
 
@@ -26,12 +26,9 @@ inductive tdeq (h) (o): relation term ≝
 .
 
 interpretation
-   "degree-based equivalence (terms)"
+   "context-free degree-based equivalence (term)"
    'LazyEq h o T1 T2 = (tdeq h o T1 T2).
 
-definition cdeq: ∀h. sd h → relation3 lenv term term ≝
-                 λh,o,L. tdeq h o.
-
 (* Basic inversion lemmas ***************************************************)
 
 fact tdeq_inv_sort1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀s1. X = ⋆s1 →
@@ -174,7 +171,7 @@ qed-.
 lemma tdneq_inv_pair: ∀h,o,I1,I2,V1,V2,T1,T2.
                       (②{I1}V1.T1 ≡[h, o] ②{I2}V2.T2 → ⊥) → 
                       ∨∨ I1 = I2 → ⊥
-                      |  V1 ≡[h, o] V2 → ⊥
+                      |  (V1 ≡[h, o] V2 → ⊥)
                       |  (T1 ≡[h, o] T2 → ⊥).
 #h #o #I1 #I2 #V1 #V2 #T1 #T2 #H12
 elim (eq_item2_dec I1 I2) /3 width=1 by or3_intro0/ #H destruct