From 765848c4d9a3f5434fae623f3e623d1b73ac76a5 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Sun, 19 Aug 2018 18:14:55 +0200 Subject: [PATCH] severe bug found in parallel zeta MIME-Version: 1.0 Content-Type: text/plain; charset=utf8 Content-Transfer-Encoding: 8bit + final commit: component "examples" corrected + ex_cpr_omega improved: Ω ➡ Ω in exactly three steps (long awaited property) --- .../apps_2/examples/ex_cpr_omega.ma | 78 ++++++++++++++++--- .../apps_2/examples/ex_fpbg_refl.ma | 36 ++++----- .../lambdadelta/apps_2/models/deq_cpr.ma | 6 +- 3 files changed, 87 insertions(+), 33 deletions(-) diff --git a/matita/matita/contribs/lambdadelta/apps_2/examples/ex_cpr_omega.ma b/matita/matita/contribs/lambdadelta/apps_2/examples/ex_cpr_omega.ma index 7ebb0a345..bb4edfd67 100644 --- a/matita/matita/contribs/lambdadelta/apps_2/examples/ex_cpr_omega.ma +++ b/matita/matita/contribs/lambdadelta/apps_2/examples/ex_cpr_omega.ma @@ -16,26 +16,82 @@ include "basic_2/rt_transition/cpr.ma". (* EXAMPLES *****************************************************************) -(* A reduction cycle in two steps: the term Omega ***************************) +(* A reduction cycle in exactly three steps: the term Omega *****************) -definition Delta: term → term ≝ λW. +ⓛW.ⓐ#0.#0. +definition Delta (s): term ≝ +ⓛ⋆s.ⓐ#0.#0. -definition Omega1: term → term ≝ λW. ⓐ(Delta W).(Delta W). +definition Omega1 (s): term ≝ ⓐ(Delta s).(Delta s). -definition Omega2: term → term ≝ λW. +ⓓⓝW.(Delta W).ⓐ#0.#0. +definition Omega2 (s): term ≝ +ⓓⓝ⋆s.(Delta s).ⓐ#0.#0. + +definition Omega3 (s): term ≝ +ⓓⓝ⋆s.(Delta s).(Omega1 s). (* Basic properties *********************************************************) -lemma Delta_lifts: ∀W1,W2,f. ⬆*[f] W1 ≘ W2 → - ⬆*[f] (Delta W1) ≘ (Delta W2). +lemma Delta_lifts (f) (s): ⬆*[f] (Delta s) ≘ (Delta s). /4 width=1 by lifts_lref, lifts_bind, lifts_flat/ qed. +(* Basic inversion properties ***********************************************) + +lemma cpr_inv_Delta1_body_sn (h) (G) (L) (s): + ∀X. ⦃G,L.ⓛ⋆s⦄ ⊢ ⓐ#O.#O ➡[h] X → ⓐ#O.#O = X. +#h #G #L #s #X #H +lapply (cpm_inv_appl1 … H) -H * * +[ #W2 #T2 #HW2 #HT2 #H destruct + elim (cpr_inv_zero1 … HW2) -HW2 + elim (cpr_inv_zero1 … HT2) -HT2 + [ #H1 #H2 destruct // + | * #Y #X1 #X2 #_ #_ #H #_ destruct + | #_ * #Y #X1 #X2 #_ #_ #H destruct + | #_ * #Y #X1 #X2 #_ #_ #H destruct + ] +| #p #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #H #_ destruct +| #p #V2 #V #W1 #W2 #T1 #T2 #_ #_ #_ #_ #H #_ destruct +] +qed-. + +lemma cpr_inv_Delta_sn (h) (G) (L) (s): + ∀X. ⦃G, L⦄ ⊢ Delta s ➡[h] X → Delta s = X. +#h #G #L #s #X #H +elim (cpm_inv_abst1 … H) -H #X1 #X2 #H1 #H2 #H destruct +lapply (cpr_inv_sort1 … H1) -H1 #H destruct +<(cpr_inv_Delta1_body_sn … H2) -X2 // +qed-. + (* Main properties **********************************************************) -theorem cpr_Omega_12: ∀h,G,L,W. ⦃G, L⦄ ⊢ Omega1 W ➡[h] Omega2 W. +theorem cpr_Omega_12 (h) (G) (L) (s): ⦃G, L⦄ ⊢ Omega1 s ➡[h] Omega2 s. /2 width=1 by cpm_beta/ qed. -theorem cpr_Omega_21: ∀h,G,L,W. ⦃G, L⦄ ⊢ Omega2 W ➡[h] Omega1 W. -#h #G #L #W1 elim (lifts_total W1 (𝐔❴1❵)) -/5 width=5 by lifts_flat, cpm_zeta, cpm_eps, cpm_appl, cpm_delta, Delta_lifts/ -qed. +theorem cpr_Omega_23 (h) (G) (L) (s): ⦃G, L⦄ ⊢ Omega2 s ➡[h] Omega3 s. +/5 width=3 by cpm_eps, cpm_appl, cpm_bind, cpm_delta, Delta_lifts/ qed. + +theorem cpr_Omega_31 (h) (G) (L) (s): ⦃G, L⦄ ⊢ Omega3 s ➡[h] Omega1 s. +/4 width=3 by cpm_zeta, Delta_lifts, lifts_flat/ qed. + +(* Main inversion properties ************************************************) + +theorem cpr_inv_Omega1_sn (h) (G) (L) (s): + ∀X. ⦃G, L⦄ ⊢ Omega1 s ➡[h] X → + ∨∨ Omega1 s = X | Omega2 s = X. +#h #G #L #s #X #H elim (cpm_inv_appl1 … H) -H * +[ #W2 #T2 #HW2 #HT2 #H destruct + <(cpr_inv_Delta_sn … HW2) -W2 + <(cpr_inv_Delta_sn … HT2) -T2 + /3 width=1 by refl, or_introl/ +| #p #V2 #W1 #W2 #T1 #T2 #HV #HW #HT whd in ⊢ (??%?→?); #H1 #H2 destruct + <(cpr_inv_Delta_sn … HV) -V2 + >(cpr_inv_sort1 … HW) -W2 + <(cpr_inv_Delta1_body_sn … HT) -T2 // +| #p #V2 #V #W1 #W2 #T1 #T2 #_ #_ #_ #_ whd in ⊢ (??%?→?); #H #_ destruct +] +qed-. + +theorem cpr_Omega_21_false (h) (G) (L) (s): ⦃G, L⦄ ⊢ Omega2 s ➡[h] Omega1 s → ⊥. +#h #G #L #s #H elim (cpm_inv_bind1 … H) -H * +[ #W #T #_ #_ whd in ⊢ (??%?→?); #H destruct +| #X #H #_ #_ #_ + elim (lifts_inv_flat2 … H) -H #X1 #X2 #H1 #_ #_ + @(lifts_inv_lref2_uni_lt … H1) -H1 // +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/apps_2/examples/ex_fpbg_refl.ma b/matita/matita/contribs/lambdadelta/apps_2/examples/ex_fpbg_refl.ma index f15bbd19c..257015d95 100644 --- a/matita/matita/contribs/lambdadelta/apps_2/examples/ex_fpbg_refl.ma +++ b/matita/matita/contribs/lambdadelta/apps_2/examples/ex_fpbg_refl.ma @@ -20,38 +20,38 @@ include "basic_2/rt_computation/fpbg_fpbs.ma". (* Reflexivity of proper rst-computation: the term ApplOmega ****************) -definition ApplDelta: term → nat → term ≝ λW,s. +ⓛW.ⓐ⋆s.ⓐ#0.#0. +definition ApplDelta (s0) (s): term ≝ +ⓛ⋆s0.ⓐ⋆s.ⓐ#0.#0. -definition ApplOmega1: term → nat → term ≝ λW,s. ⓐ(ApplDelta W s).(ApplDelta W s). +definition ApplOmega1 (s0) (s): term ≝ ⓐ(ApplDelta s0 s).(ApplDelta s0 s). -definition ApplOmega2: term → nat → term ≝ λW,s. +ⓓⓝW.(ApplDelta W s).ⓐ⋆s.ⓐ#0.#0. +definition ApplOmega2 (s0) (s): term ≝ +ⓓⓝ⋆s0.(ApplDelta s0 s).ⓐ⋆s.ⓐ#0.#0. -definition ApplOmega3: term → nat → term ≝ λW,s. ⓐ⋆s.(ApplOmega1 W s). +definition ApplOmega3 (s0) (s): term ≝ +ⓓⓝ⋆s0.(ApplDelta s0 s).ⓐ⋆s.(ApplOmega1 s0 s). + +definition ApplOmega4 (s0) (s): term ≝ ⓐ⋆s.(ApplOmega1 s0 s). (* Basic properties *********************************************************) -lemma ApplDelta_lifts (f:rtmap): - ∀W1,W2,s. ⬆*[f] W1 ≘ W2 → - ⬆*[f] (ApplDelta W1 s) ≘ (ApplDelta W2 s). +lemma ApplDelta_lifts (f:rtmap) (s0) (s): + ⬆*[f] (ApplDelta s0 s) ≘ (ApplDelta s0 s). /5 width=1 by lifts_sort, lifts_lref, lifts_bind, lifts_flat/ qed. -lemma cpr_ApplOmega_12 (h): ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega1 W s ➡[h] ApplOmega2 W s. +lemma cpr_ApplOmega_12 (h) (G) (L) (s0) (s): ⦃G, L⦄ ⊢ ApplOmega1 s0 s ➡[h] ApplOmega2 s0 s. /2 width=1 by cpm_beta/ qed. -lemma cpr_ApplOmega_23 (h): ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega2 W s ➡[h] ApplOmega3 W s. -#h #G #L #W1 #s elim (lifts_total W1 (𝐔❴1❵)) #W2 #HW12 -@(cpm_zeta … (ApplOmega3 W2 s)) /4 width=1 by ApplDelta_lifts, lifts_flat/ -@cpm_appl // @cpm_appl @(cpm_delta … (ApplDelta W1 s)) -/2 width=1 by ApplDelta_lifts, cpm_eps/ -qed. +lemma cpr_ApplOmega_23 (h) (G) (L) (s0) (s): ⦃G, L⦄ ⊢ ApplOmega2 s0 s ➡[h] ApplOmega3 s0 s. +/6 width=3 by cpm_eps, cpm_appl, cpm_bind, cpm_delta, ApplDelta_lifts/ qed. + +lemma cpr_ApplOmega_34 (h) (G) (L) (s0) (s): ⦃G, L⦄ ⊢ ApplOmega3 s0 s ➡[h] ApplOmega4 s0 s. +/4 width=3 by cpm_zeta, ApplDelta_lifts, lifts_sort, lifts_flat/ qed. -lemma cpxs_ApplOmega_13 (h): ∀G,L,W,s. ⦃G, L⦄ ⊢ ApplOmega1 W s ⬈*[h] ApplOmega3 W s. -/4 width=4 by cpxs_strap1, cpm_fwd_cpx/ qed. +lemma cpxs_ApplOmega_14 (h) (G) (L) (s0) (s): ⦃G, L⦄ ⊢ ApplOmega1 s0 s ⬈*[h] ApplOmega4 s0 s. +/5 width=4 by cpxs_strap1, cpm_fwd_cpx/ qed. -lemma fqup_ApplOmega_13: ∀G,L,W,s. ⦃G, L, ApplOmega3 W s⦄ ⊐+ ⦃G, L, ApplOmega1 W s⦄. +lemma fqup_ApplOmega_41 (G) (L) (s0) (s): ⦃G,L,ApplOmega4 s0 s⦄ ⊐+ ⦃G,L,ApplOmega1 s0 s⦄. /2 width=1 by/ qed. (* Main properties **********************************************************) -theorem fpbg_refl (h) (o): ∀G,L,W,s. ⦃G, L, ApplOmega1 W s⦄ >[h,o] ⦃G, L, ApplOmega1 W s⦄. +theorem fpbg_refl (h) (o) (G) (L) (s0) (s): ⦃G,L,ApplOmega1 s0 s⦄ >[h,o] ⦃G,L,ApplOmega1 s0 s⦄. /3 width=5 by fpbs_fpbg_trans, fqup_fpbg, cpxs_fpbs/ qed. diff --git a/matita/matita/contribs/lambdadelta/apps_2/models/deq_cpr.ma b/matita/matita/contribs/lambdadelta/apps_2/models/deq_cpr.ma index 32d65f5f5..181acf2e6 100644 --- a/matita/matita/contribs/lambdadelta/apps_2/models/deq_cpr.ma +++ b/matita/matita/contribs/lambdadelta/apps_2/models/deq_cpr.ma @@ -51,13 +51,11 @@ lemma cpr_fwd_deq (h) (M): is_model M → is_extensional M → | @(mq … H1M) [4,5: /3 width=2 by seq_sym, me/ |1,2: skip ] /2 width=1 by/ ] -| #G #L #V #U1 #U2 #T2 #_ #IH #HTU2 #gv #lv #Hlv +| #G #L #V #U1 #T1 #T2 #HTU1 #_ #IH #gv #lv #Hlv @(mq … H1M) [4: /3 width=2 by seq_sym, md/ | skip |3: @(seq_trans … H1M) [2: @mz // | skip ] - /3 width=1 by li_abbr/ - |2: skip ] -L -U1 - /3 width=1 by lifts_SO_fwd_vpush, seq_sym/ + ] /3 width=3 by lifts_SO_fwd_vpush, seq_sym/ | #G #L #V #T1 #T2 #_ #IH #gv #lv #Hlv @(seq_trans … H1M) [2: @(me … H1M) | skip ] /2 width=1 by/ -- 2.39.2