From: Ferruccio Guidi Date: Fri, 3 Nov 2017 18:17:56 +0000 (+0000) Subject: - exclusion binder in local environments allows to complete lfsx_lfsx ! X-Git-Tag: make_still_working~415 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=a09b60bd574adf1a7d3e423023009cb20c79d449 - exclusion binder in local environments allows to complete lfsx_lfsx ! - advances on csx ... --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfdeq.ma new file mode 100644 index 000000000..67e63613a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfdeq.ma @@ -0,0 +1,35 @@ +(**************************************************************************) +(* ___ *) +(* ||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/rt_transition/cpx_lfdeq.ma". +include "basic_2/rt_computation/csx_csx.ma". + +(* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********) + +(* Properties with degree-based equivalence for local environments **********) + +(* Basic_2A1: uses: csx_lleq_conf *) +lemma csx_lfdeq_conf: ∀h,o,G,L1,T. ⦃G, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → + ∀L2. L1 ≡[h, o, T] L2 → ⦃G, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. +#h #o #G #L1 #T #H +@(csx_ind … H) -T #T1 #_ #IH #L2 #HL12 +@csx_intro #T2 #HT12 #HnT12 +elim (lfdeq_cpx_trans … HL12 … HT12) -HT12 +/5 width=4 by cpx_lfdeq_conf_sn, csx_tdeq_trans, tdeq_trans/ +qed-. + +(* Basic_2A1: uses: csx_lleq_conf *) +lemma csx_lfdeq_trans: ∀h,o,L1,L2,T. L1 ≡[h, o, T] L2 → + ∀G. ⦃G, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → ⦃G, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. +/3 width=3 by csx_lfdeq_conf, lfdeq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfpxs.ma new file mode 100644 index 000000000..bc21143cd --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lfpxs.ma @@ -0,0 +1,26 @@ +(**************************************************************************) +(* ___ *) +(* ||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/rt_computation/csx_lfpx.ma". +include "basic_2/rt_computation/lfpxs_fqup.ma". + +(* STRONGLY NORMALIZING TERMS FOR UNCOUNTED PARALLEL RT-TRANSITION **********) + +(* Properties with uncounted parallel rt-computation on referred entries ****) + +(* Basic_2A1: uses: csx_lpxs_conf *) +lemma csx_lfpxs_conf: ∀h,o,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈*[h, T] L2 → + ⦃G, L1⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → ⦃G, L2⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄. +#h #o #G #L1 #L2 #T #H @(lfpxs_ind_sn … H) -L2 /3 by lfpxs_step_dx, csx_lfpx_conf/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lleq.ma deleted file mode 100644 index 71009c473..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lleq.ma +++ /dev/null @@ -1,30 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||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/reduction/cpx_lleq.ma". -include "basic_2/computation/csx.ma". - -(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************) - -(* Properties on lazy equivalence for local environments ********************) - -lemma csx_lleq_conf: ∀h,o,G,L1,T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T → - ∀L2. L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T. -#h #o #G #L1 #T #H @(csx_ind … H) -T -/4 width=6 by csx_intro, cpx_lleq_conf_dx, lleq_cpx_trans/ -qed-. - -lemma csx_lleq_trans: ∀h,o,G,L1,L2,T. - L1 ≡[T, 0] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, o] T → ⦃G, L1⦄ ⊢ ⬊*[h, o] T. -/3 width=3 by csx_lleq_conf, lleq_sym/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lpxs.ma deleted file mode 100644 index 9d06a4142..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_lpxs.ma +++ /dev/null @@ -1,25 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||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/csx_lpx.ma". -include "basic_2/computation/lpxs.ma". - -(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************) - -(* Properties on sn extended parallel computation for local environments ****) - -lemma csx_lpxs_conf: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ➡*[h, o] L2 → - ∀T. ⦃G, L1⦄ ⊢ ⬊*[h, o] T → ⦃G, L2⦄ ⊢ ⬊*[h, o] T. -#h #o #G #L1 #L2 #H @(lpxs_ind … H) -L2 /3 by lpxs_strap1, csx_lpx_conf/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_lfsx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_lfsx.ma index 3d547a49c..24dcd4c1f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_lfsx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_lfsx.ma @@ -43,7 +43,7 @@ lemma lfsx_fwd_pair_sn: ∀h,o,I,G,L,V,T. G ⊢ ⬈*[h, o, ②{I}V.T] 𝐒⦃L G ⊢ ⬈*[h, o, V] 𝐒⦃L⦄. #h #o #I #G #L #V #T #H @(lfsx_ind … H) -L #L1 #_ #IHL1 @lfsx_intro -#L2 #H #HnL12 elim (lfpx_pair_sn_split … o I … T H) -H +#L2 #H #HnL12 elim (lfpx_pair_sn_split … H o I T) -H /6 width=3 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_pair_sn/ qed-. @@ -52,13 +52,28 @@ lemma lfsx_fwd_flat_dx: ∀h,o,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓕ{I}V.T] 𝐒⦃L G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄. #h #o #I #G #L #V #T #H @(lfsx_ind … H) -L #L1 #_ #IHL1 @lfsx_intro -#L2 #H #HnL12 elim (lfpx_flat_dx_split … o I … V … H) -H +#L2 #H #HnL12 elim (lfpx_flat_dx_split … H o I V) -H /6 width=3 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_flat_dx/ qed-. +(* Basic_2A1: uses: lsx_fwd_bind_dx *) +(* Note: the exclusion binder (ⓧ) makes this more elegant and much simpler *) +lemma lfsx_fwd_bind_dx: ∀h,o,p,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓑ{p,I}V.T] 𝐒⦃L⦄ → + G ⊢ ⬈*[h, o, T] 𝐒⦃L.ⓧ⦄. +#h #o #p #I #G #L #V #T #H @(lfsx_ind … H) -L +#L1 #_ #IH @lfsx_intro +#L2 #H #HT elim (lfpx_bind_dx_split_void … H o p I V) -H +/6 width=5 by lfsx_lfdeq_trans, lfdeq_trans, lfdeq_fwd_bind_dx_void/ +qed-. + (* Advanced inversion lemmas ************************************************) (* Basic_2A1: uses: lsx_inv_flat *) lemma lfsx_inv_flat: ∀h,o,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓕ{I}V.T] 𝐒⦃L⦄ → G ⊢ ⬈*[h, o, V] 𝐒⦃L⦄ ∧ G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄. /3 width=3 by lfsx_fwd_pair_sn, lfsx_fwd_flat_dx, conj/ qed-. + +(* Basic_2A1: uses: lsx_inv_bind *) +lemma lfsx_inv_bind: ∀h,o,p,I,G,L,V,T. G ⊢ ⬈*[h, o, ⓑ{p,I}V.T] 𝐒⦃L⦄ → + G ⊢ ⬈*[h, o, V] 𝐒⦃L⦄ ∧ G ⊢ ⬈*[h, o, T] 𝐒⦃L.ⓧ⦄. +/3 width=4 by lfsx_fwd_pair_sn, lfsx_fwd_bind_dx, conj/ 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 1e49546fd..598bfebb4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/partial.txt +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/partial.txt @@ -1,5 +1,5 @@ cpxs.ma cpxs_tdeq.ma cpxs_theq.ma cpxs_theq_vector.ma cpxs_drops.ma cpxs_fqus.ma cpxs_lsubr.ma cpxs_lfdeq.ma cpxs_aaa.ma cpxs_lfpx.ma cpxs_cnx.ma cpxs_cpxs.ma lfpxs.ma lfpxs_length.ma lfpxs_drops.ma lfpxs_fqup.ma lfpxs_lfdeq.ma lfpxs_aaa.ma lfpxs_cpxs.ma lfpxs_lfpxs.ma -csx.ma csx_simple.ma csx_simple_theq.ma csx_drops.ma csx_lsubr.ma csx_aaa.ma csx_gcp.ma csx_gcr.ma csx_lfpx.ma csx_cnx.ma csx_cpxs.ma csx_csx.ma +csx.ma csx_simple.ma csx_simple_theq.ma csx_drops.ma csx_lsubr.ma csx_lfdeq.ma csx_aaa.ma csx_gcp.ma csx_gcr.ma csx_lfpx.ma csx_cnx.ma csx_cpxs.ma csx_lfpxs.ma csx_csx.ma csx_vector.ma csx_cnx_vector.ma csx_csx_vector.ma lfsx.ma lfsx_fqup.ma lfsx_lfpxs.ma lfsx_lfsx.ma diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma index 414bfddd4..3e7750845 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/lfpx_lfdeq.ma @@ -22,18 +22,22 @@ include "basic_2/rt_transition/lfpx.ma". (**) (* should be in lfpx_frees.ma *) (* Properties with degree-based equivalence for local environments **********) -lemma lfpx_pair_sn_split: ∀h,o,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → +lemma lfpx_pair_sn_split: ∀h,G,L1,L2,V. ⦃G, L1⦄ ⊢ ⬈[h, V] L2 → ∀o,I,T. ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ②{I}V.T] L & L ≡[h, o, V] L2. /3 width=5 by lfpx_frees_conf, lfxs_pair_sn_split/ qed-. -lemma lfpx_flat_dx_split: ∀h,o,I,G,L1,L2,V,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → +lemma lfpx_flat_dx_split: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈[h, T] L2 → ∀o,I,V. ∃∃L. ⦃G, L1⦄ ⊢ ⬈[h, ⓕ{I}V.T] L & L ≡[h, o, T] L2. /3 width=5 by lfpx_frees_conf, lfxs_flat_dx_split/ qed-. -lemma lfpx_bind_dx_split: ∀h,o,p,I,G,L1,L2,V1,T. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L2 → +lemma lfpx_bind_dx_split: ∀h,I,G,L1,L2,V1,T. ⦃G, L1.ⓑ{I}V1⦄ ⊢ ⬈[h, T] L2 → ∀o,p. ∃∃L,V. ⦃G, L1⦄ ⊢ ⬈[h, ⓑ{p,I}V1.T] L & L.ⓑ{I}V ≡[h, o, T] L2 & ⦃G, L1⦄ ⊢ V1 ⬈[h] V. /3 width=5 by lfpx_frees_conf, lfxs_bind_dx_split/ qed-. +lemma lfpx_bind_dx_split_void: ∀h,G,K1,L2,T. ⦃G, K1.ⓧ⦄ ⊢ ⬈[h, T] L2 → ∀o,p,I,V. + ∃∃K2. ⦃G, K1⦄ ⊢ ⬈[h, ⓑ{p,I}V.T] K2 & K2.ⓧ ≡[h, o, T] L2. +/3 width=5 by lfpx_frees_conf, lfxs_bind_dx_split_void/ qed-. + lemma cpx_tdeq_conf_lexs: ∀h,o,G. R_confluent2_lfxs … (cpx h G) (cdeq h o) (cpx h G) (cdeq h o). #h #o #G #L0 #T0 #T1 #H @(cpx_ind … H) -G -L0 -T0 -T1 /2 width=3 by ex2_intro/ [ #G #L0 #s0 #X0 #H0 #L1 #HL01 #L2 #HL02 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma index 95b89cd80..c1d816668 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lfxs_lfxs.ma @@ -89,6 +89,26 @@ elim (HR … Hf … H0) -HR -Hf -H0 /4 width=7 by sle_lexs_trans, ex3_2_intro, ex2_intro/ qed-. +lemma lfxs_bind_dx_split_void: ∀R1,R2. (∀L. reflexive … (R1 L)) → (∀L. reflexive … (R2 L)) → + lexs_frees_confluent … (cext2 R1) cfull → + ∀L1,L2,T. L1.ⓧ ⪤*[R1, T] L2 → ∀p,I,V. + ∃∃L. L1 ⪤*[R1, ⓑ{p,I}V.T] L & L.ⓧ ⪤*[R2, T] L2. +#R1 #R2 #HR1 #HR2 #HR #L1 #L2 #T * #f #Hf #HL12 #p #I #V +elim (frees_total L1 (ⓑ{p,I}V.T)) #g #Hg +elim (frees_inv_bind_void … Hg) #y1 #y2 #_ #H #Hy +lapply(frees_mono … H … Hf) -H #H2 +lapply (tl_eq_repl … H2) -H2 #H2 +lapply (sor_eq_repl_back2 … Hy … H2) -y2 #Hy +lapply (sor_inv_sle_dx … Hy) -y1 #Hfg +lapply (sle_inv_tl_sn … Hfg) -Hfg #Hfg +elim (lexs_sle_split (cext2 R1) (cext2 R2) … HL12 … Hfg) -HL12 /2 width=1 by ext2_refl/ #Y #H #HL2 +lapply (sle_lexs_trans … H … Hfg) // #H0 +elim (lexs_inv_next1 … H) -H #Z #L #HL1 #H +elim (ext2_inv_unit_sn … H) -H #H destruct +elim (HR … Hf … H0) -HR -Hf -H0 +/4 width=7 by sle_lexs_trans, ex2_intro/ (* note: 2 ex2_intro *) +qed-. + (* Main properties **********************************************************) (* Basic_2A1: uses: llpx_sn_bind llpx_sn_bind_O *) 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 8b1cc6982..7310545d7 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 @@ -89,7 +89,7 @@ table { [ { "strongly normalizing rt-computation" * } { [ "lcosx ( ? ⊢ ~⬊*[?,?,?] ? )" "lcosx_cpx" * ] [ "lsx_drop" + "lsx_lpx" + "lsx_lpxs" + "llsx_csx" * ] - [ "csx_lleq" + "csx_lpxs" + "csx_fpbs" * ] + [ "csx_fpbs" * ] } ] [ { "parallel qrst-computation" * } { @@ -110,7 +110,7 @@ table { [ { "uncounted context-sensitive rt-computation" * } { [ "lfsx ( ? ⊢ ⬈*[?,?,?] 𝐒⦃?⦄ )" "lfsx_fqup" + "lfsx_lfpxs" + "lfsx_lfsx" * ] [ "csx_vector ( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_cnx_vector" + "csx_csx_vector" * ] - [ "csx ( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_simple" + "csx_simple_theq" + "csx_drops" + "csx_lsubr" + "csx_aaa" + "csx_gcp" + "csx_gcr" + "csx_lfpx" + "csx_cnx" + "csx_cpxs" + "csx_csx" * ] + [ "csx ( ⦃?,?⦄ ⊢ ⬈*[?,?] 𝐒⦃?⦄ )" "csx_simple" + "csx_simple_theq" + "csx_drops" + "csx_lsubr" + "csx_lfdeq" + "csx_aaa" + "csx_gcp" + "csx_gcr" + "csx_lfpx" + "csx_cnx" + "csx_cpxs" + "csx_lfpxs" + "csx_csx" * ] [ "lfpxs ( ⦃?,?⦄ ⊢ ⬈*[?,?] ? )" "lfpxs_length" + "lfpxs_drops" + "lfpxs_fqup" + "lfpxs_lfdeq" + "lfpxs_aaa" + "lfpxs_cpxs" + "lfpxs_lfpxs" * ] [ "cpxs ( ⦃?,?⦄ ⊢ ? ⬈*[?] ? )" "cpxs_tdeq" + "cpxs_theq" + "cpxs_theq_vector" + "cpxs_drops" + "cpxs_fqus" + "cpxs_lsubr" + "cpxs_lfdeq" + "cpxs_aaa" + "cpxs_lfpx" + "cpxs_cnx" + "cpxs_cpxs" * ] }