]> 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 670fb947324e1e92f02a33ce56a57125af4205bc..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_lpxs.ma".
 
 (* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************)
 
-(* Advanced forward lemmas **************************************************)
+(* 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.
 (*
-lemma lsx_fwd_bind_dx_lpxs: ∀h,g,a,I,G,L1,V. ⦃G, L1⦄ ⊢ ⬊*[h, g] V →
-                            ∀L2,T. G ⊢ ⋕⬊*[h, g, ⓑ{a,I}V.T] L2 → ⦃G, L1⦄ ⊢ ➡*[h, g] L2 →
-                            G ⊢ ⋕⬊*[h, g, T] L2.ⓑ{I}V.
-#h #g #a #I #G #L1 #V #H @(csx_ind_alt … H) -V
-#V1 #_ #IHV1 #L2 #T #H @(lsx_ind … H) -L2
-#L2 #HL2 #IHL2 #HL12 @lsx_intro
-#Y #H #HnT elim (lpxs_inv_pair1 … H) -H
-#L3 #V3 #HL23 #HV13 #H destruct
-lapply (lpxs_trans … HL12 … HL23) #HL13
-elim (eq_term_dec V1 V3)
-[ -HV13 -HL2 -IHV1 -HL12 #H destruct
-  @IHL2 -IHL2 // -HL23 -HL13 /3 width=2 by lleq_fwd_bind_O/
-| -IHL2 -HL23 -HnT #HnV13
-  @(IHV1 … HnV13) -IHV1 -HnV13 /2 width=3 by lpxs_cpxs_trans/ -HL12 -HL13 -HV13
-  @(lsx_lpxs_trans) … HL2)
+#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/
+]
 *)
-
-(* Advanced inversion lemmas ************************************************)
-
-
-
 (* Main properties **********************************************************)
 
-axiom csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → G ⊢ ⋕⬊*[h, g, T] L.
-(*
-#h #g #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
-#G1 #L1 #T1 #IH #G2 #L2 * *
-[ #k #HG #HL #HT destruct //
-| #i #HG #HL #HT destruct
-  #H
-| #p #HG #HL #HT destruct //
-| #a #I #V2 #T2 #HG #HL #HT #H destruct
-  elim (csx_fwd_bind … H) -H
-  #HV2 #HT2
-| #I #V2 #T2 #HG #HL #HT #H destruct
-  elim (csx_fwd_flat … H) -H /3 width=1 by lsx_flat/
-*)
+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.