From b11e1907f99bea1de50db890d849ba5469d2e0e7 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Thu, 2 Mar 2017 16:00:48 +0000 Subject: [PATCH] - advances on csx - corrections in the web page --- .../basic_2/rt_computation/cpxs_tdeq.ma | 11 ---- .../basic_2/rt_computation/csx_csx.ma | 51 +++++++++++++++++++ .../basic_2/rt_computation/csx_main.ma | 34 ------------- .../basic_2/rt_computation/partial.txt | 2 +- .../{cpx_lfdeq.ma => lfpx_lfdeq.ma} | 18 ++++++- .../basic_2/rt_transition/partial.txt | 4 +- .../lambdadelta/basic_2/syntax/tdeq.ma | 13 +++++ .../lambdadelta/basic_2/syntax/tdeq_tdeq.ma | 10 ++++ .../lambdadelta/basic_2/web/basic_2_src.tbl | 11 ++-- 9 files changed, 100 insertions(+), 54 deletions(-) create mode 100644 matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma delete mode 100644 matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_main.ma rename matita/matita/contribs/lambdadelta/basic_2/rt_transition/{cpx_lfdeq.ma => lfpx_lfdeq.ma} (89%) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma index 027135514..c3e2e2065 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma @@ -15,21 +15,10 @@ include "basic_2/syntax/tdeq_tdeq.ma". include "basic_2/rt_computation/cpxs.ma". include "basic_2/rt_transition/cpx_lfdeq.ma". -include "basic_2/static/lfdeq_fqup.ma". include "basic_2/rt_transition/lfpx_fqup.ma". (* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS ************) -axiom tdeq_dec: ∀h,o,T1,T2. Decidable (tdeq h o T1 T2). - -axiom tdeq_canc_sn: ∀h,o. left_cancellable … (tdeq h o). - -lemma tdeq_cpx_trans: ∀h,o,U1,T1. U1 ≡[h, o] T1 → ∀G,L,T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → - ∃∃U2. ⦃G, L⦄ ⊢ U1 ⬈[h] U2 & U2 ≡[h, o] T2. -#h #o #U1 #T1 #HUT1 #G #L #T2 #HT12 -elim (cpx_tdeq_conf_lexs … o … HT12 … U1 … L … L) /3 width=3 by tdeq_sym, ex2_intro/ -qed-. - lemma tdeq_cpxs_trans: ∀h,o,U1,T1. U1 ≡[h, o] T1 → ∀G,L,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → ∃∃U2. ⦃G, L⦄ ⊢ U1 ⬈*[h] U2 & U2 ≡[h, o] T2. #h #o #U1 #T1 #HUT1 #G #L #T2 #HT12 @(cpxs_ind … HT12) -T2 /2 width=3 by ex2_intro/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma new file mode 100644 index 000000000..e1df7c029 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx.ma @@ -0,0 +1,51 @@ +(**************************************************************************) +(* ___ *) +(* ||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/syntax/tdeq_tdeq.ma". +include "basic_2/rt_transition/lfpx_lfdeq.ma". +include "basic_2/rt_computation/csx.ma". + +(* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********) + +(* Advanced properties ******************************************************) + +lemma csx_tdeq_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ → + ∀T2. T1 ≡[h, o] T2 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄. +#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T #_ #IH #T2 #HT2 +@csx_intro #T1 #HT21 #HnT21 elim (tdeq_cpx_trans … HT2 … HT21) -HT21 +/4 width=5 by tdeq_repl/ +qed-. + +lemma csx_cpx_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ → + ∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄. +#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T1 #HT1 #IHT1 #T2 #HLT12 +elim (tdeq_dec h o T1 T2) /3 width=4 by csx_tdeq_trans/ +qed-. + +(* Basic_1: was just: sn3_cast *) +lemma csx_cast: ∀h,o,G,L,W. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃W⦄ → + ∀T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃ⓝW.T⦄. +#h #o #G #L #W #HW @(csx_ind … HW) -W +#W #HW #IHW #T #HT @(csx_ind … HT) -T +#T #HT #IHT @csx_intro +#X #H1 #H2 elim (cpx_inv_cast1 … H1) -H1 +[ * #W0 #T0 #HLW0 #HLT0 #H destruct + elim (tdeq_false_inv_pair … H2) -H2 + [ -W -T #H elim H -H // + | -HW -IHT /3 width=3 by csx_cpx_trans/ + | -HW -HT -IHW /4 width=3 by csx_cpx_trans, cpx_pair_sn/ + ] +|*: /3 width=3 by csx_cpx_trans/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_main.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_main.ma deleted file mode 100644 index d3556cef5..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_main.ma +++ /dev/null @@ -1,34 +0,0 @@ -(* -lemma csx_tdeq_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ → - ∀T2. T1 ≡[h, o] T2 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄. -#h #o #G #L #T1 #H @(csx_ind … H) -H #T #HT #IH #T2 #HT2 -@csx_intro #T0 #HT20 #HnT20 - -lemma csx_tdeq_trans: ∀h,o,T1,T2. T1 ≡[h, o] T2 → - ∀G,L. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄. -#h #o #T1 #T2 #H elim H -T1 -T2 // -[ #s1 #s2 #d #Hs1 #Hs2 #G #L #H -| #I #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #G #L #H - -lemma csx_cpx_trans: ∀h,o,G,L,T1. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T1⦄ → - ∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T2⦄. -#h #o #G #L #T1 #H @(csx_ind … H) -T1 #T1 #HT1 #IHT1 #T2 #HLT12 -elim (tdeq_dec h o T1 T2) #HT12 /3 width=4 by/ -IHT1 -HLT12 -qed-. - -(* Basic_1: was just: sn3_cast *) -lemma csx_cast: ∀h,o,G,L,W. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃W → - ∀T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃T → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐒⦃ⓝW.T. -#h #o #G #L #W #HW @(csx_ind … HW) -W #W #HW #IHW #T #HT @(csx_ind … HT) -T #T #HT #IHT -@csx_intro #X #H1 #H2 -elim (cpx_inv_cast1 … H1) -H1 -[ * #W0 #T0 #HLW0 #HLT0 #H destruct - elim (eq_false_inv_tpair_sn … H2) -H2 - [ /3 width=3 by csx_cpx_trans/ - | -HLW0 * #H destruct /3 width=1 by/ - ] -|2,3: /3 width=3 by csx_cpx_trans/ -] -qed. - -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/partial.txt b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/partial.txt index fb8f2ebf9..deef34eac 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/partial.txt +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/partial.txt @@ -1,2 +1,2 @@ cpxs.ma -csx.ma csx_cnx.ma +csx.ma csx_cnx.ma csx_csx.ma diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma similarity index 89% rename from matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma rename to matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma index 1a4d722e7..d23144557 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpx_lfdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/relocation/lifts_tdeq.ma". -include "basic_2/static/lfdeq.ma". +include "basic_2/static/lfdeq_fqup.ma". include "basic_2/rt_transition/lfpx.ma". (* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *************) @@ -100,6 +100,22 @@ lemma cpx_tdeq_conf_lexs: ∀h,o,G. R_confluent2_lfxs … (cpx h G) (cdeq h o) ( /4 width=9 by cpx_theta, tdeq_pair, ex2_intro/ (* note: 2 tdeq_pair *) ] qed-. + +lemma cpx_tdeq_conf: ∀h,o,G,L,T0,T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → + ∀T2. T0 ≡[h, o] T2 → + ∃∃T. T1 ≡[h, o] T & ⦃G, L⦄ ⊢ T2 ⬈[h] T. +#h #o #G #L #T0 #T1 #HT01 #T2 #HT02 +elim (cpx_tdeq_conf_lexs … HT01 … HT02 L … L) -HT01 -HT02 +/2 width=3 by lfxs_refl, ex2_intro/ +qed-. + +lemma tdeq_cpx_trans: ∀h,o,G,L,T2,T0. T2 ≡[h, o] T0 → + ∀T1. ⦃G, L⦄ ⊢ T0 ⬈[h] T1 → + ∃∃T. ⦃G, L⦄ ⊢ T2 ⬈[h] T & T ≡[h, o] T1. +#h #o #G #L #T2 #T0 #HT20 #T1 #HT01 +elim (cpx_tdeq_conf … HT01 T2) -HT01 /3 width=3 by tdeq_sym, ex2_intro/ +qed-. + (* lemma cpx_lleq_conf: ∀h,o,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ⬈[h, o] T2 → ∀L1. L2 ≡[T1, 0] L1 → ⦃G, L1⦄ ⊢ T1 ⬈[h, o] T2. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/partial.txt b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/partial.txt index 9ed6e4104..1adc75c38 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/partial.txt +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/partial.txt @@ -1,6 +1,6 @@ cpg.ma cpg_simple.ma cpg_drops.ma cpg_lsubr.ma -cpx.ma cpx_simple.ma cpx_drops.ma cpx_fqus.ma cpx_lsubr.ma cpx_lfxs.ma cpx_lfdeq.ma -lfpx.ma lfpx_length.ma lfpx_drops.ma lfpx_fqup.ma lfpx_frees.ma lfpx_aaa.ma +cpx.ma cpx_simple.ma cpx_drops.ma cpx_fqus.ma cpx_lsubr.ma cpx_lfxs.ma +lfpx.ma lfpx_length.ma lfpx_drops.ma lfpx_fqup.ma lfpx_frees.ma lfpx_lfdeq.ma lfpx_aaa.ma cnx.ma cnx_simple.ma cnx_drops.ma cpm.ma cpm_simple.ma cpm_drops.ma cpm_lsubr.ma cpm_lfxs.ma cpm_cpx.ma cpr.ma cpr_drops.ma diff --git a/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma index 4e5e41b4d..4b077059d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq.ma @@ -168,3 +168,16 @@ lemma tdeq_dec: ∀h,o,T1,T2. Decidable (tdeq h o T1 T2). elim (tdeq_inv_pair … H) -H /2 width=1 by/ ] qed-. + +(* Negated inversion lemmas *************************************************) + +lemma tdeq_false_inv_pair: ∀h,o,I1,I2,V1,V2,T1,T2. + (②{I1}V1.T1 ≡[h, o] ②{I2}V2.T2 → ⊥) → + ∨∨ I1 = I2 → ⊥ + | V1 ≡[h, o] V2 → ⊥ + | (T1 ≡[h, o] T2 → ⊥). +#h #o #I1 #I2 #V1 #V2 #T1 #T2 #H12 +elim (eq_item2_dec I1 I2) /3 width=1 by or3_intro0/ #H destruct +elim (tdeq_dec h o V1 V2) /3 width=1 by or3_intro1/ +elim (tdeq_dec h o T1 T2) /4 width=1 by tdeq_pair, or3_intro2/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq_tdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq_tdeq.ma index 39b06145c..9f47fd88f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq_tdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/syntax/tdeq_tdeq.ma @@ -28,3 +28,13 @@ theorem tdeq_trans: ∀h,o. Transitive … (tdeq h o). elim (tdeq_inv_pair1 … H) -H /3 width=1 by tdeq_pair/ ] qed-. + +theorem tdeq_canc_sn: ∀h,o. left_cancellable … (tdeq h o). +/3 width=3 by tdeq_trans, tdeq_sym/ qed-. + +theorem tdeq_canc_dx: ∀h,o. right_cancellable … (tdeq h o). +/3 width=3 by tdeq_trans, tdeq_sym/ qed-. + +theorem tdeq_repl: ∀h,o,T1,T2. T1 ≡[h, o] T2 → + ∀U1. T1 ≡[h, o] U1 → ∀U2. T2 ≡[h, o] U2 → U1 ≡[h, o] U2. +/3 width=3 by tdeq_canc_sn, tdeq_trans/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl index 4a5a696ac..267559ff5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl +++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl @@ -121,6 +121,7 @@ table { [ "lpxs ( ⦃?,?⦄ ⊢ ➡*[?,?] ? )" "lpxs_drop" + "lpxs_lleq" + "lpxs_aaa" + "lpxs_cpxs" + "lpxs_lpxs" * ] [ "cpxs_tsts" + "cpxs_tsts_vector" + "cpxs_lreq" + "cpxs_lift" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ] *) + [ "csx ( ⦃?,?⦄ ⊢ ⬈[?,?] 𝐒⦃?⦄ )" "csx_cnx" + "csx_csx" * ] [ "cpxs ( ⦃?,?⦄ ⊢ ? ⬈*[?] ? )" * ] } ] @@ -128,7 +129,7 @@ table { ] class "water" [ { "rt-transition" * } { - [ { "parallel qrst-rtransition" * } { + [ { "parallel qrst-transition" * } { (* [ "fpbq ( ⦃?,?,?⦄ ≽[?,?] ⦃?,?,?⦄ )" "fpbq_alt ( ⦃?,?,?⦄ ≽≽[?,?] ⦃?,?,?⦄ )" + "fpbq_aaa" * ] *) [ "fpb ( ⦃?,?,?⦄ ≻[?,?] ⦃?,?,?⦄ )" (* "fpb_lleq" + "fpb_fleq" *) * ] } @@ -148,8 +149,8 @@ table { ] [ { "uncounted context-sensitive rt-transition" * } { [ "cnx ( ⦃?,?⦄ ⊢ ⬈[?,?] 𝐍⦃?⦄ )" "cnx_simple" + "cnx_drops" * ] - [ "lfpx ( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "lfpx_length" + "lfpx_drops" + "lfpx_fqup" + "lfpx_frees" + "lfpx_aaa" * ] - [ "cpx ( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" "cpx_simple" + "cpx_drops" + "cpx_fqus" + "cpx_lsubr" + "cpx_lfxs" + "cpx_lfdeq" * ] + [ "lfpx ( ⦃?,?⦄ ⊢ ⬈[?,?] ? )" "lfpx_length" + "lfpx_drops" + "lfpx_fqup" + "lfpx_frees" + "lfpx_lfdeq" + "lfpx_aaa" * ] + [ "cpx ( ⦃?,?⦄ ⊢ ? ⬈[?] ? )" "cpx_simple" + "cpx_drops" + "cpx_fqus" + "cpx_lsubr" + "cpx_lfxs" * ] } ] [ { "counted context-sensitive rt-transition" * } { @@ -161,7 +162,7 @@ table { class "green" [ { "static typing" * } { [ { "generic reducibility" * } { - [ "lsubc ( ? ⊢ ? ⫃[?] ? )" "lsubc_drop" + "lsubc_drops" + "lsubc_lsubr" + "lsubc_lsuba" * ] + [ "lsubc ( ? ⊢ ? ⫃[?] ? )" "lsubc_drops" + "lsubc_lsubr" + "lsubc_lsuba" * ] [ "gcp_cr ( ⦃?,?,?⦄ ϵ[?] 〚?〛 )" "gcp_aaa" * ] [ "gcp" *] } @@ -182,7 +183,7 @@ table { ] [ { "context-sensitive free variables" * } { [ "lsubf ( ⦃?,?⦄ ⫃𝐅* ⦃?,?⦄ )" "lsubf_frees" * ] - [ "frees ( ? ⊢ 𝐅*⦃?⦄ ≡ ? )" "frees_weight" + "frees_drops" + "frees_fqup" + "frees_frees" * ] + [ "frees ( ? ⊢ 𝐅*⦃?⦄ ≡ ? )" "frees_drops" + "frees_fqup" + "frees_frees" * ] } ] [ { "restricted ref. for local env." * } { -- 2.39.2