]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/etc/cpys/cpys.etc
update in static_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc / cpys / cpys.etc
index a0603f13e3c76e3f3ae4496d6c9ce4e5bb6e1abd..7da756f978ea3f7827363cc7bc865e2b8d08dfc7 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/notation/relations/extpsubststar_4.ma".
-include "basic_2/grammar/genv.ma".
-include "basic_2/grammar/cl_shift.ma".
-include "basic_2/relocation/ldrop_append.ma".
-include "basic_2/relocation/lsuby.ma".
+include "basic_2/notation/relations/psubststar_6.ma".
+include "basic_2/substitution/cpy.ma".
 
 (* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************)
 
-(* avtivate genv *)
-inductive cpys: relation4 genv lenv term term ≝
-| cpys_atom : ∀I,G,L. cpys G L (⓪{I}) (⓪{I})
-| cpys_delta: ∀I,G,L,K,V,V2,W2,i.
-              ⇩[0, i] L ≡ K.ⓑ{I}V → cpys G K V V2 →
-              ⇧[0, i + 1] V2 ≡ W2 → cpys G L (#i) W2
-| cpys_bind : ∀a,I,G,L,V1,V2,T1,T2.
-              cpys G L V1 V2 → cpys G (L.ⓑ{I}V1) T1 T2 →
-              cpys G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2)
-| cpys_flat : ∀I,G,L,V1,V2,T1,T2.
-              cpys G L V1 V2 → cpys G L T1 T2 →
-              cpys G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2)
-.
-
-interpretation
-   "context-sensitive extended multiple substitution (term)"
-   'ExtPSubstStar G L T1 T2 = (cpys G L T1 T2).
+definition cpys: ynat → ynat → relation4 genv lenv term term ≝
+                 λl,m,G. LTC … (cpy l m G).
 
-(* Basic properties *********************************************************)
+interpretation "context-sensitive extended multiple substritution (term)"
+   'PSubstStar G L T1 l m T2 = (cpys l m G L T1 T2).
 
-lemma lsuby_cpys_trans: ∀G. lsub_trans … (cpys G) lsuby.
-#G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2
-[ //
-| #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12
-  elim (lsuby_fwd_ldrop2_pair … HL12 … HLK1) -HL12 -HLK1 *
-  /3 width=7 by cpys_delta/
-| /4 width=1 by lsuby_pair, cpys_bind/
-| /3 width=1 by cpys_flat/
-]
+(* Basic eliminators ********************************************************)
+
+lemma cpys_ind: ∀G,L,T1,l,m. ∀R:predicate term. R T1 →
+                (∀T,T2. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T → ⦃G, L⦄ ⊢ T ▶[l, m] T2 → R T → R T2) →
+                ∀T2. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → R T2.
+#G #L #T1 #l #m #R #HT1 #IHT1 #T2 #HT12
+@(TC_star_ind … HT1 IHT1 … HT12) //
 qed-.
 
-(* Note: this is "∀L. reflexive … (cpys L)" *)
-lemma cpys_refl: ∀G,T,L. ⦃G, L⦄ ⊢ T ▶*× T.
-#G #T elim T -T // * /2 width=1 by cpys_bind, cpys_flat/
-qed.
+lemma cpys_ind_dx: ∀G,L,T2,l,m. ∀R:predicate term. R T2 →
+                   (∀T1,T. ⦃G, L⦄ ⊢ T1 ▶[l, m] T → ⦃G, L⦄ ⊢ T ▶*[l, m] T2 → R T → R T1) →
+                   ∀T1. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → R T1.
+#G #L #T2 #l #m #R #HT2 #IHT2 #T1 #HT12
+@(TC_star_ind_dx … HT2 IHT2 … HT12) //
+qed-.
 
-lemma cpys_pair_sn: ∀I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ▶*× V2 →
-                    ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ▶*× ②{I}V2.T.
-* /2 width=1 by cpys_bind, cpys_flat/
-qed.
+(* Basic properties *********************************************************)
 
-lemma cpys_bind_ext: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ▶*× V2 →
-                     ∀J,T1,T2. ⦃G, L.ⓑ{J}V1⦄ ⊢ T1 ▶*× T2 →
-                     ∀a,I. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*× ⓑ{a,I}V2.T2.
-/4 width=4 by lsuby_cpys_trans, cpys_bind, lsuby_pair/ qed.
-
-lemma cpys_delift: ∀I,G,K,V,T1,L,d. ⇩[0, d] L ≡ (K.ⓑ{I}V) →
-                   ∃∃T2,T.  ⦃G, L⦄ ⊢ T1 ▶*× T2 & ⇧[d, 1] T ≡ T2.
-#I #G #K #V #T1 elim T1 -T1
-[ * /2 width=4 by cpys_atom, lift_sort, lift_gref, ex2_2_intro/
-  #i #L #d elim (lt_or_eq_or_gt i d) #Hid [1,3: /3 width=4 by cpys_atom, lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ ]
-  destruct
-  elim (lift_total V 0 (i+1)) #W #HVW
-  elim (lift_split … HVW i i) /3 width=7 by cpys_delta, ex2_2_intro/
-| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK
-  elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2
-  [ elim (IHU1 (L. ⓑ{I}W1) (d+1)) -IHU1 /3 width=9 by cpys_bind, ldrop_ldrop, lift_bind, ex2_2_intro/
-  | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8 by cpys_flat, lift_flat, ex2_2_intro/
-  ]
-]
+lemma cpy_cpys: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2.
+/2 width=1 by inj/ qed.
+
+lemma cpys_strap1: ∀G,L,T1,T,T2,l,m.
+                   ⦃G, L⦄ ⊢ T1 ▶*[l, m] T → ⦃G, L⦄ ⊢ T ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2.
+normalize /2 width=3 by step/ qed-.
+
+lemma cpys_strap2: ∀G,L,T1,T,T2,l,m.
+                   ⦃G, L⦄ ⊢ T1 ▶[l, m] T → ⦃G, L⦄ ⊢ T ▶*[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2.
+normalize /2 width=3 by TC_strap/ qed-.
+
+lemma lsuby_cpys_trans: ∀G,l,m. lsub_trans … (cpys l m G) (lsuby l m).
+/3 width=5 by lsuby_cpy_trans, LTC_lsub_trans/
 qed-.
 
-lemma cpys_append: ∀G. l_appendable_sn … (cpys G).
-#G #K #T1 #T2 #H elim H -G -K -T1 -T2
-/2 width=3 by cpys_bind, cpys_flat/
-#I #G #K #K0 #V1 #V2 #W2 #i #HK0 #_ #HVW2 #IHV12 #L
-lapply (ldrop_fwd_length_lt2 … HK0) #H
-@(cpys_delta … I … (L@@K0) V1 … HVW2) // 
-@(ldrop_O1_append_sn_le … HK0) /2 width=2 by lt_to_le/ (**) (* /3/ does not work *)
+lemma cpys_refl: ∀G,L,l,m. reflexive … (cpys l m G L).
+/2 width=1 by cpy_cpys/ qed.
+
+lemma cpys_bind: ∀G,L,V1,V2,l,m. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 →
+                 ∀I,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯l, m] T2 →
+                 ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[l, m] ⓑ{a,I}V2.T2.
+#G #L #V1 #V2 #l #m #HV12 @(cpys_ind … HV12) -V2
+[ #I #T1 #T2 #HT12 @(cpys_ind … HT12) -T2 /3 width=5 by cpys_strap1, cpy_bind/
+| /3 width=5 by cpys_strap1, cpy_bind/
+]
 qed.
 
-(* Basic inversion lemmas ***************************************************)
+lemma cpys_flat: ∀G,L,V1,V2,l,m. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 →
+                 ∀T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 →
+                 ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[l, m] ⓕ{I}V2.T2.
+#G #L #V1 #V2 #l #m #HV12 @(cpys_ind … HV12) -V2
+[ #T1 #T2 #HT12 @(cpys_ind … HT12) -T2 /3 width=5 by cpys_strap1, cpy_flat/
+| /3 width=5 by cpys_strap1, cpy_flat/
+qed.
 
-fact cpys_inv_atom1_aux: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ▶*× T2 → ∀J. T1 = ⓪{J} →
-                         T2 = ⓪{J} ∨
-                         ∃∃I,K,V,V2,i. ⇩[O, i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ▶*× V2 &
-                                       ⇧[O, i + 1] V2 ≡ T2 & J = LRef i.
-#G #L #T1 #T2 * -L -T1 -T2
-[ #I #G #L #J #H destruct /2 width=1 by or_introl/
-| #I #G #L #K #V #V2 #T2 #i #HLK #HV2 #HVT2 #J #H destruct /3 width=9 by ex4_5_intro, or_intror/
-| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
-| #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #H destruct
-]
+lemma cpys_weak: ∀G,L,T1,T2,l1,m1. ⦃G, L⦄ ⊢ T1 ▶*[l1, m1] T2 →
+                 ∀l2,m2. l2 ≤ l1 → l1 + m1 ≤ l2 + m2 →
+                 ⦃G, L⦄ ⊢ T1 ▶*[l2, m2] T2.
+#G #L #T1 #T2 #l1 #m1 #H #l1 #l2 #Hl21 #Hlm12 @(cpys_ind … H) -T2
+/3 width=7 by cpys_strap1, cpy_weak/
 qed-.
 
-lemma cpys_inv_atom1: ∀J,G,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ▶*× T2 →
-                      T2 = ⓪{J} ∨
-                      ∃∃I,K,V,V2,i. ⇩[O, i] L ≡ K.ⓑ{I}V & ⦃G, K⦄ ⊢ V ▶*× V2 &
-                                    ⇧[O, i + 1] V2 ≡ T2 & J = LRef i.
-/2 width=3 by cpys_inv_atom1_aux/ qed-.
-
-lemma cpys_inv_sort1: ∀G,L,T2,k. ⦃G, L⦄ ⊢ ⋆k ▶*× T2 → T2 = ⋆k.
-#G #L #T2 #k #H elim (cpys_inv_atom1 … H) -H // *
-#I #K #V #V2 #i #_ #_ #_ #H destruct
+lemma cpys_weak_top: ∀G,L,T1,T2,l,m.
+                     ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[l, |L| - l] T2.
+#G #L #T1 #T2 #l #m #H @(cpys_ind … H) -T2
+/3 width=4 by cpys_strap1, cpy_weak_top/
 qed-.
 
-lemma cpys_inv_lref1: ∀G,L,T2,i. ⦃G, L⦄ ⊢ #i ▶*× T2 →
-                      T2 = #i ∨
-                      ∃∃I,K,V,V2. ⇩[O, i] L ≡ K. ⓑ{I}V & ⦃G, K⦄ ⊢ V ▶*× V2 &
-                                  ⇧[O, i + 1] V2 ≡ T2.
-#G #L #T2 #i #H elim (cpys_inv_atom1 … H) -H /2 width=1 by or_introl/ *
-#I #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=7 by ex3_4_intro, or_intror/
+lemma cpys_weak_full: ∀G,L,T1,T2,l,m.
+                      ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶*[0, |L|] T2.
+#G #L #T1 #T2 #l #m #H @(cpys_ind … H) -T2
+/3 width=5 by cpys_strap1, cpy_weak_full/
 qed-.
 
-lemma cpys_inv_lref1_ge: ∀G,L,T2,i. ⦃G, L⦄ ⊢ #i ▶*× T2 → |L| ≤ i → T2 = #i.
-#G #L #T2 #i #H elim (cpys_inv_lref1 … H) -H // *
-#I #K #V1 #V2 #HLK #_ #_ #HL -V2 lapply (ldrop_fwd_length_lt2 … HLK) -K -I -V1
-#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
+(* Basic forward lemmas *****************************************************)
+
+lemma cpys_fwd_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶*[lt, mt] U2 →
+                   ∀T1,l,m. ⬆[l, m] T1 ≡ U1 →
+                   l ≤ lt → l + m ≤ lt + mt →
+                   ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶*[l+m, lt+mt-(l+m)] U2 & ⬆[l, m] T2 ≡ U2.
+#G #L #U1 #U2 #lt #mt #H #T1 #l #m #HTU1 #Hllt #Hlmlmt @(cpys_ind … H) -U2
+[ /2 width=3 by ex2_intro/
+| -HTU1 #U #U2 #_ #HU2 * #T #HU1 #HTU
+  elim (cpy_fwd_up … HU2 … HTU) -HU2 -HTU /3 width=3 by cpys_strap1, ex2_intro/
+]
 qed-.
 
-lemma cpys_inv_gref1: ∀G,L,T2,p.  ⦃G, L⦄ ⊢ §p ▶*× T2 → T2 = §p.
-#G #L #T2 #p #H elim (cpys_inv_atom1 … H) -H // *
-#I #K #V #V2 #i #_ #_ #_ #H destruct
+lemma cpys_fwd_tw: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 → ♯{T1} ≤ ♯{T2}.
+#G #L #T1 #T2 #l #m #H @(cpys_ind … H) -T2 //
+#T #T2 #_ #HT2 #IHT1 lapply (cpy_fwd_tw … HT2) -HT2
+/2 width=3 by transitive_le/
 qed-.
 
-fact cpys_inv_bind1_aux: ∀G,L,U1,U2. ⦃G, L⦄ ⊢ U1 ▶*× U2 →
-                         ∀a,J,V1,T1. U1 = ⓑ{a,J}V1.T1 →
-                         ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*× V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ T1 ▶*× T2 &
-                                  U2 = ⓑ{a,J}V2.T2.
-#G #L #U1 #U2 * -L -U1 -U2
-[ #I #G #L #b #J #W #U1 #H destruct
-| #I #G #L #K #V #V2 #W2 #i #_ #_ #_ #b #J #W #U1 #H destruct
-| #a #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #b #J #W #U1 #H destruct /2 width=5 by ex3_2_intro/
-| #I #G #L #V1 #V2 #T1 #T2 #_ #_ #b #J #W #U1 #H destruct
-]
+(* Basic inversion lemmas ***************************************************)
+
+(* Note: this can be derived from cpys_inv_atom1 *)
+lemma cpys_inv_sort1: ∀G,L,T2,k,l,m. ⦃G, L⦄ ⊢ ⋆k ▶*[l, m] T2 → T2 = ⋆k.
+#G #L #T2 #k #l #m #H @(cpys_ind … H) -T2 //
+#T #T2 #_ #HT2 #IHT1 destruct
+>(cpy_inv_sort1 … HT2) -HT2 //
 qed-.
 
-lemma cpys_inv_bind1: ∀a,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*× U2 →
-                      ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*× V2 & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*× T2 &
-                               U2 = ⓑ{a,I}V2.T2.
-/2 width=3 by cpys_inv_bind1_aux/ qed-.
-
-lemma cpys_inv_bind1_ext: ∀a,I,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*× U2 → ∀J.
-                          ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*× V2 & ⦃G, L.ⓑ{J}V1⦄ ⊢ T1 ▶*× T2 &
-                                   U2 = ⓑ{a,I}V2.T2.
-#a #I #G #L #V1 #T1 #U2 #H #J elim (cpys_inv_bind1 … H) -H
-#V2 #T2 #HV12 #HT12 #H destruct
-/4 width=5 by lsuby_cpys_trans, lsuby_pair, ex3_2_intro/
+(* Note: this can be derived from cpys_inv_atom1 *)
+lemma cpys_inv_gref1: ∀G,L,T2,p,l,m. ⦃G, L⦄ ⊢ §p ▶*[l, m] T2 → T2 = §p.
+#G #L #T2 #p #l #m #H @(cpys_ind … H) -T2 //
+#T #T2 #_ #HT2 #IHT1 destruct
+>(cpy_inv_gref1 … HT2) -HT2 //
 qed-.
 
-fact cpys_inv_flat1_aux: ∀G,L,U,U2. ⦃G, L⦄ ⊢ U ▶*× U2 →
-                         ∀J,V1,U1. U = ⓕ{J}V1.U1 →
-                         ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*× V2 & ⦃G, L⦄ ⊢ U1 ▶*× T2 &
-                                  U2 = ⓕ{J}V2.T2.
-#G #L #U #U2 * -L -U -U2
-[ #I #G #L #J #W #U1 #H destruct
-| #I #G #L #K #V #V2 #W2 #i #_ #_ #_ #J #W #U1 #H destruct
-| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #J #W #U1 #H destruct
-| #I #G #L #V1 #V2 #T1 #T2 #HV12 #HT12 #J #W #U1 #H destruct /2 width=5 by ex3_2_intro/
+lemma cpys_inv_bind1: ∀a,I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[l, m] U2 →
+                      ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 &
+                               ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯l, m] T2 &
+                               U2 = ⓑ{a,I}V2.T2.
+#a #I #G #L #V1 #T1 #U2 #l #m #H @(cpys_ind … H) -U2
+[ /2 width=5 by ex3_2_intro/
+| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct
+  elim (cpy_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H
+  lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V1) ?) -HT2
+  /3 width=5 by cpys_strap1, lsuby_succ, ex3_2_intro/
 ]
 qed-.
 
-lemma cpys_inv_flat1: ∀I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ▶*× U2 →
-                      ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*× V2 & ⦃G, L⦄ ⊢ U1 ▶*× T2 &
+lemma cpys_inv_flat1: ∀I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[l, m] U2 →
+                      ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[l, m] V2 & ⦃G, L⦄ ⊢ T1 ▶*[l, m] T2 &
                                U2 = ⓕ{I}V2.T2.
-/2 width=3 by cpys_inv_flat1_aux/ qed-.
-
-(* Basic forward lemmas *****************************************************)
+#I #G #L #V1 #T1 #U2 #l #m #H @(cpys_ind … H) -U2
+[ /2 width=5 by ex3_2_intro/
+| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct
+  elim (cpy_inv_flat1 … HU2) -HU2
+  /3 width=5 by cpys_strap1, ex3_2_intro/
+]
+qed-.
 
-lemma cpys_fwd_bind1: ∀a,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*× T → ∀b,J.
-                      ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,J}V1.T1 ▶*× ⓑ{b,J}V2.T2 &
-                               T = ⓑ{a,I}V2.T2.
-#a #I #G #L #V1 #T1 #T #H #b #J elim (cpys_inv_bind1_ext … H J) -H
-#V2 #T2 #HV12 #HT12 #H destruct /3 width=4 by cpys_bind, ex2_2_intro/
+lemma cpys_inv_refl_O2: ∀G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▶*[l, 0] T2 → T1 = T2.
+#G #L #T1 #T2 #l #H @(cpys_ind … H) -T2 //
+#T #T2 #_ #HT2 #IHT1 <(cpy_inv_refl_O2 … HT2) -HT2 //
 qed-.
 
-lemma cpys_fwd_shift1: ∀G,L1,L,T1,T. ⦃G, L⦄ ⊢ L1 @@ T1 ▶*× T →
-                       ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2.
-#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 (cpys_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 trans_eq/ (**) (* explicit constructor *)
-]
+lemma cpys_inv_lift1_eq: ∀G,L,U1,U2,l,m.
+                         ⦃G, L⦄ ⊢ U1 ▶*[l, yinj m] U2 → ∀T1. ⬆[l, m] T1 ≡ U1 → U1 = U2.
+#G #L #U1 #U2 #l #m #H #T1 #HTU1 @(cpys_ind … H) -U2
+/2 width=7 by cpy_inv_lift1_eq/
 qed-.