]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma
- exclusion binder added in local environments
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / syntax / tdeq.ma
index 4e5e41b4db741f9e53392f332b32e630c2df1304..9cf106377078485474223c9cda0f4649c672418b 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 ****************************************)
 
@@ -29,9 +29,6 @@ interpretation
    "degree-based equivalence (terms)"
    '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 →
@@ -131,7 +128,7 @@ lemma tdeq_sym: ∀h,o. symmetric … (tdeq h o).
 /2 width=3 by tdeq_sort, tdeq_lref, tdeq_gref, tdeq_pair/
 qed-.
 
-lemma tdeq_dec: ∀h,o,T1,T2. Decidable (tdeq h o T1 T2).
+lemma tdeq_dec: ∀h,o,T1,T2. Decidable (T1 ≡[h, o] T2).
 #h #o #T1 elim T1 -T1 [ * #s1 | #I1 #V1 #T1 #IHV #IHT ] * [1,3,5,7: * #s2 |*: #I2 #V2 #T2 ]
 [ elim (deg_total h o s1) #d1 #H1
   elim (deg_total h o s2) #d2 #H2
@@ -168,3 +165,16 @@ lemma tdeq_dec: ∀h,o,T1,T2. Decidable (tdeq h o T1 T2).
   elim (tdeq_inv_pair … H) -H /2 width=1 by/
 ]
 qed-.
+
+(* Negated inversion lemmas *************************************************)
+
+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 → ⊥)
+                      |  (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
+elim (tdeq_dec h o V1 V2) /3 width=1 by or3_intro1/
+elim (tdeq_dec h o T1 T2) /4 width=1 by tdeq_pair, or3_intro2/
+qed-.