X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcpx.ma;h=7d20bc47fe9fc848197567877d811a868c2ef74a;hb=b4b5f03ffca4f250a1dc02f277b70e4f33ac8a9b;hp=c7dde663e8dfa0e6fc0ea7d379f44a6461995da9;hpb=632e2b3b8444d363c5f8acb87c3695fc781fe35b;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx.ma index c7dde663e..7d20bc47f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx.ma @@ -18,7 +18,7 @@ include "basic_2/rt_transition/cpg.ma". (* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************) definition cpx (h): relation4 genv lenv term term ≝ - λG,L,T1,T2. ∃c. ⦃G, L⦄ ⊢ T1 ⬈[c, h] T2. + λG,L,T1,T2. ∃c. ⦃G, L⦄ ⊢ T1 ⬈[eq_f, c, h] T2. interpretation "uncounted context-sensitive parallel rt-transition (term)" @@ -26,9 +26,6 @@ interpretation (* Basic properties *********************************************************) -lemma cpx_atom: ∀h,I,G,L. ⦃G, L⦄ ⊢ ⓪{I} ⬈[h] ⓪{I}. -/2 width=2 by cpg_atom, ex_intro/ qed. - (* Basic_2A1: was: cpx_st *) lemma cpx_ess: ∀h,G,L,s. ⦃G, L⦄ ⊢ ⋆s ⬈[h] ⋆(next h s). /2 width=2 by cpg_ess, ex_intro/ qed. @@ -46,17 +43,17 @@ lemma cpx_lref: ∀h,I,G,K,V,T,U,i. ⦃G, K⦄ ⊢ #i ⬈[h] T → qed. lemma cpx_bind: ∀h,p,I,G,L,V1,V2,T1,T2. - ⦃G, L⦄ ⊢ V1 ⬈[h] V2 → ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ⬈[h] T2 → - ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ⬈[h] ⓑ{p,I}V2.T2. + ⦃G, L⦄ ⊢ V1 ⬈[h] V2 → ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ⬈[h] T2 → + ⦃G, L⦄ ⊢ ⓑ{p,I}V1.T1 ⬈[h] ⓑ{p,I}V2.T2. #h #p #I #G #L #V1 #V2 #T1 #T2 * #cV #HV12 * /3 width=2 by cpg_bind, ex_intro/ qed. lemma cpx_flat: ∀h,I,G,L,V1,V2,T1,T2. - ⦃G, L⦄ ⊢ V1 ⬈[h] V2 → ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → - ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ⬈[h] ⓕ{I}V2.T2. -#h #I #G #L #V1 #V2 #T1 #T2 * #cV #HV12 * -/3 width=2 by cpg_flat, ex_intro/ + ⦃G, L⦄ ⊢ V1 ⬈[h] V2 → ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → + ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ⬈[h] ⓕ{I}V2.T2. +#h * #G #L #V1 #V2 #T1 #T2 * #cV #HV12 * +/3 width=5 by cpg_appl, cpg_cast, ex_intro/ qed. lemma cpx_zeta: ∀h,G,L,V,T1,T,T2. ⦃G, L.ⓓV⦄ ⊢ T1 ⬈[h] T → @@ -91,13 +88,15 @@ lemma cpx_theta: ∀h,p,G,L,V1,V,V2,W1,W2,T1,T2. /3 width=4 by cpg_theta, ex_intro/ qed. +(* Basic_2A1: includes: cpx_atom *) lemma cpx_refl: ∀h,G,L. reflexive … (cpx h G L). -/2 width=2 by ex_intro/ qed. +/3 width=2 by cpg_refl, ex_intro/ qed. + +(* Advanced properties ******************************************************) lemma cpx_pair_sn: ∀h,I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 → ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ⬈[h] ②{I}V2.T. -#h #I #G #L #V1 #V2 * -/3 width=2 by cpg_pair_sn, ex_intro/ +#h * /2 width=2 by cpx_flat, cpx_bind/ qed. (* Basic inversion lemmas ***************************************************) @@ -164,23 +163,6 @@ lemma cpx_inv_abst1: ∀h,p,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ⬈[h] U2 /3 width=5 by ex3_2_intro, ex_intro/ qed-. -lemma cpx_inv_flat1: ∀h,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ⬈[h] U2 → - ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 & ⦃G, L⦄ ⊢ U1 ⬈[h] T2 & - U2 = ⓕ{I}V2.T2 - | (⦃G, L⦄ ⊢ U1 ⬈[h] U2 ∧ I = Cast) - | (⦃G, L⦄ ⊢ V1 ⬈[h] U2 ∧ I = Cast) - | ∃∃p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 & ⦃G, L⦄ ⊢ W1 ⬈[h] W2 & - ⦃G, L.ⓛW1⦄ ⊢ T1 ⬈[h] T2 & - U1 = ⓛ{p}W1.T1 & - U2 = ⓓ{p}ⓝW2.V2.T2 & I = Appl - | ∃∃p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ⬈[h] V & ⬆*[1] V ≡ V2 & - ⦃G, L⦄ ⊢ W1 ⬈[h] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ⬈[h] T2 & - U1 = ⓓ{p}W1.T1 & - U2 = ⓓ{p}W2.ⓐV2.T2 & I = Appl. -#h #I #G #L #V1 #U1 #U2 * #c #H elim (cpg_inv_flat1 … H) -H * -/4 width=14 by or5_intro0, or5_intro1, or5_intro2, or5_intro3, or5_intro4, ex7_7_intro, ex6_6_intro, ex3_2_intro, ex_intro, conj/ -qed-. - lemma cpx_inv_appl1: ∀h,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐ V1.U1 ⬈[h] U2 → ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 & ⦃G, L⦄ ⊢ U1 ⬈[h] T2 & U2 = ⓐV2.T2 @@ -203,6 +185,43 @@ lemma cpx_inv_cast1: ∀h,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ⬈[h] U2 → /4 width=5 by or3_intro0, or3_intro1, or3_intro2, ex3_2_intro, ex_intro/ qed-. +(* Advanced inversion lemmas ************************************************) + +lemma cpx_inv_zero1_pair: ∀h,I,G,K,V1,T2. ⦃G, K.ⓑ{I}V1⦄ ⊢ #0 ⬈[h] T2 → + T2 = #0 ∨ + ∃∃V2. ⦃G, K⦄ ⊢ V1 ⬈[h] V2 & ⬆*[1] V2 ≡ T2. +#h #I #G #L #V1 #T2 * #c #H elim (cpg_inv_zero1_pair … H) -H * +/4 width=3 by ex2_intro, ex_intro, or_intror, or_introl/ +qed-. + +lemma cpx_inv_lref1_pair: ∀h,I,G,K,V,T2,i. ⦃G, K.ⓑ{I}V⦄ ⊢ #⫯i ⬈[h] T2 → + T2 = #(⫯i) ∨ + ∃∃T. ⦃G, K⦄ ⊢ #i ⬈[h] T & ⬆*[1] T ≡ T2. +#h #I #G #L #V #T2 #i * #c #H elim (cpg_inv_lref1_pair … H) -H * +/4 width=3 by ex2_intro, ex_intro, or_introl, or_intror/ +qed-. + +lemma cpx_inv_flat1: ∀h,I,G,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ⬈[h] U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 & ⦃G, L⦄ ⊢ U1 ⬈[h] T2 & + U2 = ⓕ{I}V2.T2 + | (⦃G, L⦄ ⊢ U1 ⬈[h] U2 ∧ I = Cast) + | (⦃G, L⦄ ⊢ V1 ⬈[h] U2 ∧ I = Cast) + | ∃∃p,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ⬈[h] V2 & ⦃G, L⦄ ⊢ W1 ⬈[h] W2 & + ⦃G, L.ⓛW1⦄ ⊢ T1 ⬈[h] T2 & + U1 = ⓛ{p}W1.T1 & + U2 = ⓓ{p}ⓝW2.V2.T2 & I = Appl + | ∃∃p,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ⬈[h] V & ⬆*[1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ⬈[h] W2 & ⦃G, L.ⓓW1⦄ ⊢ T1 ⬈[h] T2 & + U1 = ⓓ{p}W1.T1 & + U2 = ⓓ{p}W2.ⓐV2.T2 & I = Appl. +#h * #G #L #V1 #U1 #U2 #H +[ elim (cpx_inv_appl1 … H) -H * + /3 width=14 by or5_intro0, or5_intro3, or5_intro4, ex7_7_intro, ex6_6_intro, ex3_2_intro/ +| elim (cpx_inv_cast1 … H) -H [ * ] + /3 width=14 by or5_intro0, or5_intro1, or5_intro2, ex3_2_intro, conj/ +] +qed-. + (* Basic forward lemmas *****************************************************) lemma cpx_fwd_bind1_minus: ∀h,I,G,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ⬈[h] T → ∀p.