From d71ad5c0d52dff8bc4b77216fbcb0b65ecd7d391 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Thu, 2 Oct 2014 21:20:21 +0000 Subject: [PATCH 1/1] previous commit was not complete :( --- .../basic_2/computation/fpbg_fpbs.ma | 88 +++++++++++++++++++ .../basic_2/examples/ex_fpbg_refl.ma | 54 ++++++++++++ .../lambdadelta/basic_2/web/basic_2.ldw.xml | 12 +-- 3 files changed, 148 insertions(+), 6 deletions(-) create mode 100644 matita/matita/contribs/lambdadelta/basic_2/computation/fpbg_fpbs.ma create mode 100644 matita/matita/contribs/lambdadelta/basic_2/examples/ex_fpbg_refl.ma diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/fpbg_fpbs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/fpbg_fpbs.ma new file mode 100644 index 000000000..4b8662295 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fpbg_fpbs.ma @@ -0,0 +1,88 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/computation/lpxs_lleq.ma". +include "basic_2/computation/fpbs_lift.ma". +include "basic_2/computation/fpbg_fleq.ma". + +(* "QRST" PROPER PARALLEL COMPUTATION FOR CLOSURES **************************) + +(* Properties on "qrst" parallel reduction on closures **********************) + +lemma fpbg_fpb_trans: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. + ⦃G1, L1, T1⦄ >≡[h, g] ⦃G, L, T⦄ → ⦃G, L, T⦄ ≽[h, g] ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄. +#h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fpb_fpbu … H2) -H2 +/3 width=5 by fpbg_fleq_trans, fpbg_strap1, fpbu_fpbc/ +qed-. + +lemma fpb_fpbg_trans: ∀h,g,G1,G,G2,L1,L,L2,T1,T,T2. + ⦃G1, L1, T1⦄ ≽[h, g] ⦃G, L, T⦄ → ⦃G, L, T⦄ >≡[h, g] ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄. +#h #g #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 elim (fpb_fpbu … H1) -H1 +/3 width=5 by fleq_fpbg_trans, fpbg_strap2, fpbu_fpbc/ +qed-. + +(* Properties on "qrst" parallel compuutation on closures *******************) + +lemma fpbs_fpbg_trans: ∀h,g,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G, L, T⦄ → + ∀G2,L2,T2. ⦃G, L, T⦄ >≡[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄. +#h #g #G1 #G #L1 #L #T1 #T #H @(fpbs_ind … H) -G -L -T /3 width=5 by fpb_fpbg_trans/ +qed-. + +(* Note: this is used in the closure proof *) +lemma fpbg_fpbs_trans: ∀h,g,G,G2,L,L2,T,T2. ⦃G, L, T⦄ ≥[h, g] ⦃G2, L2, T2⦄ → + ∀G1,L1,T1. ⦃G1, L1, T1⦄ >≡[h, g] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄. +#h #g #G #G2 #L #L2 #T #T2 #H @(fpbs_ind_dx … H) -G -L -T /3 width=5 by fpbg_fpb_trans/ +qed-. + +lemma fpbu_fpbs_fpbg: ∀h,g,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G, L, T⦄ → + ∀G2,L2,T2. ⦃G, L, T⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄. +/3 width=5 by fpbg_fpbs_trans, fpbu_fpbg/ qed. + +(* Note: this is used in the closure proof *) +lemma fqup_fpbg: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐+ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄. +#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim (fqup_inv_step_sn … H) -H +/3 width=5 by fqus_fpbs, fpbu_fqu, fpbu_fpbs_fpbg/ +qed. + +lemma cpxs_fpbg: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → + (T1 = T2 → ⊥) → ⦃G, L, T1⦄ >≡[h, g] ⦃G, L, T2⦄. +#h #g #G #L #T1 #T2 #H #H0 elim (cpxs_neq_inv_step_sn … H … H0) -H -H0 +/4 width=5 by cpxs_fpbs, fpbu_cpx, fpbu_fpbs_fpbg/ +qed. + +lemma lstas_fpbg: ∀h,g,G,L,T1,T2,l2. ⦃G, L⦄ ⊢ T1 •*[h, l2] T2 → (T1 = T2 → ⊥) → + ∀l1. l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 → ⦃G, L, T1⦄ >≡[h, g] ⦃G, L, T2⦄. +/3 width=5 by lstas_cpxs, cpxs_fpbg/ qed. + +lemma lpxs_fpbg: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → + (L1 ≡[T, 0] L2 → ⊥) → ⦃G, L1, T⦄ >≡[h, g] ⦃G, L2, T⦄. +#h #g #G #L1 #L2 #T #H #H0 elim (lpxs_nlleq_inv_step_sn … H … H0) -H -H0 +/4 width=5 by fpbu_fpbs_fpbg, fpbu_lpx, lpxs_lleq_fpbs/ +qed. + +lemma fpbs_fpbg: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ ≡[0] ⦃G2, L2, T2⦄ ∨ + ⦃G1, L1, T1⦄ >≡[h, g] ⦃G2, L2, T2⦄. +#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2 +[ /2 width=1 by or_introl/ +| #G #G2 #L #L2 #T #T2 #_ #H2 * #H1 elim (fpb_fpbu … H2) -H2 #H2 + [ /3 width=5 by fleq_trans, or_introl/ + | /5 width=5 by fpbc_fpbg, fleq_fpbc_trans, fpbu_fpbc, or_intror/ + | /3 width=5 by fpbg_fleq_trans, or_intror/ + | /4 width=5 by fpbg_strap1, fpbu_fpbc, or_intror/ + ] +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/examples/ex_fpbg_refl.ma b/matita/matita/contribs/lambdadelta/basic_2/examples/ex_fpbg_refl.ma new file mode 100644 index 000000000..a0f03d62b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/examples/ex_fpbg_refl.ma @@ -0,0 +1,54 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/computation/fpbg_fpbs.ma". + +(* EXAMPLES *****************************************************************) + +(* Reflexivity of proper qrst-computation: the term ApplOmega ***************) + +definition ApplDelta: term → nat → term ≝ λW,k. +ⓛW.ⓐ⋆k.ⓐ#0.#0. + +definition ApplOmega1: term → nat → term ≝ λW,k. ⓐ(ApplDelta W k).(ApplDelta W k). + +definition ApplOmega2: term → nat → term ≝ λW,k. +ⓓⓝW.(ApplDelta W k).ⓐ⋆k.ⓐ#0.#0. + +definition ApplOmega3: term → nat → term ≝ λW,k. ⓐ⋆k.(ApplOmega1 W k). + +(* Basic properties *********************************************************) + +lemma ApplDelta_lift: ∀W1,W2,k,d,e. ⬆[d, e] W1 ≡ W2 → + ⬆[d, e] (ApplDelta W1 k) ≡ (ApplDelta W2 k). +/5 width=1 by lift_flat, lift_bind, lift_lref_lt/ qed. + +lemma cpr_ApplOmega_12: ∀G,L,W,k. ⦃G, L⦄ ⊢ ApplOmega1 W k ➡ ApplOmega2 W k. +/2 width=1 by cpr_beta/ qed. + +lemma cpr_ApplOmega_23: ∀G,L,W,k. ⦃G, L⦄ ⊢ ApplOmega2 W k ➡ ApplOmega3 W k. +#G #L #W1 #k elim (lift_total W1 0 1) #W2 #HW12 +@(cpr_zeta … (ApplOmega3 W2 k)) /4 width=1 by ApplDelta_lift, lift_flat/ +@cpr_flat // @cpr_flat @(cpr_delta … (ApplDelta W1 k) ? 0) +[3,5,8,10: /2 width=2 by ApplDelta_lift/ |4,9: /2 width=1 by cpr_eps/ |*: skip ] +qed. + +lemma cpxs_ApplOmega_13: ∀h,g,G,L,W,k. ⦃G, L⦄ ⊢ ApplOmega1 W k ➡*[h,g] ApplOmega3 W k. +/4 width=3 by cpxs_strap1, cpr_cpx/ qed. + +lemma fqup_ApplOmega_13: ∀G,L,W,k. ⦃G, L, ApplOmega3 W k⦄ ⊐+ ⦃G, L, ApplOmega1 W k⦄. +/2 width=1 by/ qed. + +(* Main properties **********************************************************) + +theorem fpbg_refl: ∀h,g,G,L,W,k. ⦃G, L, ApplOmega1 W k⦄ >≡[h,g] ⦃G, L, ApplOmega1 W k⦄. +/3 width=5 by fpbs_fpbg_trans, fqup_fpbg, cpxs_fpbs/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2.ldw.xml b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2.ldw.xml index e6c828efb..ca4970bcb 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2.ldw.xml +++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2.ldw.xml @@ -44,16 +44,16 @@ for context-sensitive computation on terms. - "Big tree" strong normalization + strong qrst-normalization for simply typed terms. lazy equivalence on local environments - serves as irrelevant step in "big tree" computation on closures + addded as q-step to rst-computation on closures (anniversary milestone). - Parametrized slicing for local environments + Parametrized slicing of local environments comprises both versions of this operation (one from basic_1, the other used in basic_2 till now). @@ -61,10 +61,10 @@ Passive support for global environments. - Reaxiomatized β-reductum as in extended β-reduction + Reaxiomatized β-reductum as in rt-reduction. - Context-sensitive extended strong normalization + Context-sensitive strong rt-normalization for simply typed terms. @@ -74,7 +74,7 @@ Mutual recursive preservation of stratified native validity - for "big tree" computation on closures. + for rst-computation on closures. Confluence for context-free parallel reduction on closures. -- 2.39.2