X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcpm_drops.ma;h=91544044c3e92b201ccf1619c05b53acf5faf0f1;hb=ca7327c20c6031829fade8bb84a3a1bb66113f54;hp=bc9463c958390c5a16ec878d6d4b3f5bf6d30fcf;hpb=db020b4218272e2e35641ce3bc3b0a9b3afda899;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma index bc9463c95..91544044c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpm_drops.ma @@ -21,58 +21,58 @@ include "basic_2/rt_transition/cpm.ma". (* Basic_1: includes: pr0_lift pr2_lift *) (* Basic_2A1: includes: cpr_lift *) -lemma cpm_lifts_sn: ∀n,h,G. d_liftable2_sn … lifts (λL. cpm h G L n). -#n #h #G #K #T1 #T2 * #c #Hc #HT12 #b #f #L #HLK #U1 #HTU1 +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: ∀n,h,G. d_liftable2_bi … lifts (λL. cpm h G L n). -#n #h #G #K #T1 #T2 * /3 width=11 by cpg_lifts_bi, ex2_intro/ +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-. (* Inversion lemmas with generic slicing for local environments *************) (* Basic_1: includes: pr0_gen_lift pr2_gen_lift *) (* Basic_2A1: includes: cpr_inv_lift1 *) -lemma cpm_inv_lifts_sn: ∀n,h,G. d_deliftable2_sn … lifts (λL. cpm h G L n). -#n #h #G #L #U1 #U2 * #c #Hc #HU12 #b #f #K #HLK #T1 #HTU1 +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: ∀n,h,G. d_deliftable2_bi … lifts (λL. cpm h G L n). -#n #h #G #L #U1 #U2 * /3 width=11 by cpg_inv_lifts_bi, ex2_intro/ +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-. (* Advanced properties ******************************************************) (* Basic_1: includes: pr2_delta1 *) (* Basic_2A1: includes: cpr_delta *) -lemma cpm_delta_drops: ∀n,h,G,L,K,V,V2,W2,i. - ⬇*[i] L ≘ K.ⓓV → ⦃G,K⦄ ⊢ V ➡[n,h] V2 → - ⬆*[↑i] V2 ≘ W2 → ⦃G,L⦄ ⊢ #i ➡[n,h] W2. -#n #h #G #L #K #V #V2 #W2 #i #HLK * +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: ∀n,h,G,L,K,V,V2,W2,i. - ⬇*[i] L ≘ K.ⓛV → ⦃G,K⦄ ⊢ V ➡[n,h] V2 → - ⬆*[↑i] V2 ≘ W2 → ⦃G,L⦄ ⊢ #i ➡[↑n,h] W2. -#n #h #G #L #K #V #V2 #W2 #i #HLK * +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: ∀n,h,I,G,L,T2. ⦃G,L⦄ ⊢ ⓪{I} ➡[n,h] T2 → - ∨∨ T2 = ⓪{I} ∧ n = 0 +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 ➡[n,h] V2 & - ⬆*[↑i] V2 ≘ T2 & I = LRef i - | ∃∃m,K,V,V2,i. ⬇*[i] L ≘ K.ⓛV & ⦃G,K⦄ ⊢ V ➡[m,h] V2 & - ⬆*[↑i] V2 ≘ T2 & I = LRef i & n = ↑m. -#n #h #I #G #L #T2 * #c #Hc #H elim (cpg_inv_atom1_drops … H) -H * + | ∃∃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 * [ #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 @@ -85,13 +85,13 @@ lemma cpm_inv_atom1_drops: ∀n,h,I,G,L,T2. ⦃G,L⦄ ⊢ ⓪{I} ➡[n,h] T2 → ] qed-. -lemma cpm_inv_lref1_drops: ∀n,h,G,L,T2,i. ⦃G,L⦄ ⊢ #i ➡[n,h] T2 → +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 ➡[n,h] V2 & - ⬆*[↑i] V2 ≘ T2 - | ∃∃m,K,V,V2. ⬇*[i] L ≘ K. ⓛV & ⦃G,K⦄ ⊢ V ➡[m,h] V2 & - ⬆*[↑i] V2 ≘ T2 & n = ↑m. -#n #h #G #L #T2 #i * #c #Hc #H elim (cpg_inv_lref1_drops … H) -H * + | ∃∃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/ | #cV #K #V1 #V2 #HLK #HV12 #HVT2 #H destruct @@ -104,10 +104,10 @@ qed-. (* Advanced forward lemmas **************************************************) -fact cpm_fwd_plus_aux (n) (h): ∀G,L,T1,T2. ⦃G,L⦄ ⊢ T1 ➡[n,h] 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 ➡[n1,h] T & ⦃G,L⦄ ⊢ T ➡[n2,h] T2. -#n #h #G #L #T1 #T2 #H @(cpm_ind … H) -G -L -T1 -T2 -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 /2 width=3 by ex2_intro/ @@ -119,19 +119,19 @@ fact cpm_fwd_plus_aux (n) (h): ∀G,L,T1,T2. ⦃G,L⦄ ⊢ T1 ➡[n,h] T2 → ] | #n #G #K #V1 #V2 #W2 #_ #IH #HVW2 #n1 #n2 #H destruct elim IH [|*: // ] -IH #V #HV1 #HV2 - elim (lifts_total V 𝐔❴↑O❵) #W #HVW + elim (lifts_total V 𝐔❨↑O❩) #W #HVW /5 width=11 by cpm_lifts_bi, cpm_delta, drops_refl, drops_drop, ex2_intro/ | #n #G #K #V1 #V2 #W2 #HV12 #IH #HVW2 #x1 #n2 #H elim (plus_inv_S3_sn … H) -H * [ #H1 #H2 destruct -IH /3 width=3 by cpm_ell, ex2_intro/ | #n1 #H1 #H2 destruct -HV12 elim (IH n1) [|*: // ] -IH #V #HV1 #HV2 - elim (lifts_total V 𝐔❴↑O❵) #W #HVW + elim (lifts_total V 𝐔❨↑O❩) #W #HVW /5 width=11 by cpm_lifts_bi, cpm_ell, drops_refl, drops_drop, ex2_intro/ ] | #n #I #G #K #T2 #U2 #i #_ #IH #HTU2 #n1 #n2 #H destruct elim IH [|*: // ] -IH #T #HT1 #HT2 - elim (lifts_total T 𝐔❴↑O❵) #U #HTU + elim (lifts_total T 𝐔❨↑O❩) #U #HTU /5 width=11 by cpm_lifts_bi, cpm_lref, drops_refl, drops_drop, ex2_intro/ | #n #p #I #G #L #V1 #V2 #T1 #T2 #HV12 #_ #_ #IHT #n1 #n2 #H destruct elim IHT [|*: // ] -IHT #T #HT1 #HT2 @@ -165,6 +165,6 @@ fact cpm_fwd_plus_aux (n) (h): ∀G,L,T1,T2. ⦃G,L⦄ ⊢ T1 ➡[n,h] T2 → ] qed-. -lemma cpm_fwd_plus (h) (G) (L): ∀n1,n2,T1,T2. ⦃G,L⦄ ⊢ T1 ➡[n1+n2,h] T2 → - ∃∃T. ⦃G,L⦄ ⊢ T1 ➡[n1,h] T & ⦃G,L⦄ ⊢ T ➡[n2,h] 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-.