]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma
- name changes in the rediction rules
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / cpx.ma
index 7b99bd8de0098a3536129dee3d27f92dd7d2b430..edfc74f5bc8e8cdd791d3876580639ef97a0c71c 100644 (file)
@@ -21,7 +21,7 @@ include "basic_2/reduction/cpr.ma".
 (* avtivate genv *)
 inductive cpx (h) (g): relation4 genv lenv term term ≝
 | cpx_atom : ∀I,G,L. cpx h g G L (⓪{I}) (⓪{I})
-| cpx_sort : ∀G,L,k,l. deg h g k (l+1) → cpx h g G L (⋆k) (⋆(next h k))
+| cpx_st   : ∀G,L,k,l. deg h g k (l+1) → cpx h g G L (⋆k) (⋆(next h k))
 | cpx_delta: ∀I,G,L,K,V,V2,W2,i.
              ⇩[i] L ≡ K.ⓑ{I}V → cpx h g G K V V2 →
              ⇧[0, i+1] V2 ≡ W2 → cpx h g G L (#i) W2
@@ -33,8 +33,8 @@ inductive cpx (h) (g): relation4 genv lenv term term ≝
              cpx h g G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
 | cpx_zeta : ∀G,L,V,T1,T,T2. cpx h g G (L.ⓓV) T1 T →
              ⇧[0, 1] T2 ≡ T → cpx h g G L (+ⓓV.T1) T2
-| cpx_tau  : ∀G,L,V,T1,T2. cpx h g G L T1 T2 → cpx h g G L (ⓝV.T1) T2
-| cpx_ti   : ∀G,L,V1,V2,T. cpx h g G L V1 V2 → cpx h g G L (ⓝV1.T) V2
+| cpx_eps  : ∀G,L,V,T1,T2. cpx h g G L T1 T2 → cpx h g G L (ⓝV.T1) T2
+| cpx_ct   : ∀G,L,V1,V2,T. cpx h g G L V1 V2 → cpx h g G L (ⓝV1.T) V2
 | cpx_beta : ∀a,G,L,V1,V2,W1,W2,T1,T2.
              cpx h g G L V1 V2 → cpx h g G L W1 W2 → cpx h g G (L.ⓛW1) T1 T2 →
              cpx h g G L (ⓐV1.ⓛ{a}W1.T1) (ⓓ{a}ⓝW2.V2.T2)
@@ -53,12 +53,12 @@ interpretation
 lemma lsubr_cpx_trans: ∀h,g,G. lsub_trans … (cpx h g G) lsubr.
 #h #g #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2
 [ //
-| /2 width=2 by cpx_sort/
+| /2 width=2 by cpx_st/
 | #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
   elim (lsubr_fwd_ldrop2_bind … HL12 … HLK1) -HL12 -HLK1 *
-  /4 width=7 by cpx_delta, cpx_ti/
+  /4 width=7 by cpx_delta, cpx_ct/
 |4,9: /4 width=1 by cpx_bind, cpx_beta, lsubr_bind/
-|5,7,8: /3 width=1 by cpx_flat, cpx_tau, cpx_ti/
+|5,7,8: /3 width=1 by cpx_flat, cpx_eps, cpx_ct/
 |6,10: /4 width=3 by cpx_zeta, cpx_theta, lsubr_bind/
 ]
 qed-.
@@ -70,7 +70,7 @@ qed.
 
 lemma cpr_cpx: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2.
 #h #g #G #L #T1 #T2 #H elim H -L -T1 -T2
-/2 width=7 by cpx_delta, cpx_bind, cpx_flat, cpx_zeta, cpx_tau, cpx_beta, cpx_theta/
+/2 width=7 by cpx_delta, cpx_bind, cpx_flat, cpx_zeta, cpx_eps, cpx_beta, cpx_theta/
 qed.
 
 lemma cpx_pair_sn: ∀h,g,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 →
@@ -94,15 +94,6 @@ lemma cpx_delift: ∀h,g,I,G,K,V,T1,L,d. ⇩[d] L ≡ (K.ⓑ{I}V) →
 ]
 qed-.
 
-lemma cpx_append: ∀h,g,G. l_appendable_sn … (cpx h g G).
-#h #g #G #K #T1 #T2 #H elim H -G -K -T1 -T2
-/2 width=3 by cpx_sort, cpx_bind, cpx_flat, cpx_zeta, cpx_tau, cpx_ti, cpx_beta, cpx_theta/
-#I #G #K #K0 #V1 #V2 #W2 #i #HK0 #_ #HVW2 #IHV12 #L
-lapply (ldrop_fwd_length_lt2 … HK0) #H
-@(cpx_delta … I … (L@@K0) V1 … HVW2) // 
-@(ldrop_O1_append_sn_le … HK0) /2 width=2 by lt_to_le/ (**) (* /3/ does not work *)
-qed.
-
 (* Basic inversion lemmas ***************************************************)
 
 fact cpx_inv_atom1_aux: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ∀J. T1 = ⓪{J} →
@@ -311,20 +302,3 @@ elim (cpx_inv_bind1 … H) -H *
 | #T2 #_ #_ #H destruct
 ]
 qed-.
-
-lemma cpx_fwd_shift1: ∀h,g,G,L1,L,T1,T. ⦃G, L⦄ ⊢ L1 @@ T1 ➡[h, g] T →
-                      ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2.
-#h #g #G #L1 @(lenv_ind_dx … L1) -L1 normalize
-[ #L #T1 #T #HT1
-  @(ex2_2_intro … (⋆)) // (**) (* explicit constructor *)
-| #I #L1 #V1 #IH #L #T1 #X
-  >shift_append_assoc normalize #H
-  elim (cpx_inv_bind1 … H) -H *
-  [ #V0 #T0 #_ #HT10 #H destruct
-    elim (IH … HT10) -IH -HT10 #L2 #T2 #HL12 #H destruct
-    >append_length >HL12 -HL12
-    @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] /2 width=3 by refl, trans_eq/ (**) (* explicit constructor *)
-  | #T #_ #_ #H destruct
-  ]
-]
-qed-.