]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lsx_csx.ma
some advances on reduction
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lsx_csx.ma
index 48786e5e2cc48ec96ff4e53a292310b86fabcf3b..791a0e526c57fa8c989b6ec1f5bdb14484d0b6c1 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/computation/lpxs_lleq.ma".
+include "basic_2/reduction/cpx_cpys.ma".
+include "basic_2/computation/lpxs_cpye.ma".
 include "basic_2/computation/csx_alt.ma".
-include "basic_2/computation/lsx.ma".
+include "basic_2/computation/lsx_lpxs.ma".
 
 (* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
 
-(* Main properties **********************************************************)
-
-axiom csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → G ⊢ ⋕⬊*[h, g, T] L.
-
-axiom lsx_inv_csx: ∀h,g,G,L,T. G ⊢ ⋕⬊*[h, g, T] L → ⦃G, L⦄ ⊢ ⬊*[h, g] T.
+(* Advanced properties ******************************************************)
 
+axiom lpxs_cpye_csx_lsx: ∀h,g,G,L1,U. ⦃G, L1⦄ ⊢ ⬊*[h, g] U →
+                         ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → ∀T.  ⦃G, L2⦄ ⊢ T ▶*[0, ∞] 𝐍⦃U⦄ →
+                         G ⊢ ⋕⬊*[h, g, T] L2.
 (*
-#h #g #G #L1 #T1 #H @(csx_ind_alt … H) -T1
-#T1 #_ #IHT1 @lsx_intro
-#L2 #HL12 #HnT1 elim (lpxs_inv_cpxs_nlleq … HL12 HnT1) -HL12 -HnT1
-#T2 #H1T12 #HnT12 #H2T12 lapply (IHT1 … H1T12 HnT12) -IHT1 -H1T12 -HnT12
-#HT2  
+#h #g #G #L1 #U #H @(csx_ind_alt … H) -U
+#U0 #_ #IHU0 #L0 #HL10 #T #H0 @lsx_intro
+#L2 #HL02 #HnT elim (cpye_total G L2 T 0 (∞))
+#U2 #H2 elim (eq_term_dec U0 U2) #H destruct
+[ -IHU0
+| -HnT /4 width=9 by lpxs_trans, lpxs_cpxs_trans, cpx_cpye_fwd_lpxs/
+]
 *)
+(* Main properties **********************************************************)
+
+lemma csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → G ⊢ ⋕⬊*[h, g, T] L.
+#h #g #G #L #T #HT elim (cpye_total G L T 0 (∞))
+#U #HTU elim HTU
+/4 width=5 by lpxs_cpye_csx_lsx, csx_cpx_trans, cpys_cpx/
+qed.