X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Flprs_cprs.ma;h=958fcab3bc781053975b060c8f5cf558b8dfb4a0;hp=b910b2763c065d77282fd4abecb1d20acc43c315;hb=cac0166656e08399eaaf1a1e19f0ccea28c36d39;hpb=150f931929c8333dbcfff8dbe77fb2e177f44c56 diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lprs_cprs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lprs_cprs.ma index b910b2763..958fcab3b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lprs_cprs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lprs_cprs.ma @@ -12,56 +12,26 @@ (* *) (**************************************************************************) -include "basic_2/computation/cprs_cprs.ma". -include "basic_2/computation/lprs.ma". +include "basic_2/rt_computation/cprs_cprs.ma". +include "basic_2/rt_computation/lprs_cpms.ma". -(* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************) +(* PARALLEL R-COMPUTATION FOR FULL LOCAL ENVIRONMENTS ***********************) (* Advanced properties ******************************************************) -lemma lprs_pair: ∀I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → - ∀V1,V2. ⦃G, L1⦄ ⊢ V1 ➡* V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡* L2.ⓑ{I}V2. -/2 width=1 by TC_lpx_sn_pair/ qed. +(* Basic_2A1: was: lprs_pair2 *) +lemma lprs_pair_dx (h) (G): ∀L1,L2. ⦃G, L1⦄ ⊢ ➡*[h] L2 → + ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡*[h] V2 → + ∀I. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡*[h] L2.ⓑ{I}V2. +/3 width=3 by lprs_pair, lprs_cpms_trans/ qed. -(* Advanced inversion lemmas ************************************************) +(* Properties on context-sensitive parallel r-computation for terms *********) -lemma lprs_inv_pair1: ∀I,G,K1,L2,V1. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡* L2 → - ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡* K2 & ⦃G, K1⦄ ⊢ V1 ➡* V2 & - L2 = K2.ⓑ{I}V2. -/3 width=3 by TC_lpx_sn_inv_pair1, lpr_cprs_trans/ qed-. - -lemma lprs_inv_pair2: ∀I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡* K2.ⓑ{I}V2 → - ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡* K2 & ⦃G, K1⦄ ⊢ V1 ➡* V2 & - L1 = K1.ⓑ{I}V1. -/3 width=3 by TC_lpx_sn_inv_pair2, lpr_cprs_trans/ qed-. - -(* Advanced eliminators *****************************************************) - -lemma lprs_ind_alt: ∀G. ∀R:relation lenv. - R (⋆) (⋆) → ( - ∀I,K1,K2,V1,V2. - ⦃G, K1⦄ ⊢ ➡* K2 → ⦃G, K1⦄ ⊢ V1 ➡* V2 → - R K1 K2 → R (K1.ⓑ{I}V1) (K2.ⓑ{I}V2) - ) → - ∀L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → R L1 L2. -/3 width=4 by TC_lpx_sn_ind, lpr_cprs_trans/ qed-. - -(* Properties on context-sensitive parallel computation for terms ***********) - -lemma lprs_cpr_trans: ∀G. b_c_transitive … (cpr G) (λ_. lprs G). -/3 width=5 by b_c_trans_CTC2, lpr_cprs_trans/ qed-. - -(* Basic_1: was just: pr3_pr3_pr3_t *) -(* Note: alternative proof /3 width=5 by s_r_trans_CTC1, lprs_cpr_trans/ *) -lemma lprs_cprs_trans: ∀G. b_rs_transitive … (cpr G) (λ_. lprs G). -#G @b_c_to_b_rs_trans @b_c_trans_CTC2 -@b_rs_trans_TC1 /2 width=3 by lpr_cprs_trans/ (**) (* full auto too slow *) -qed-. - -lemma lprs_cprs_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 → - ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 → - ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T. -#G #L0 #T0 #T1 #HT01 #L1 #H @(lprs_ind … H) -L1 /2 width=3 by ex2_intro/ +lemma lprs_cprs_conf_dx (h) (G): ∀L0.∀T0,T1:term. ⦃G, L0⦄ ⊢ T0 ➡*[h] T1 → + ∀L1. ⦃G, L0⦄ ⊢ ➡*[h] L1 → + ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡*[h] T & ⦃G, L1⦄ ⊢ T0 ➡*[h] T. +#h #G #L0 #T0 #T1 #HT01 #L1 #H +@(lprs_ind_dx … H) -L1 /2 width=3 by ex2_intro/ #L #L1 #_ #HL1 * #T #HT1 #HT0 -L0 elim (cprs_lpr_conf_dx … HT1 … HL1) -HT1 #T2 #HT2 elim (cprs_lpr_conf_dx … HT0 … HL1) -L #T3 #HT3 @@ -69,112 +39,21 @@ elim (cprs_conf … HT2 … HT3) -T /3 width=5 by cprs_trans, ex2_intro/ qed-. -lemma lprs_cpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → - ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 → - ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T. -/3 width=3 by lprs_cprs_conf_dx, cpr_cprs/ qed-. +lemma lprs_cpr_conf_dx (h) (G): ∀L0. ∀T0,T1:term. ⦃G, L0⦄ ⊢ T0 ➡[h] T1 → + ∀L1. ⦃G, L0⦄ ⊢ ➡*[h] L1 → + ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡*[h] T & ⦃G, L1⦄ ⊢ T0 ➡*[h] T. +/3 width=3 by lprs_cprs_conf_dx, cpm_cpms/ qed-. -(* Note: this can be proved on its own using lprs_ind_dx *) -lemma lprs_cprs_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 → - ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 → - ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T. -#G #L0 #T0 #T1 #HT01 #L1 #HL01 +(* Note: this can be proved on its own using lprs_ind_sn *) +lemma lprs_cprs_conf_sn (h) (G): ∀L0. ∀T0,T1:term. ⦃G, L0⦄ ⊢ T0 ➡*[h] T1 → + ∀L1. ⦃G, L0⦄ ⊢ ➡*[h] L1 → + ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡*[h] T & ⦃G, L1⦄ ⊢ T0 ➡*[h] T. +#h #G #L0 #T0 #T1 #HT01 #L1 #HL01 elim (lprs_cprs_conf_dx … HT01 … HL01) -HT01 -/3 width=3 by lprs_cprs_trans, ex2_intro/ +/3 width=3 by lprs_cpms_trans, ex2_intro/ qed-. -lemma lprs_cpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 → - ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 → - ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T. -/3 width=3 by lprs_cprs_conf_sn, cpr_cprs/ qed-. - -lemma cprs_bind2: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 → - ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡* T2 → - ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2. -/4 width=5 by lprs_cprs_trans, lprs_pair, cprs_bind/ qed. - -(* Inversion lemmas on context-sensitive parallel computation for terms *****) - -(* Basic_1: was: pr3_gen_abst *) -lemma cprs_inv_abst1: ∀a,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* U2 → - ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 & - U2 = ⓛ{a}W2.T2. -#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5 by ex3_2_intro/ -#U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct -elim (cpr_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct -lapply (lprs_cpr_trans … HT02 (L.ⓛV1) ?) -/3 width=5 by lprs_pair, cprs_trans, cprs_strap1, ex3_2_intro/ -qed-. - -lemma cprs_inv_abst: ∀a,G,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 → - ⦃G, L⦄ ⊢ W1 ➡* W2 ∧ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2. -#a #G #L #W1 #W2 #T1 #T2 #H elim (cprs_inv_abst1 … H) -H -#W #T #HW1 #HT1 #H destruct /2 width=1 by conj/ -qed-. - -(* Basic_1: was pr3_gen_abbr *) -lemma cprs_inv_abbr1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡* U2 → ( - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 & - U2 = ⓓ{a}V2.T2 - ) ∨ - ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 & ⬆[0, 1] U2 ≘ T2 & a = true. -#a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/ -#U0 #U2 #_ #HU02 * * -[ #V0 #T0 #HV10 #HT10 #H destruct - elim (cpr_inv_abbr1 … HU02) -HU02 * - [ #V2 #T2 #HV02 #HT02 #H destruct - lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) - /4 width=5 by lprs_pair, cprs_trans, cprs_strap1, ex3_2_intro, or_introl/ - | #T2 #HT02 #HUT2 - lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) -HT02 - /4 width=3 by lprs_pair, cprs_trans, ex3_intro, or_intror/ - ] -| #U1 #HTU1 #HU01 elim (lift_total U2 0 1) - #U #HU2 lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0 - /4 width=3 by cprs_strap1, drop_drop, ex3_intro, or_intror/ -] -qed-. - -(* More advanced properties *************************************************) - -lemma lprs_pair2: ∀I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → - ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡* V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡* L2.ⓑ{I}V2. -/3 width=3 by lprs_pair, lprs_cprs_trans/ qed. - -(* Advanced inversion lemmas ************************************************) - -(* Basic_2A1: uses: scpds_inv_abst1 *) -lemma cpms_inv_abst_sn (n) (h) (G) (L): - ∀p,V1,T1,X2. ⦃G, L⦄ ⊢ ⓛ{p}V1.T1 ➡*[n, h] X2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h] V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 ➡*[n, h] T2 & - X2 = ⓛ{p}V2.T2. -#n #h #G #L #p #V1 #T1 #X2 #H -@(cpms_ind_dx … H) -X2 /2 width=5 by ex3_2_intro/ -#n1 #n2 #X #X2 #_ * #V #T #HV1 #HT1 #H1 #H2 destruct -elim (cpm_inv_abst1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H2 destruct -@ex3_2_intro [1,2,5: // ] -[ /2 width=3 by cprs_step_dx/ -| @(cpms_trans … HT1) -T1 -V2 -n1 - -/3 width=5 by cprs_step_dx, cpms_step_dx, ex3_2_intro/ - - - -#d2 * #X #d1 #Hd21 #Hd1 #H1 #H2 -lapply (da_inv_bind … Hd1) -Hd1 #Hd1 -elim (lstas_inv_bind1 … H1) -H1 #U #HTU1 #H destruct -elim (cprs_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HUT2 #H destruct -/3 width=6 by ex4_2_intro, ex3_2_intro/ -qed-. - -lemma scpds_inv_abbr_abst: ∀h,o,a1,a2,G,L,V1,W2,T1,T2,d. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, o, d] ⓛ{a2}W2.T2 → - ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, o, d] T & ⬆[0, 1] ⓛ{a2}W2.T2 ≘ T & a1 = true. -#h #o #a1 #a2 #G #L #V1 #W2 #T1 #T2 #d2 * #X #d1 #Hd21 #Hd1 #H1 #H2 -lapply (da_inv_bind … Hd1) -Hd1 #Hd1 -elim (lstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct -elim (cprs_inv_abbr1 … H2) -H2 * -[ #V2 #U2 #HV12 #HU12 #H destruct -| /3 width=6 by ex4_2_intro, ex3_intro/ -] -qed-. -*) +lemma lprs_cpr_conf_sn (h) (G): ∀L0. ∀T0,T1:term. ⦃G, L0⦄ ⊢ T0 ➡[h] T1 → + ∀L1. ⦃G, L0⦄ ⊢ ➡*[h] L1 → + ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡*[h] T & ⦃G, L1⦄ ⊢ T0 ➡*[h] T. +/3 width=3 by lprs_cprs_conf_sn, cpm_cpms/ qed-.