]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lsubr.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / csx_lsubr.ma
index 694ce6b3c2b9c9d770a5cae80fcca62f40de1f1a..5b180f3da91a1f73b0d6154d41cc8a895cd00572 100644 (file)
 include "basic_2/rt_transition/cpx_lsubr.ma".
 include "basic_2/rt_computation/csx_csx.ma".
 
-(* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********)
+(* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************)
 
 (* Advanced properties ******************************************************)
 
-fact csx_appl_beta_aux: ∀h,o,p,G,L,U1. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃U1⦄ →
-                        ∀V,W,T1. U1 = ⓓ{p}ⓝW.V.T1 → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓐV.ⓛ{p}W.T1⦄.
-#h #o #p #G #L #X #H @(csx_ind … H) -X
+fact csx_appl_beta_aux (h) (G):
+     ∀p,L,U1. ❪G,L❫ ⊢ ⬈*[h] 𝐒❪U1❫ →
+     ∀V,W,T1. U1 = ⓓ[p]ⓝW.V.T1 → ❪G,L❫ ⊢ ⬈*[h] 𝐒❪ⓐV.ⓛ[p]W.T1❫.
+#h #G #p #L #X #H @(csx_ind … H) -X
 #X #HT1 #IHT1 #V #W #T1 #H1 destruct
 @csx_intro #X #H1 #H2
 elim (cpx_inv_appl1 … H1) -H1 *
@@ -30,9 +31,9 @@ elim (cpx_inv_appl1 … H1) -H1 *
   @IHT1 -IHT1 [4: // | skip ]
   [ lapply (lsubr_cpx_trans … HLT0 (L.ⓓⓝW.V) ?) -HLT0 -H2
     /3 width=1 by cpx_bind, cpx_flat, lsubr_beta/
-  | #H elim (tdeq_inv_pair … H) -H
-    #_ #H elim (tdeq_inv_pair … H) -H
-    #_ /4 width=1 by tdeq_pair/
+  | #H elim (teqx_inv_pair … H) -H
+    #_ #H elim (teqx_inv_pair … H) -H
+    #_ /4 width=1 by teqx_pair/
   ]
 | -IHT1 -H2 #q #V0 #W0 #W2 #T0 #T2 #HLV0 #HLW02 #HLT02 #H1 #H3 destruct
   lapply (lsubr_cpx_trans … HLT02 (L.ⓓⓝW0.V) ?) -HLT02
@@ -42,23 +43,37 @@ elim (cpx_inv_appl1 … H1) -H1 *
 qed-.
 
 (* Basic_1: was just: sn3_beta *)
-lemma csx_appl_beta: ∀h,o,p,G,L,V,W,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓓ{p}ⓝW.V.T⦄ → ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓐV.ⓛ{p}W.T⦄.
+lemma csx_appl_beta (h) (G):
+      ∀p,L,V,W,T. ❪G,L❫ ⊢ ⬈*[h] 𝐒❪ⓓ[p]ⓝW.V.T❫ → ❪G,L❫ ⊢ ⬈*[h] 𝐒❪ⓐV.ⓛ[p]W.T❫.
 /2 width=3 by csx_appl_beta_aux/ qed.
 
 (* Advanced forward lemmas **************************************************)
 
-fact csx_fwd_bind_dx_unit_aux: ∀h,o,G,L,U. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃U⦄ →
-                               ∀p,I,J,V,T. U = ⓑ{p,I}V.T → ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄.
-#h #o #G #L #U #H elim H -H #U0 #_ #IH #p #I #J #V #T #H destruct
+fact csx_fwd_bind_dx_unit_aux (h) (G):
+     ∀L,U. ❪G,L❫ ⊢ ⬈*[h] 𝐒❪U❫ →
+     ∀p,I,J,V,T. U = ⓑ[p,I]V.T → ❪G,L.ⓤ[J]❫ ⊢ ⬈*[h] 𝐒❪T❫.
+#h #G #L #U #H elim H -H #U0 #_ #IH #p #I #J #V #T #H destruct
 @csx_intro #T2 #HLT2 #HT2
-@(IH (ⓑ{p,I}V.T2)) -IH /2 width=4 by cpx_bind_unit/ -HLT2
-#H elim (tdeq_inv_pair … H) -H /2 width=1 by/
+@(IH (ⓑ[p, I]V.T2)) -IH /2 width=4 by cpx_bind_unit/ -HLT2
+#H elim (teqx_inv_pair … H) -H /2 width=1 by/
 qed-.
 
-lemma csx_fwd_bind_dx_unit: ∀h,o,p,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓑ{p,I}V.T⦄ →
-                            ∀J. ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄.
+lemma csx_fwd_bind_dx_unit (h) (G):
+      ∀p,I,L,V,T. ❪G,L❫ ⊢ ⬈*[h] 𝐒❪ⓑ[p,I]V.T❫ →
+      ∀J. ❪G,L.ⓤ[J]❫ ⊢ ⬈*[h] 𝐒❪T❫.
 /2 width=6 by csx_fwd_bind_dx_unit_aux/ qed-.
 
-lemma csx_fwd_bind_unit: ∀h,o,p,I,G,L,V,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃ⓑ{p,I}V.T⦄ →
-                         ∀J. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃V⦄ ∧ ⦃G, L.ⓤ{J}⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄.
+lemma csx_fwd_bind_unit (h) (G):
+      ∀p,I,L,V,T. ❪G,L❫ ⊢ ⬈*[h] 𝐒❪ⓑ[p,I]V.T❫ →
+      ∀J. ❪G,L❫ ⊢ ⬈*[h] 𝐒❪V❫ ∧ ❪G,L.ⓤ[J]❫ ⊢ ⬈*[h] 𝐒❪T❫.
 /3 width=4 by csx_fwd_pair_sn, csx_fwd_bind_dx_unit, conj/ qed-.
+
+(* Properties with restricted refinement for local environments *************)
+
+lemma csx_lsubr_conf (h) (G):
+      ∀L1,T. ❪G,L1❫ ⊢ ⬈*[h] 𝐒❪T❫ → ∀L2. L1 ⫃ L2 → ❪G,L2❫ ⊢ ⬈*[h] 𝐒❪T❫.
+#h #G #L1 #T #H
+@(csx_ind … H) -T #T1 #_ #IH #L2 #HL12
+@csx_intro #T2 #HT12 #HnT12
+/3 width=3 by lsubr_cpx_trans/
+qed-.