]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma
- improved definition of lsx allows more invariants
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lpxs_lleq.ma
index 4609675a72913d91b2250d59abe3e9eeb293a04b..2ca1e193d79dd19b9cca88a98eaca1ec22eb541a 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/substitution/lleq_alt.ma".
+include "basic_2/reduction/lpx_lleq.ma".
+include "basic_2/computation/cpxs_cpys.ma".
 include "basic_2/computation/lpxs_ldrop.ma".
 include "basic_2/computation/lpxs_cpxs.ma".
 
 (* SN EXTENDED PARALLEL COMPUTATION FOR LOCAL ENVIRONMENTS ******************)
 
-(* Advanced properties ******************************************************)
-
-axiom lleq_lpxs_trans_nlleq: ∀h,g,G,L1s,L1d,T,d. L1s ⋕[T, d] L1d →
-                             ∀L2d. ⦃G, L1d⦄ ⊢ ➡*[h, g] L2d → (L1d ⋕[T, d] L2d → ⊥) →
-                             ∃∃L2s. ⦃G, L1s⦄ ⊢ ➡*[h, g] L2s & L2s ⋕[T, d] L2d & L1s ⋕[T, d] L2s → ⊥.
-
-(* Advanced inversion lemmas ************************************************)
-
-axiom lpxs_inv_cpxs_nlleq: ∀h,g,G,L1,L2,T1. ⦃G, L1⦄ ⊢ ➡*[h,g] L2 → (L1 ⋕[T1, 0] L2 → ⊥) →
-                           ∃∃T2. ⦃G, L1⦄ ⊢ T1 ➡*[h, g] T2 & T1 = T2 → ⊥ & ⦃G, L2⦄ ⊢ T1 ➡[h, g] T2.
-
 (* Properties on lazy equivalence for local environments ********************)
 
+lemma lleq_lpxs_trans: ∀h,g,G,L2,K2. ⦃G, L2⦄ ⊢ ➡*[h, g] K2 →
+                       ∀L1,T,d. L1 ⋕[T, d] L2 →
+                       ∃∃K1. ⦃G, L1⦄ ⊢ ➡*[h, g] K1 & K1 ⋕[T, d] K2.
+#h #g #G #L2 #K2 #H @(lpxs_ind … H) -K2 /2 width=3 by ex2_intro/
+#K #K2 #_ #HK2 #IH #L1 #T #d #HT elim (IH … HT) -L2
+#L #HL1 #HT elim (lleq_lpx_trans … HK2 … HT) -K
+/3 width=3 by lpxs_strap1, ex2_intro/
+qed-.
+
 lemma lpxs_lleq_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ →
                            ∀K1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 → K1 ⋕[T1, 0] L1 →
                            ∃∃K2. ⦃G1, K1, T1⦄ ⊃ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡*[h, g] L2 & K2 ⋕[T2, 0] L2.
@@ -93,3 +92,37 @@ elim (fqus_inv_gen … H) -H
 | * #HG #HL #HT destruct /2 width=4 by ex3_intro/
 ]
 qed-.
+
+fact lsuby_lpxs_trans_lleq_aux: ∀h,g,G,L1,L0,d,e. L1 ⊑×[d, e] L0 → e = ∞ →
+                                ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
+                                ∃∃L. L ⊑×[d, e] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
+                                     (∀T. |L1| = |L0| → |L| = |L2| → L0 ⋕[T, d] L2 ↔ L1 ⋕[T, d] L).
+#h #g #G #L1 #L0 #d #e #H elim H -L1 -L0 -d -e
+[ #L1 #d #e #_ #L2 #H >(lpxs_inv_atom1 … H) -H
+  /3 width=5 by ex3_intro, conj/
+| #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #He destruct
+| #I1 #I0 #L1 #L0 #V1 #e #HL10 #IHL10 #He #Y #H
+  elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
+  lapply (ysucc_inv_Y_dx … He) -He #He
+  elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
+  @(ex3_intro … (L.ⓑ{I1}V2)) /3 width=3 by lpxs_pair, lsuby_cpxs_trans, lsuby_pair/
+  #T #H1 #H2 lapply (injective_plus_l … H1) lapply (injective_plus_l … H2) -H1 -H2
+  #H1 #H2 elim (IH T) // #HL0dx #HL0sn
+  @conj #H @(lleq_lsuby_repl … H) -H normalize
+  /3 width=1 by lsuby_sym, lsuby_pair_O_Y/
+| #I1 #I0 #L1 #L0 #V1 #V0 #d #e #HL10 #IHL10 #He #Y #H
+  elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct
+  elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH
+  @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpxs_pair, lsuby_succ/
+  #T #H1 #H2 lapply (injective_plus_l … H1) lapply (injective_plus_l … H2) -H1 -H2
+  #H1 #H2 elim (IH T) // #HL0dx #HL0sn
+  @conj #H @(lleq_lsuby_repl … H) -H
+  /3 width=1 by lsuby_sym, lsuby_succ/ normalize //
+]
+qed-.
+
+lemma lsuby_lpxs_trans_lleq: ∀h,g,G,L1,L0,d. L1 ⊑×[d, ∞] L0 →
+                             ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 →
+                             ∃∃L. L ⊑×[d, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L &
+                                  (∀T. |L1| = |L0| → |L| = |L2| → L0 ⋕[T, d] L2 ↔ L1 ⋕[T, d] L).
+/2 width=1 by lsuby_lpxs_trans_lleq_aux/ qed-.