]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma
csx on the way ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / syntax / tdeq.ma
index c79c1940cb24246d3da9c2631939b6f4c0f40a02..4e5e41b4db741f9e53392f332b32e630c2df1304 100644 (file)
@@ -32,19 +32,6 @@ interpretation
 definition cdeq: ∀h. sd h → relation3 lenv term term ≝
                  λh,o,L. tdeq h o.
 
-(* Basic properties *********************************************************)
-
-lemma tdeq_refl: ∀h,o. reflexive … (tdeq h o).
-#h #o #T elim T -T /2 width=1 by tdeq_pair/
-* /2 width=1 by tdeq_lref, tdeq_gref/
-#s elim (deg_total h o s) /2 width=3 by tdeq_sort/
-qed.
-
-lemma tdeq_sym: ∀h,o. symmetric … (tdeq h o).
-#h #o #T1 #T2 #H elim H -T1 -T2
-/2 width=3 by tdeq_sort, tdeq_lref, tdeq_gref, tdeq_pair/
-qed-.
-
 (* Basic inversion lemmas ***************************************************)
 
 fact tdeq_inv_sort1_aux: ∀h,o,X,Y. X ≡[h, o] Y → ∀s1. X = ⋆s1 →
@@ -100,7 +87,15 @@ lemma tdeq_inv_pair1: ∀h,o,I,V1,T1,Y. ②{I}V1.T1 ≡[h, o] Y →
 lemma tdeq_inv_sort1_deg: ∀h,o,Y,s1. ⋆s1 ≡[h, o] Y → ∀d. deg h o s1 d →
                           ∃∃s2. deg h o s2 d & Y = ⋆s2.
 #h #o #Y #s1 #H #d #Hs1 elim (tdeq_inv_sort1 … H) -H
-#s2 #x #Hx <(deg_mono h o … Hx … Hs1) -s1 -d /2 width=3 by ex2_intro/  
+#s2 #x #Hx <(deg_mono h o … Hx … Hs1) -s1 -d /2 width=3 by ex2_intro/
+qed-.
+
+lemma tdeq_inv_sort_deg: ∀h,o,s1,s2. ⋆s1 ≡[h, o] ⋆s2 →
+                         ∀d1,d2. deg h o s1 d1 → deg h o s2 d2 →
+                         d1 = d2.
+#h #o #s1 #y #H #d1 #d2 #Hs1 #Hy
+elim (tdeq_inv_sort1_deg … H … Hs1) -s1 #s2 #Hs2 #H destruct
+<(deg_mono h o … Hy … Hs2) -s2 -d1 //
 qed-.
 
 lemma tdeq_inv_pair: ∀h,o,I1,I2,V1,V2,T1,T2. ②{I1}V1.T1 ≡[h, o] ②{I2}V2.T2 →
@@ -122,3 +117,54 @@ lemma tdeq_fwd_atom1: ∀h,o,I,Y. ⓪{I} ≡[h, o] Y → ∃J. Y = ⓪{J}.
 #h #o * #x #Y #H [ elim (tdeq_inv_sort1 … H) -H ]
 /3 width=4 by tdeq_inv_gref1, tdeq_inv_lref1, ex_intro/
 qed-.
+
+(* Basic properties *********************************************************)
+
+lemma tdeq_refl: ∀h,o. reflexive … (tdeq h o).
+#h #o #T elim T -T /2 width=1 by tdeq_pair/
+* /2 width=1 by tdeq_lref, tdeq_gref/
+#s elim (deg_total h o s) /2 width=3 by tdeq_sort/
+qed.
+
+lemma tdeq_sym: ∀h,o. symmetric … (tdeq h o).
+#h #o #T1 #T2 #H elim H -T1 -T2
+/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).
+#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
+  elim (eq_nat_dec d1 d2) #Hd12 destruct /3 width=3 by tdeq_sort, or_introl/
+  @or_intror #H
+  lapply (tdeq_inv_sort_deg … H … H1 H2) -H -H1 -H2 /2 width=1 by/
+|2,3,13:
+  @or_intror #H
+  elim (tdeq_inv_sort1 … H) -H #x1 #x2 #_ #_ #H destruct
+|4,6,14:
+  @or_intror #H
+  lapply (tdeq_inv_lref1 … H) -H #H destruct
+|5:
+  elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/
+  @or_intror #H
+  lapply (tdeq_inv_lref1 … H) -H #H destruct /2 width=1 by/
+|7,8,15:
+  @or_intror #H
+  lapply (tdeq_inv_gref1 … H) -H #H destruct
+|9:
+  elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/
+  @or_intror #H
+  lapply (tdeq_inv_gref1 … H) -H #H destruct /2 width=1 by/
+|10,11,12:
+  @or_intror #H
+  elim (tdeq_inv_pair1 … H) -H #X1 #X2 #_ #_ #H destruct
+|16:
+  elim (eq_item2_dec I1 I2) #HI12 destruct
+  [ elim (IHV V2) -IHV #HV12
+    elim (IHT T2) -IHT #HT12
+    [ /3 width=1 by tdeq_pair, or_introl/ ]
+  ]
+  @or_intror #H
+  elim (tdeq_inv_pair … H) -H /2 width=1 by/
+]
+qed-.