]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/cpxs.ma
- bug fix in the induction for the closure property
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / cpxs.ma
index a11bb4892f9749cf0b10f133e02f3da220a0d856..ae51de39df124532c92aaa2db3ce7bda1cb82060 100644 (file)
@@ -13,7 +13,6 @@
 (**************************************************************************)
 
 include "basic_2/notation/relations/predstar_6.ma".
-include "basic_2/unfold/sstas.ma".
 include "basic_2/reduction/cnx.ma".
 include "basic_2/computation/cprs.ma".
 
@@ -61,11 +60,6 @@ lemma lsubr_cpxs_trans: ∀h,g,G. lsub_trans … (cpxs h g G) lsubr.
 /3 width=5 by lsubr_cpx_trans, TC_lsub_trans/
 qed-.
 
-lemma sstas_cpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 •* [h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
-#h #g #G #L #T1 #T2 #H @(sstas_ind … H) -T2 //
-/3 width=4 by cpxs_strap1, ssta_cpx/
-qed.
-
 lemma cprs_cpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
 #h #g #G #L #T1 #T2 #H @(cprs_ind … H) -T2 // /3 width=3/
 qed.
@@ -152,3 +146,14 @@ lemma cpxs_inv_cnx1: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T ➡*[h, g] U → ⦃G, L
 #T0 #T #H1T0 #_ #IHT #H2T0
 lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1/
 qed-.
+
+lemma cpxs_neq_inv_step_sn: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) →
+                            ∃∃T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T & T1 = T → ⊥ & ⦃G, L⦄ ⊢ T ➡*[h, g] T2.
+#h #g #G #L #T1 #T2 #H @(cpxs_ind_dx … H) -T1
+[ #H elim H -H //
+| #T1 #T #H1 #H2 #IH2 #H12 elim (eq_term_dec T1 T) #H destruct
+  [ -H1 -H2 /3 width=1 by/
+  | -IH2 /3 width=4 by ex3_intro/ (**) (* auto fails without clear *)
+  ]
+]
+qed-.