]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma
milestone update in basic_2, update in ground and static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cpm_drops.ma
index 91544044c3e92b201ccf1619c05b53acf5faf0f1..38c78dea6644e6979362d7ed92d28b6880125796 100644 (file)
@@ -21,13 +21,13 @@ include "basic_2/rt_transition/cpm.ma".
 
 (* Basic_1: includes: pr0_lift pr2_lift *)
 (* Basic_2A1: includes: cpr_lift *)
-lemma cpm_lifts_sn: ∀h,n,G. d_liftable2_sn … lifts (λL. cpm h G L n).
+lemma cpm_lifts_sn (h) (n) (G): d_liftable2_sn … lifts (λL. cpm h G L n).
 #h #n #G #K #T1 #T2 * #c #Hc #HT12 #b #f #L #HLK #U1 #HTU1
 elim (cpg_lifts_sn … HT12 … HLK … HTU1) -K -T1
 /3 width=5 by ex2_intro/
 qed-.
 
-lemma cpm_lifts_bi: ∀h,n,G. d_liftable2_bi … lifts (λL. cpm h G L n).
+lemma cpm_lifts_bi (h) (n) (G): d_liftable2_bi … lifts (λL. cpm h G L n).
 #h #n #G #K #T1 #T2 * /3 width=11 by cpg_lifts_bi, ex2_intro/
 qed-.
 
@@ -35,13 +35,13 @@ qed-.
 
 (* Basic_1: includes: pr0_gen_lift pr2_gen_lift *)
 (* Basic_2A1: includes: cpr_inv_lift1 *)
-lemma cpm_inv_lifts_sn: ∀h,n,G. d_deliftable2_sn … lifts (λL. cpm h G L n).
+lemma cpm_inv_lifts_sn (h) (n) (G): d_deliftable2_sn … lifts (λL. cpm h G L n).
 #h #n #G #L #U1 #U2 * #c #Hc #HU12 #b #f #K #HLK #T1 #HTU1
 elim (cpg_inv_lifts_sn … HU12 … HLK … HTU1) -L -U1
 /3 width=5 by ex2_intro/
 qed-.
 
-lemma cpm_inv_lifts_bi: ∀h,n,G. d_deliftable2_bi … lifts (λL. cpm h G L n).
+lemma cpm_inv_lifts_bi (h) (n) (G): d_deliftable2_bi … lifts (λL. cpm h G L n).
 #h #n #G #L #U1 #U2 * /3 width=11 by cpg_inv_lifts_bi, ex2_intro/
 qed-.
 
@@ -49,33 +49,34 @@ qed-.
 
 (* Basic_1: includes: pr2_delta1 *)
 (* Basic_2A1: includes: cpr_delta *)
-lemma cpm_delta_drops: ∀h,n,G,L,K,V,V2,W2,i.
-                       ⇩[i] L ≘ K.ⓓV → ❪G,K❫ ⊢ V ➡[h,n] V2 →
-                       ⇧[↑i] V2 ≘ W2 → ❪G,L❫ ⊢ #i ➡[h,n] W2.
+lemma cpm_delta_drops (h) (n) (G) (L):
+      ∀K,V,V2,W2,i.
+      ⇩[i] L ≘ K.ⓓV → ❪G,K❫ ⊢ V ➡[h,n] V2 →
+      ⇧[↑i] V2 ≘ W2 → ❪G,L❫ ⊢ #i ➡[h,n] W2.
 #h #n #G #L #K #V #V2 #W2 #i #HLK *
 /3 width=8 by cpg_delta_drops, ex2_intro/
 qed.
 
-lemma cpm_ell_drops: ∀h,n,G,L,K,V,V2,W2,i.
-                     ⇩[i] L ≘ K.ⓛV → ❪G,K❫ ⊢ V ➡[h,n] V2 →
-                     ⇧[↑i] V2 ≘ W2 → ❪G,L❫ ⊢ #i ➡[h,↑n] W2.
+lemma cpm_ell_drops (h) (n) (G) (L):
+      ∀K,V,V2,W2,i.
+      ⇩[i] L ≘ K.ⓛV → ❪G,K❫ ⊢ V ➡[h,n] V2 →
+      ⇧[↑i] V2 ≘ W2 → ❪G,L❫ ⊢ #i ➡[h,↑n] W2.
 #h #n #G #L #K #V #V2 #W2 #i #HLK *
 /3 width=8 by cpg_ell_drops, isrt_succ, ex2_intro/
 qed.
 
 (* Advanced inversion lemmas ************************************************)
 
-lemma cpm_inv_atom1_drops: ∀h,n,I,G,L,T2. ❪G,L❫ ⊢ ⓪[I] ➡[h,n] T2 →
-                           ∨∨ T2 = ⓪[I] ∧ n = 0
-                            | ∃∃s. T2 = ⋆(⫯[h]s) & I = Sort s & n = 1
-                            | ∃∃K,V,V2,i. ⇩[i] L ≘ K.ⓓV & ❪G,K❫ ⊢ V ➡[h,n] V2 &
-                                          ⇧[↑i] V2 ≘ T2 & I = LRef i
-                            | ∃∃m,K,V,V2,i. ⇩[i] L ≘ K.ⓛV & ❪G,K❫ ⊢ V ➡[h,m] V2 &
-                                            ⇧[↑i] V2 ≘ T2 & I = LRef i & n = ↑m.
-#h #n #I #G #L #T2 * #c #Hc #H elim (cpg_inv_atom1_drops … H) -H *
+lemma cpm_inv_atom1_drops (h) (n) (G) (L):
+      ∀I,T2. ❪G,L❫ ⊢ ⓪[I] ➡[h,n] T2 →
+      ∨∨ ∧∧ T2 = ⓪[I] & n = 0
+       | ∃∃s. T2 = ⋆(⫯[h]s) & I = Sort s & n = 1
+       | ∃∃K,V,V2,i. ⇩[i] L ≘ K.ⓓV & ❪G,K❫ ⊢ V ➡[h,n] V2 & ⇧[↑i] V2 ≘ T2 & I = LRef i
+       | ∃∃m,K,V,V2,i. ⇩[i] L ≘ K.ⓛV & ❪G,K❫ ⊢ V ➡[h,m] V2 & ⇧[↑i] V2 ≘ T2 & I = LRef i & n = ↑m.
+#h #n #G #L #I #T2 * #c #Hc #H elim (cpg_inv_atom1_drops … H) -H *
 [ #H1 #H2 destruct lapply (isrt_inv_00 … Hc) -Hc
   /3 width=1 by or4_intro0, conj/
-| #s #H1 #H2 #H3 destruct lapply (isrt_inv_01 … Hc) -Hc
+| #s1 #s2 #H1 #H2 #H3 #H4 destruct lapply (isrt_inv_01 … Hc) -Hc
   /4 width=3 by or4_intro1, ex3_intro, sym_eq/ (**) (* sym_eq *)
 | #cV #i #K #V1 #V2 #HLK #HV12 #HVT2 #H1 #H2 destruct
   /4 width=8 by ex4_4_intro, ex2_intro, or4_intro2/
@@ -85,12 +86,11 @@ lemma cpm_inv_atom1_drops: ∀h,n,I,G,L,T2. ❪G,L❫ ⊢ ⓪[I] ➡[h,n] T2 →
 ]
 qed-.
 
-lemma cpm_inv_lref1_drops: ∀h,n,G,L,T2,i. ❪G,L❫ ⊢ #i ➡[h,n] T2 →
-                           ∨∨ T2 = #i ∧ n = 0
-                            | ∃∃K,V,V2. ⇩[i] L ≘ K.ⓓV & ❪G,K❫ ⊢ V ➡[h,n] V2 &
-                                        ⇧[↑i] V2 ≘ T2
-                            | ∃∃m,K,V,V2. ⇩[i] L ≘ K. ⓛV & ❪G,K❫ ⊢ V ➡[h,m] V2 &
-                                          ⇧[↑i] V2 ≘ T2 & n = ↑m.
+lemma cpm_inv_lref1_drops (h) (n) (G) (L):
+      ∀T2,i. ❪G,L❫ ⊢ #i ➡[h,n] T2 →
+      ∨∨ ∧∧ T2 = #i & n = 0
+       | ∃∃K,V,V2. ⇩[i] L ≘ K.ⓓV & ❪G,K❫ ⊢ V ➡[h,n] V2 & ⇧[↑i] V2 ≘ T2
+       | ∃∃m,K,V,V2. ⇩[i] L ≘ K. ⓛV & ❪G,K❫ ⊢ V ➡[h,m] V2 & ⇧[↑i] V2 ≘ T2 & n = ↑m.
 #h #n #G #L #T2 #i * #c #Hc #H elim (cpg_inv_lref1_drops … H) -H *
 [ #H1 #H2 destruct lapply (isrt_inv_00 … Hc) -Hc
   /3 width=1 by or3_intro0, conj/
@@ -104,9 +104,10 @@ qed-.
 
 (* Advanced forward lemmas **************************************************)
 
-fact cpm_fwd_plus_aux (h) (n): ∀G,L,T1,T2. ❪G,L❫ ⊢ T1 ➡[h,n] T2 →
-                               ∀n1,n2. n1+n2 = n →
-                               ∃∃T. ❪G,L❫ ⊢ T1 ➡[h,n1] T & ❪G,L❫ ⊢ T ➡[h,n2] T2.
+fact cpm_fwd_plus_aux (h) (n) (G) (L):
+     ∀T1,T2. ❪G,L❫ ⊢ T1 ➡[h,n] T2 →
+     ∀n1,n2. n1+n2 = n →
+     ∃∃T. ❪G,L❫ ⊢ T1 ➡[h,n1] T & ❪G,L❫ ⊢ T ➡[h,n2] T2.
 #h #n #G #L #T1 #T2 #H @(cpm_ind … H) -G -L -T1 -T2 -n
 [ #I #G #L #n1 #n2 #H
   elim (plus_inv_O3 … H) -H #H1 #H2 destruct
@@ -165,6 +166,7 @@ fact cpm_fwd_plus_aux (h) (n): ∀G,L,T1,T2. ❪G,L❫ ⊢ T1 ➡[h,n] T2 →
 ]
 qed-.
 
-lemma cpm_fwd_plus (h) (G) (L): ∀n1,n2,T1,T2. ❪G,L❫ ⊢ T1 ➡[h,n1+n2] T2 →
-                                ∃∃T. ❪G,L❫ ⊢ T1 ➡[h,n1] T & ❪G,L❫ ⊢ T ➡[h,n2] T2.
+lemma cpm_fwd_plus (h) (G) (L):
+      ∀n1,n2,T1,T2. ❪G,L❫ ⊢ T1 ➡[h,n1+n2] T2 →
+      ∃∃T. ❪G,L❫ ⊢ T1 ➡[h,n1] T & ❪G,L❫ ⊢ T ➡[h,n2] T2.
 /2 width=3 by cpm_fwd_plus_aux/ qed-.