From: Ferruccio Guidi Date: Sun, 21 Apr 2013 15:51:01 +0000 (+0000) Subject: - we commit the "reduction" component X-Git-Tag: make_still_working~1181 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=09af7a9751464291ec3f32fb80c92fe1accdbf88;p=helm.git - we commit the "reduction" component - we park the focalized reduction for the moment - some refactoring --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cnf_cif.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cnf_cif.etc new file mode 100644 index 000000000..5a1ae5a5b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cnf_cif.etc @@ -0,0 +1,106 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cif.ma". +include "basic_2/reducibility/cnf_lift.ma". + +(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) + +(* Main properties **********************************************************) + +lemma tps_cif_eq: ∀L,T1,T2,d,e. L ⊢ T1 ▶[d, e] T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +[ // +| #L #K #V #W #i #d #e #_ #_ #HLK #_ #H -d -e + elim (cif_inv_delta … HLK ?) // +| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H + elim (cif_inv_bind … H) -H #HV1 #HT1 * #H destruct + lapply (IHV12 … HV1) -IHV12 -HV1 #H destruct /3 width=1/ +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H + elim (cif_inv_flat … H) -H #HV1 #HT1 #_ #_ /3 width=1/ +] +qed. + +lemma tpss_cif_eq: ∀L,T1,T2,d,e. L ⊢ T1 ▶*[d, e] T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2. +#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT1 #HT1 +lapply (IHT1 HT1) -IHT1 #H destruct /2 width=5/ +qed. + +lemma tpr_cif_eq: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ 𝐈⦃T1⦄ → T1 = T2. +#T1 #T2 #H elim H -T1 -T2 +[ // +| * #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #L #H + [ elim (cif_inv_appl … H) -H #HV1 #HT1 #_ + >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 // + | elim (cif_inv_ri2 … H) /2 width=1/ + ] +| #a #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #L #H + elim (cif_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +| #a * #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV1 #IHT1 #L #H + [ lapply (tps_lsubr_trans … HT2 (L.ⓓV2) ?) -HT2 /2 width=1/ #HT2 + elim (cif_inv_bind … H) -H #HV1 #HT1 * #H destruct + lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct + lapply (IHT1 … HT1) -IHT1 #H destruct + lapply (tps_cif_eq … HT2 ?) -HT2 // + | <(tps_inv_refl_SO2 … HT2 ?) -HT2 // + elim (cif_inv_ib2 … H) -H /2 width=1/ /3 width=2/ + ] +| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #L #H + elim (cif_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +| #V1 #T1 #T #T2 #_ #_ #_ #L #H + elim (cif_inv_ri2 … H) /2 width=1/ +| #V1 #T1 #T2 #_ #_ #L #H + elim (cif_inv_ri2 … H) /2 width=1/ +] +qed. + +lemma cpr_cif_eq: ∀L,T1,T2. L ⊢ T1 ➡ T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2. +#L #T1 #T2 * #T0 #HT10 #HT02 #HT1 +lapply (tpr_cif_eq … HT10 … HT1) -HT10 #H destruct /2 width=5/ +qed. + +theorem cif_cnf: ∀L,T. L ⊢ 𝐈⦃T⦄ → L ⊢ 𝐍⦃T⦄. +/3 width=3/ qed. + +(* Note: this property is unusual *) +lemma cnf_crf_false: ∀L,T. L ⊢ 𝐑⦃T⦄ → L ⊢ 𝐍⦃T⦄ → ⊥. +#L #T #H elim H -L -T +[ #L #K #V #i #HLK #H + elim (cnf_inv_delta … HLK H) +| #L #V #T #_ #IHV #H + elim (cnf_inv_appl … H) -H /2 width=1/ +| #L #V #T #_ #IHT #H + elim (cnf_inv_appl … H) -H /2 width=1/ +| #I #L #V #T * #H1 #H2 destruct + [ elim (cnf_inv_zeta … H2) + | elim (cnf_inv_tau … H2) + ] +|5,6: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct + [1,3: elim (cnf_inv_abbr … H2) -H2 /2 width=1/ + |*: elim (cnf_inv_abst … H2) -H2 /2 width=1/ + ] +| #a #L #V #W #T #H + elim (cnf_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +| #a #L #V #W #T #H + elim (cnf_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +] +qed. + +theorem cnf_cif: ∀L,T. L ⊢ 𝐍⦃T⦄ → L ⊢ 𝐈⦃T⦄. +/2 width=4/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_aaa.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_aaa.etc new file mode 100644 index 000000000..7c1273334 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_aaa.etc @@ -0,0 +1,25 @@ +(**************************************************************************) +(* ___ *) +(* ||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/static/aaa_ltpss_sn.ma". +include "basic_2/reducibility/ltpr_aaa.ma". +include "basic_2/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Properties about atomic arity assignment on terms ************************) + +lemma aaa_cpr_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → ∀T2. L ⊢ T1 ➡ T2 → L ⊢ T2 ⁝ A. +#L #T1 #A #HT1 #T2 * /3 width=5/ +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_cpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_cpr.etc new file mode 100644 index 000000000..d231c849f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_cpr.etc @@ -0,0 +1,63 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/tpr_tpr.ma". +include "basic_2/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Advanced properties ******************************************************) + +lemma cpr_bind_sn: ∀a,I,L,V1,V2,T1,T2. L ⊢ V1 ➡ V2 → T1 ➡ T2 → + L ⊢ ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2. +#a #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 #HT12 +@ex2_intro [2: @(tpr_delta … HV1 HT12) | skip ] /2 width=3/ (* /3 width=5/ is too slow *) +qed. + +(* Basic_1: was only: pr2_gen_cbind *) +lemma cpr_bind_dx: ∀a,I,L,V1,V2,T1,T2. V1 ➡ V2 → L. ⓑ{I} V2 ⊢ T1 ➡ T2 → + L ⊢ ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2. +#a #I #L #V1 #V2 #T1 #T2 #HV12 * #T #HT1 normalize #HT2 +elim (tpss_split_up … HT2 1 ? ?) -HT2 // #T0 (ltpr_fwd_length … HL12) in HT2; #HT2 +elim (ltpr_tpr_tpss_conf … HL12 … HT2) -L1 /3 width=3/ +qed. + +lemma cpr_ltpr_conf_tpss: ∀L1,L2. L1 ➡ L2 → ∀T1,T2. L1 ⊢ T1 ➡ T2 → + ∀d,e,U1. L1 ⊢ T1 ▶* [d, e] U1 → + ∃∃U2. L2 ⊢ U1 ➡ U2 & L2 ⊢ T2 ➡ U2. +#L1 #L2 #HL12 #T1 #T2 #HT12 #d #e #U1 #HTU1 +elim (cpr_ltpr_conf_eq … HT12 … HL12) -HT12 #T #HT1 #HT2 +elim (cpr_tpss_ltpr … HL12 … HT1 … HTU1) -L1 -HT1 #U2 #HU12 #HTU2 +lapply (tpss_weak_full … HTU2) -HTU2 #HTU2 /3 width=5/ (**) (* /4 width=5/ is too slow *) +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_ltpss_dx.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_ltpss_dx.etc new file mode 100644 index 000000000..5d90e6865 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_ltpss_dx.etc @@ -0,0 +1,28 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/ltpss_dx_ltpss_dx.ma". +include "basic_2/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Properties concerning dx partial unfold on local environments ************) + +lemma ltpss_dx_cpr_conf: ∀L1,T,U1. L1 ⊢ T ➡ U1 → + ∀L2,d,e. L1 ▶* [d, e] L2 → + ∃∃U2. L2 ⊢ T ➡ U2 & L2 ⊢ U1 ▶* [d, e] U2. +#L1 #T #U1 * #U #HTU #HU1 #L2 #d #e #HL12 +lapply (ltpss_dx_fwd_length … HL12) #H >H in HU1; -H #HU1 +elim (ltpss_dx_tpss_conf … HU1 … HL12) -L1 /3 width=3/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_ltpss_sn.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_ltpss_sn.etc new file mode 100644 index 000000000..f946e4cfb --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_ltpss_sn.etc @@ -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/unfold/ltpss_sn_alt.ma". +include "basic_2/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Properties concerning sn partial unfold on local environments ************) + +lemma ltpss_sn_cpr_trans: ∀L1,L2,d,e. L1 ⊢ ▶* [d, e] L2 → + ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡ T2. +#L1 #L2 #d #e #HL12 #T1 #T2 * +lapply (ltpss_sn_weak_full … HL12) +<(ltpss_sn_fwd_length … HL12) -HL12 /3 width=5/ +qed. + +lemma ltpss_sn_cpr_conf: ∀L1,T,U1. L1 ⊢ T ➡ U1 → + ∀L2,d,e. L1 ⊢ ▶* [d, e] L2 → + ∃∃U2. L2 ⊢ T ➡ U2 & L1 ⊢ U1 ▶* [d, e] U2. +#L1 #T #U1 * #U #HTU #HU1 #L2 #d #e #HL12 +lapply (ltpss_sn_fwd_length … HL12) #H >H in HU1; -H #HU1 +elim (ltpss_sn_tpss_conf … HU1 … HL12) -HU1 -HL12 /3 width=3/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_tpss.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_tpss.etc new file mode 100644 index 000000000..d8c7225d9 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_tpss.etc @@ -0,0 +1,39 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/ltpr_tpss.ma". +include "basic_2/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Properties on partial unfold for terms ***********************************) + +lemma cpr_tpss_trans: ∀L,T1,T. L ⊢ T1 ➡ T → + ∀T2,d,e. L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ➡ T2. +#L #T1 #T * #T0 #HT10 #HT0 #T2 #d #e #HT2 +lapply (tpss_weak_full … HT2) -HT2 #HT2 +lapply (tpss_trans_eq … HT0 HT2) -T /2 width=3/ +qed. + +lemma cpr_tps_trans: ∀L,T1,T. L ⊢ T1 ➡ T → + ∀T2,d,e. L ⊢ T ▶ [d, e] T2 → L ⊢ T1 ➡ T2. +/3 width=5/ qed. + +lemma cpr_tpss_conf: ∀L,T0,T1. L ⊢ T0 ➡ T1 → + ∀T2,d,e. L ⊢ T0 ▶* [d, e] T2 → + ∃∃T. L ⊢ T1 ▶* [d, e] T & L ⊢ T2 ➡ T. +#L #T0 #T1 * #U0 #HTU0 #HU0T1 #T2 #d #e #HT02 +elim (tpr_tpss_conf … HTU0 … HT02) -T0 #T0 #HT20 #HUT0 +elim (tpss_conf_eq … HU0T1 … HUT0) -U0 /3 width=5/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr.etc new file mode 100644 index 000000000..ea9f585a6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr.etc @@ -0,0 +1,48 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +notation "hvbox( ⦃ term 46 L1 ⦄ ➡ break ⦃ term 46 L2 ⦄ )" + non associative with precedence 45 + for @{ 'FocalizedPRed $L1 $L2 }. + +include "basic_2/unfold/ltpss_sn.ma". +include "basic_2/reducibility/ltpr.ma". + +(* FOCALIZED PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ***********************) + +definition lfpr: relation lenv ≝ + λL1,L2. ∃∃L. L1 ➡ L & L ⊢ ▶* [0, |L|] L2 +. + +interpretation + "focalized parallel reduction (environment)" + 'FocalizedPRed L1 L2 = (lfpr L1 L2). + +(* Basic properties *********************************************************) + +(* Note: lemma 250 *) +lemma lfpr_refl: ∀L. ⦃L⦄ ➡ ⦃L⦄. +/2 width=3/ qed. + +lemma ltpss_sn_lfpr: ∀L1,L2,d,e. L1 ⊢ ▶* [d, e] L2 → ⦃L1⦄ ➡ ⦃L2⦄. +/3 width=5/ qed. + +lemma ltpr_lfpr: ∀L1,L2. L1 ➡ L2 → ⦃L1⦄ ➡ ⦃L2⦄. +/3 width=3/ qed. + +(* Basic inversion lemmas ***************************************************) + +lemma lfpr_inv_atom1: ∀L2. ⦃⋆⦄ ➡ ⦃L2⦄ → L2 = ⋆. +#L2 * #L #HL >(ltpr_inv_atom1 … HL) -HL #HL2 >(ltpss_sn_inv_atom1 … HL2) -HL2 // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_aaa.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_aaa.etc new file mode 100644 index 000000000..6f6c49df3 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_aaa.etc @@ -0,0 +1,25 @@ +(**************************************************************************) +(* ___ *) +(* ||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/static/aaa_ltpss_sn.ma". +include "basic_2/reducibility/ltpr_aaa.ma". +include "basic_2/reducibility/lfpr.ma". + +(* FOCALIZED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS **********************) + +(* Properties about atomic arity assignment on terms ************************) + +lemma aaa_lfpr_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. ⦃L1⦄ ➡ ⦃L2⦄ → L2 ⊢ T ⁝ A. +#L1 #T #A #HT #L2 * /3 width=5/ +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_cpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_cpr.etc new file mode 100644 index 000000000..8a01f263a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_cpr.etc @@ -0,0 +1,52 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/ltpss_sn_ltpss_sn.ma". +include "basic_2/reducibility/ltpr_ldrop.ma". +include "basic_2/reducibility/cpr.ma". +include "basic_2/reducibility/lfpr.ma". + +(* FOCALIZED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS **********************) + +(* Advanced properties ******************************************************) + +lemma lfpr_pair_cpr: ∀L1,L2. ⦃L1⦄ ➡ ⦃L2⦄ → ∀V1,V2. L2 ⊢ V1 ➡ V2 → + ∀I. ⦃L1. ⓑ{I} V1⦄ ➡ ⦃L2. ⓑ{I} V2⦄. +#L1 #L2 * #L #HL1 #HL2 #V1 #V2 * +<(ltpss_sn_fwd_length … HL2) #V #HV1 #HV2 #I +lapply (ltpss_sn_tpss_trans_eq … HV2 … HL2) -HV2 #V2 +@(ex2_intro … (L.ⓑ{I}V)) /2 width=1/ (**) (* explicit constructor *) +qed. + +(* Properties on supclosure *************************************************) +(* +lamma fsub_cpr_trans: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ➡ U2 → + ∃∃L,U1. ⦃L1⦄ ➡ ⦃L⦄ & L ⊢ T1 ➡ U1 & ⦃L, U1⦄ ⊃ ⦃L2, U2⦄. +#L1 #L2 #T1 #T2 #HT12 #U2 * #T #H1 #H2 +elim (fsub_tpr_trans … HT12 … H1) -T2 #L #U #HL1 #HT1U #HUT +elim (fsup_tpss_trans_full … HUT … H2) -T -HUT -H2 #L #U #HL1 #HT1U #HUT + + + + + + + #H elim H -L1 -L2 -T1 -T2 [1,2,3,4,5: /3 width=5/ ] +#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 +elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HK2 +elim (lift_total T d e) #U #HTU +elim (ldrop_ltpr_trans … HLK1 … HK1) -HLK1 -HK1 #L #HL1 #HLK +lapply (tpr_lift … HT1 … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ +qed-. +*) \ No newline at end of file diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_lfpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_lfpr.etc new file mode 100644 index 000000000..7031e792b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/lfpr_lfpr.etc @@ -0,0 +1,39 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/ltpr_ltpss_sn.ma". +include "basic_2/reducibility/ltpr_ltpr.ma". +include "basic_2/reducibility/lfpr.ma". + +(* FOCALIZED PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ***********************) + +(* Main properties **********************************************************) + +theorem lfpr_conf: ∀L0,L1,L2. ⦃L0⦄ ➡ ⦃L1⦄ → ⦃L0⦄ ➡ ⦃L2⦄ → + ∃∃L. ⦃L1⦄ ➡ ⦃L⦄ & ⦃L2⦄ ➡ ⦃L⦄. +#K0 #L1 #L2 * #K1 #HK01 #HKL1 * #K2 #HK02 #HKL2 +lapply (ltpr_fwd_length … HK01) #H +>(ltpr_fwd_length … HK02) in H; #H +elim (ltpr_conf … HK01 … HK02) -K0 #K #HK1 #HK2 +lapply (ltpss_sn_fwd_length … HKL1) #H1 +lapply (ltpss_sn_fwd_length … HKL2) #H2 +>H1 in HKL1 H; -H1 #HKL1 +>H2 in HKL2; -H2 #HKL2 #H +elim (ltpr_ltpss_sn_conf … HKL1 … HK1) -K1 #K1 #HK1 #HLK1 +elim (ltpr_ltpss_sn_conf … HKL2 … HK2) -K2 #K2 #HK2 #HLK2 +elim (ltpss_sn_conf … HK1 … HK2) -K #K #HK1 #HK2 +lapply (ltpr_fwd_length … HLK1) #H1 +lapply (ltpr_fwd_length … HLK2) #H2 +/3 width=5/ +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_aaa.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_aaa.etc new file mode 100644 index 000000000..5208a6405 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_aaa.etc @@ -0,0 +1,85 @@ +(**************************************************************************) +(* ___ *) +(* ||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/static/aaa_ltpss_dx.ma". +include "basic_2/static/lsuba_aaa.ma". +include "basic_2/reducibility/ltpr_ldrop.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +(* Properties about atomic arity assignment on terms ************************) + +lemma aaa_ltpr_tpr_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → ∀L2. L1 ➡ L2 → + ∀T2. T1 ➡ T2 → L2 ⊢ T2 ⁝ A. +#L1 #T1 @(f2_ind … fw … L1 T1) -L1 -T1 #n #IH #L1 * * +[1,2,3: + #i #Hn #X #H1 #L2 #HL12 #Y #H2 destruct + >(tpr_inv_atom1 … H2) -Y +|4,5: [ #a ] * #V1 #T1 #Hn #X #H1 #L2 #HL12 #Y #H2 destruct +] +[ >(aaa_inv_sort … H1) -X // +| elim (aaa_inv_lref … H1) #I #K1 #V1 #HLK1 #HA + lapply (ldrop_pair2_fwd_fw … HLK1 (#i)) #HKV1 + elim (ltpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #Y #H #HLK2 + elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct + lapply (IH … HKV1 … HA … HK12 … HV12) -L1 -K1 -V1 /2 width=5/ +| elim (aaa_inv_gref … H1) +| elim (aaa_inv_abbr … H1) -H1 #B #HB #HA + elim (tpr_inv_abbr1 … H2) -H2 * + [ #V2 #T #T2 #HV12 #HT1 #HT2 #H destruct + lapply (tps_lsubr_trans … HT2 (L2.ⓓV2) ?) -HT2 /2 width=1/ #HT2 + lapply (IH … HB … HL12 … HV12) -HB /width=5/ #HB + lapply (IH … HA … (L2.ⓓV2) … HT1) -IH -HA -HT1 /width=5/ -T1 /2 width=1/ -L1 -V1 /3 width=5/ + | -B #T #HT1 #HXT #H destruct + lapply (IH … HA … (L2.ⓓV1) … HT1) /width=5/ -T1 /2 width=1/ -L1 #HA + @(aaa_inv_lift … HA … HXT) /2 width=1/ + ] +| elim (aaa_inv_abst … H1) -H1 #B #A #HB #HA #H destruct + elim (tpr_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HT12 #H destruct + lapply (IH … HB … HL12 … HV12) -HB /width=5/ #HB + lapply (IH … HA … (L2.ⓛV2) … HT12) -IH -HA -HT12 /width=5/ -T1 /2 width=1/ +| elim (aaa_inv_appl … H1) -H1 #B #HB #HA + elim (tpr_inv_appl1 … H2) -H2 * + [ #V2 #T2 #HV12 #HT12 #H destruct + lapply (IH … HB … HL12 … HV12) -HB -HV12 /width=5/ #HB + lapply (IH … HA … HL12 … HT12) -IH -HA -HL12 -HT12 /width=5/ /2 width=3/ + | #b #V2 #W2 #T0 #T2 #HV12 #HT02 #H1 #H2 destruct + elim (aaa_inv_abst … HA) -HA #B0 #A0 #HB0 #HA0 #H destruct + lapply (IH … HB … HL12 … HV12) -HB -HV12 /width=5/ #HB + lapply (IH … HB0 … HL12 W2 ?) -HB0 /width=5/ #HB0 + lapply (IH … HA0 … (L2.ⓛW2) … HT02) -IH -HA0 -HT02 // /2 width=1/ -T0 -L1 -V1 /4 width=7/ + | #b #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HW02 #HT02 #HV02 #H1 #H2 destruct + elim (aaa_inv_abbr … HA) -HA #B0 #HW0 #HT0 + lapply (IH … HW0 … HL12 … HW02) -HW0 /width=5/ #HW2 + lapply (IH … HB … HL12 … HV10) -HB -HV10 /width=5/ #HV0 + lapply (IH … HT0 … (L2.ⓓW2) … HT02) -IH -HT0 -HT02 // /2 width=1/ -V1 -T0 -L1 -W0 #HT2 + @(aaa_abbr … HW2) -HW2 + @(aaa_appl … HT2) -HT2 /3 width=7/ (**) (* explict constructors, /5 width=7/ is too slow *) + ] +| elim (aaa_inv_cast … H1) -H1 #HV1 #HT1 + elim (tpr_inv_cast1 … H2) -H2 + [ * #V2 #T2 #HV12 #HT12 #H destruct + lapply (IH … HV1 … HL12 … HV12) -HV1 -HV12 /width=5/ #HV2 + lapply (IH … HT1 … HL12 … HT12) -IH -HT1 -HL12 -HT12 /width=5/ -L1 -V1 -T1 /2 width=1/ + | -HV1 #HT1X + lapply (IH … HT1 … HL12 … HT1X) -IH -HT1 -HL12 -HT1X /width=5/ + ] +] +qed. + +lemma aaa_ltpr_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. L1 ➡ L2 → L2 ⊢ T ⁝ A. +/2 width=5/ qed. + +lemma aaa_tpr_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → ∀T2. T1 ➡ T2 → L ⊢ T2 ⁝ A. +/2 width=5/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_ltpss_dx.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_ltpss_dx.etc new file mode 100644 index 000000000..00730b732 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_ltpss_dx.etc @@ -0,0 +1,36 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/ltpr_tpss.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +(* Properties concerning dx parallel unfold on local environments ***********) + +lemma ltpr_ltpss_dx_conf: ∀L1,K1,d,e. L1 ▶* [d, e] K1 → ∀L2. L1 ➡ L2 → + ∃∃K2. L2 ▶* [d, e] K2 & K1 ➡ K2. +#L1 #K1 #d #e #H elim H -L1 -K1 -d -e +[ /2 width=3/ +| #L1 #I #V1 #X #H + elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct /3 width=5/ +| #L1 #K1 #I #V1 #W1 #e #_ #HVW1 #IHLK1 #X #H + elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct + elim (IHLK1 … HL12) -L1 #K2 #HLK2 #HK12 + elim (ltpr_tpr_tpss_conf … HK12 … HV12 … HVW1) -V1 /3 width=5/ +| #L1 #K1 #I #V1 #W1 #d #e #_ #HVW1 #IHLK1 #X #H + elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct + elim (IHLK1 … HL12) -L1 #K2 #HLK2 #HK12 + elim (ltpr_tpr_tpss_conf … HK12 … HV12 … HVW1) -V1 /3 width=5/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_ltpss_sn.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_ltpss_sn.etc new file mode 100644 index 000000000..79ffb7c9e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_ltpss_sn.etc @@ -0,0 +1,31 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/ltpss_sn_alt.ma". +include "basic_2/reducibility/ltpr_ltpss_dx.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +(* Properties on sn parallel unfold on local environments *******************) + +(* Note: this can also be proved like ltpr_ltpss_dx_conf *) +lemma ltpr_ltpss_sn_conf: ∀L1,K1,d,e. L1 ⊢ ▶* [d, e] K1 → ∀L2. L1 ➡ L2 → + ∃∃K2. L2 ⊢ ▶* [d, e] K2 & K1 ➡ K2. +#L1 #K1 #d #e #H +lapply (ltpss_sn_ltpssa … H) -H #H +@(ltpssa_ind … H) -K1 /2 width=3/ +#K #K1 #_ #HK1 #IHK #L2 #HL12 +elim (IHK … HL12) -L1 #K2 #HLK2 #HK2 +elim (ltpr_ltpss_dx_conf … HK1 … HK2) -K /3 width=3/ +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_tps.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_tps.etc new file mode 100644 index 000000000..7f08be62b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_tps.etc @@ -0,0 +1,57 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/ltpr_ldrop.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +(* Properties concerning parallel substitution on terms *********************) + +(* Basic_1: was: pr0_subst1_fwd *) +lemma ltpr_tps_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶ [d, e] T2 → ∀L2. L1 ➡ L2 → + ∃∃T. L2 ⊢ T1 ▶ [d, e] T & T2 ➡ T. +#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e +[ /2 width=3/ +| #L1 #K1 #V1 #W1 #i #d #e #Hdi #Hide #HLK1 #HVW1 #L2 #HL12 + elim (ltpr_ldrop_conf … HLK1 … HL12) -L1 #X #H #HLK2 + elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct -K1 + elim (lift_total V2 0 (i+1)) #W2 #HVW2 + lapply (tpr_lift … HV12 … HVW1 … HVW2) -V1 /3 width=4/ +| #L1 #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12 + elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2 + elim (IHT12 (L2.ⓑ{I}V) ?) /2 width=1/ -L1 /3 width=5/ +| #L1 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12 + elim (IHV12 … HL12) -IHV12 + elim (IHT12 … HL12) -L1 /3 width=5/ +] +qed-. + +(* Basic_1: was: pr0_subst1_back *) +lemma ltpr_tps_trans: ∀L2,T1,T2,d,e. L2 ⊢ T1 ▶ [d, e] T2 → ∀L1. L1 ➡ L2 → + ∃∃T. L1 ⊢ T1 ▶ [d, e] T & T ➡ T2. +#L2 #T1 #T2 #d #e #H elim H -L2 -T1 -T2 -d -e +[ /2 width=3/ +| #L2 #K2 #V2 #W2 #i #d #e #Hdi #Hide #HLK2 #HVW2 #L1 #HL12 + elim (ltpr_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H + elim (ltpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct -K2 + elim (lift_total V1 0 (i+1)) #W1 #HVW1 + lapply (tpr_lift … HV12 … HVW1 … HVW2) -V2 /3 width=4/ +| #L2 #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12 + elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2 + elim (IHT12 (L1.ⓑ{I}V) ?) /2 width=1/ -L2 /3 width=5/ +| #L2 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12 + elim (IHV12 … HL12) -IHV12 + elim (IHT12 … HL12) -L2 /3 width=5/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_tpss.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_tpss.etc new file mode 100644 index 000000000..71586fa36 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpr_tpss.etc @@ -0,0 +1,91 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/ltpss_dx_ltpss_dx.ma". +include "basic_2/reducibility/ltpr_tps.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) + +(* Properties on partial unfold for terms ***********************************) + +(* Basic_1: was: pr0_subst1 *) +lemma ltpr_tpr_tps_conf: ∀T1,T2. T1 ➡ T2 → + ∀L1,d,e,U1. L1 ⊢ T1 ▶ [d, e] U1 → + ∀L2. L1 ➡ L2 → + ∃∃U2. U1 ➡ U2 & L2 ⊢ T2 ▶* [d, e] U2. +#T1 #T2 #H elim H -T1 -T2 +[ #I #L1 #d #e #U1 #H #L2 #HL12 + elim (ltpr_tps_conf … H … HL12) -L1 /3 width=3/ +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_flat1 … H) -H #W1 #U1 #HVW1 #HTU1 #H destruct + elim (IHV12 … HVW1 … HL12) -V1 + elim (IHT12 … HTU1 … HL12) -T1 -HL12 /3 width=5/ +| #a #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct + elim (tps_inv_bind1 … HY) -HY #WW #TT1 #_ #HTT1 #H destruct + elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 + elim (IHT12 … HTT1 (L2. ⓛWW) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 + lapply (tpss_lsubr_trans … HTT2 (L2. ⓓVV2) ?) -HTT2 /3 width=5/ +| #a #I #V1 #V2 #T1 #T #T2 #HV12 #_ #HT2 #IHV12 #IHT1 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_bind1 … H) -H #W1 #U1 #HVW1 #HTU1 #H destruct + elim (IHV12 … HVW1 … HL12) -V1 #W2 #HW12 #HVW2 + elim (IHT1 … HTU1 (L2. ⓑ{I} W2) ?) -T1 /2 width=1/ -HL12 #U #HU1 #HTU + elim (tpss_strip_neq … HTU … HT2 ?) -T /2 width=1/ #U2 #HU2 #HTU2 + lapply (tps_lsubr_trans … HU2 (L2. ⓑ{I} V2) ?) -HU2 /2 width=1/ #HU2 + elim (ltpss_dx_tps_conf … HU2 (L2. ⓑ{I} W2) (d + 1) e ?) -HU2 /2 width=1/ #U3 #HU3 #HU23 + lapply (tps_lsubr_trans … HU3 (⋆. ⓑ{I} W2) ?) -HU3 /2 width=1/ #HU3 + lapply (tpss_lsubr_trans … HU23 (L2. ⓑ{I} W2) ?) -HU23 /2 width=1/ #HU23 + lapply (tpss_trans_eq … HTU2 … HU23) -U2 /3 width=5/ +| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct + elim (tps_inv_bind1 … HY) -HY #WW1 #TT1 #HWW1 #HTT1 #H destruct + elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 + elim (IHW12 … HWW1 … HL12) -W1 #WW2 #HWW12 #HWW2 + elim (IHT12 … HTT1 (L2. ⓓWW2) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 + elim (lift_total VV2 0 1) #VV #H2VV + lapply (tpss_lift_ge … HVV2 (L2. ⓓWW2) … HV2 … H2VV) -V2 /2 width=1/ #HVV + @ex2_intro [2: @tpr_theta |1: skip |3: @tpss_bind [2: @tpss_flat ] ] /width=11/ (**) (* /4 width=11/ is too slow *) +| #V #T1 #T #T2 #_ #HT2 #IHT1 #L1 #d #e #X #H #L2 #HL12 + elim (tps_inv_bind1 … H) -H #W #U1 #_ #HTU1 #H destruct -V + elim (IHT1 … HTU1 (L2.ⓓW) ?) -T1 /2 width=1/ -HL12 #U #HU1 #HTU + elim (tpss_inv_lift1_ge … HTU L2 … HT2 ?) -T (length_inv_zero_sn … H) /2 width=1/ +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma fpr_inv_pair1_sn: ∀I,K1,L2,V1,T1,T2. ⦃⋆.ⓑ{I}V1@@K1, T1⦄ ➡ ⦃L2, T2⦄ → + ∃∃K2,V2. V1 ➡ V2 & + ⋆.ⓑ{I}V2 ⊢ ⦃K1, T1⦄ ➡ ⦃K2, T2⦄ & + L2 = ⋆.ⓑ{I}V2@@K2. +#I1 #K1 #L2 #V1 #T1 #T2 * >append_length #H +elim (length_inv_pos_sn_append … H) -H #I2 #K2 #V2 #HK12 #H destruct +>shift_append_assoc >shift_append_assoc normalize in ⊢ (%→?); #H +elim (tpr_inv_bind1 … H) -H * +[ #V0 #T #T0 #HV10 #HT1 #HT0 #H destruct /5 width=5/ +| #T0 #_ #_ #H destruct +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_aaa.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_aaa.etc new file mode 100644 index 000000000..72802e499 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_aaa.etc @@ -0,0 +1,27 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr_aaa.ma". +include "basic_2/reducibility/cfpr_cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) + +(* Properties about atomic arity assignment on terms ************************) + +lemma aaa_fpr_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → + ∀L2,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → L2 ⊢ T2 ⁝ A. +#L1 #T1 #A #HT1 #L2 #T2 #H +elim (fpr_inv_all … H) -H +/4 width=5 by aaa_cpr_conf, aaa_ltpr_conf, aaa_ltpss_sn_conf/ +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_cfpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_cfpr.etc new file mode 100644 index 000000000..f442be28b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_cfpr.etc @@ -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/reducibility/cpr_cpr.ma". +include "basic_2/reducibility/cfpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) + +(* Main properties **********************************************************) + +theorem cfpr_conf: ∀L. bi_confluent … (cfpr L). +#L #L0 #L1 #T0 #T1 * #HL01 #HT01 #L2 #T2 * >HL01 #HL12 #HT02 +elim (cpr_conf … HT01 HT02) -L0 -T0 #X #H1 #H2 +elim (cpr_fwd_shift1 … H1) #L0 #T0 #HL10 #H destruct /3 width=5/ +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_cpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_cpr.etc new file mode 100644 index 000000000..a1e20a025 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_cpr.etc @@ -0,0 +1,64 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr_tpss.ma". +include "basic_2/reducibility/cpr_cpr.ma". +include "basic_2/reducibility/cfpr_ltpss.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) + +(* Advanced properties ******************************************************) + +lemma fpr_all: ∀L1,L. L1 ➡ L → ∀L2,T1,T2. L ⊢ T1 ➡ T2 → + L ⊢ ▶* [0, |L|] L2 → ⦃L1, T1⦄ ➡ ⦃L2, T2⦄. +#L1 #L #H elim H -L1 -L +[ #L2 #T1 #T2 #HT12 #HL2 + lapply (ltpss_sn_inv_atom1 … HL2) -HL2 #H destruct + lapply (cpr_inv_atom … HT12) -HT12 /2 width=1/ +| #I #L1 #L #V1 #V #_ #HV1 #IH #X #T1 #T2 #HT12 #H + elim (ltpss_sn_inv_tpss21 … H ?) -H // append_length >append_length #H + lapply (injective_plus_r … H) -H #H + @(ex3_intro … (⋆.ⓑ{I}V@@Y)) append_length H -H >commutative_plus /3 width=1/ +] +qed-. + +lemma fpr_inv_all: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → + ∃∃L. L1 ➡ L & L ⊢ T1 ➡ T2 & L ⊢ ▶* [0, |L|] L2. +#L1 #L2 #T1 #T2 #H +lapply (fpr_cfpr … H) -H #H +elim (cfpr_inv_all … H) -H /2 width=4/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_ltpss.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_ltpss.etc new file mode 100644 index 000000000..66e102e63 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/cfpr_ltpss.etc @@ -0,0 +1,39 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cpr_lift.ma". +include "basic_2/reducibility/cpr_ltpss_sn.ma". +include "basic_2/reducibility/cfpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) + +(* Advanced inversion lemmas ************************************************) + +lemma cfpr_inv_pair1: ∀I,L,K1,L2,V1,T1,T2. L ⊢ ⦃⋆.ⓑ{I}V1@@K1, T1⦄ ➡ ⦃L2, T2⦄ → + ∃∃K2,V,V2. V1 ➡ V & L ⊢ V ▶* [0, |L|] V2 & + L.ⓑ{I}V ⊢ ⦃K1, T1⦄ ➡ ⦃K2, T2⦄ & + L2 = ⋆.ⓑ{I}V2@@K2. +* #L #K1 #L2 #V1 #T1 #T2 * >append_length #H +elim (length_inv_pos_sn_append … H) -H #I2 #K2 #V2 #HK12 #H destruct +>shift_append_assoc >shift_append_assoc normalize in ⊢ (??%%→?); #H +[ elim (cpr_inv_abbr1 … H) -H * + [ #V #V0 #T0 #HV1 #HV0 #HT10 #H destruct /3 width=7/ + | #T0 #_ #_ #H destruct + ] +| elim (cpr_inv_abst1 … H Abst V2) -H + #V #T * #V0 #HV10 #HV0 #HT1 #H destruct + lapply (ltpss_sn_cpr_trans (L.ⓛV0) … 0 (|L|+1) … HT1) -HT1 /2 width=1/ #HT12 + /3 width=7/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr.etc new file mode 100644 index 000000000..3bb168f9e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr.etc @@ -0,0 +1,68 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +notation "hvbox( ⦃ term 46 L1, break term 46 T1 ⦄ ➡ break ⦃ term 46 L2 , break term 46 T2 ⦄ )" + non associative with precedence 45 + for @{ 'FocalizedPRed $L1 $T1 $L2 $T2 }. + +include "basic_2/reducibility/tpr.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON CLOSURES ******************************) + +definition fpr: bi_relation lenv term ≝ + λL1,T1,L2,T2. |L1| = |L2| ∧ L1 @@ T1 ➡ L2 @@ T2. + +interpretation + "context-free parallel reduction (closure)" + 'FocalizedPRed L1 T1 L2 T2 = (fpr L1 T1 L2 T2). + +(* Basic properties *********************************************************) + +lemma fpr_refl: bi_reflexive … fpr. +/2 width=1/ qed. + +lemma fpr_shift: ∀I1,I2,L1,L2,V1,V2,T1,T2. + ⦃L1, -ⓑ{I1}V1.T1⦄ ➡ ⦃L2, -ⓑ{I2}V2.T2⦄ → + ⦃L1.ⓑ{I1}V1, T1⦄ ➡ ⦃L2.ⓑ{I2}V2, T2⦄. +#I1 #I2 #L1 #L2 #V1 #V2 #T1 #T2 * #HL12 #HT12 +@conj // normalize // (**) (* explicit constructor *) +qed. + +(* Basic inversion lemmas ***************************************************) + +lemma fpr_inv_atom1: ∀L2,T1,T2. ⦃⋆, T1⦄ ➡ ⦃L2, T2⦄ → T1 ➡ T2 ∧ L2 = ⋆. +#L2 #T1 #T2 * #H +lapply (length_inv_zero_sn … H) -H #H destruct /2 width=1/ +qed-. + +lemma fpr_inv_atom3: ∀L1,T1,T2. ⦃L1,T1⦄ ➡ ⦃⋆,T2⦄ → T1 ➡ T2 ∧ L1 = ⋆. +#L1 #T1 #T2 * #H +lapply (length_inv_zero_dx … H) -H #H destruct /2 width=1/ +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma fpr_fwd_pair1: ∀I1,K1,L2,V1,T1,T2. ⦃K1.ⓑ{I1}V1, T1⦄ ➡ ⦃L2, T2⦄ → + ∃∃I2,K2,V2. ⦃K1, -ⓑ{I1}V1.T1⦄ ➡ ⦃K2, -ⓑ{I2}V2.T2⦄ & + L2 = K2.ⓑ{I2}V2. +#I1 #K1 #L2 #V1 #T1 #T2 * #H +elim (length_inv_pos_sn … H) -H #I2 #K2 #V2 #HK12 #H destruct /3 width=5/ +qed-. + +lemma fpr_fwd_pair3: ∀I2,L1,K2,V2,T1,T2. ⦃L1, T1⦄ ➡ ⦃K2.ⓑ{I2}V2, T2⦄ → + ∃∃I1,K1,V1. ⦃K1, -ⓑ{I1}V1.T1⦄ ➡ ⦃K2, -ⓑ{I2}V2.T2⦄ & + L1 = K1.ⓑ{I1}V1. +#I2 #L1 #K2 #V2 #T1 #T2 * #H +elim (length_inv_pos_dx … H) -H #I1 #K1 #V1 #HK12 #H destruct /3 width=5/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr_cpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr_cpr.etc new file mode 100644 index 000000000..baa630f5b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr_cpr.etc @@ -0,0 +1,125 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/cfpr_cpr.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON CLOSURES ******************************) + +(* Properties on context-sensitive parallel reduction for terms *************) + +lemma ltpr_tpr_fpr: ∀L1,L2. L1 ➡ L2 → ∀T1,T2. T1 ➡ T2 → ⦃L1, T1⦄ ➡ ⦃L2, T2⦄. +/3 width=4/ qed. + +lemma cpr_fpr: ∀L,T1,T2. L ⊢ T1 ➡ T2 → ⦃L, T1⦄ ➡ ⦃L, T2⦄. +/2 width=4/ qed. + +lemma fpr_lift: ∀K1,K2,T1,T2. ⦃K1, T1⦄ ➡ ⦃K2, T2⦄ → + ∀d,e,L1. ⇩[d, e] L1 ≡ K1 → + ∀U1,U2. ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + ∃∃L2. ⦃L1, U1⦄ ➡ ⦃L2, U2⦄ & ⇩[d, e] L2 ≡ K2. +#K1 #K2 #T1 #T2 #HT12 #d #e #L1 #HLK1 #U1 #U2 #HTU1 #HTU2 +elim (fpr_inv_all … HT12) -HT12 #K #HK1 #HT12 #HK2 +elim (ldrop_ltpr_trans … HLK1 … HK1) -K1 #L #HL1 #HLK +lapply (cpr_lift … HLK … HTU1 … HTU2 HT12) -T1 -T2 #HU12 +elim (le_or_ge (|K|) d) #Hd +[ elim (ldrop_ltpss_sn_trans_ge … HLK … HK2 …) +| elim (ldrop_ltpss_sn_trans_be … HLK … HK2 …) +] // -Hd #L2 #HL2 #HLK2 +lapply (ltpss_sn_weak_full … HL2) -K /3 width=4/ +qed-. + +(* Advanced properties ******************************************************) + +lemma fpr_flat_dx: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ∀V1,V2. V1 ➡ V2 → + ∀I. ⦃L1, ⓕ{I}V1.T1⦄ ➡ ⦃L2, ⓕ{I}V2.T2⦄. +#L1 #L2 #T1 #T2 #HT12 +elim (fpr_inv_all … HT12) -HT12 /4 width=4/ +qed. + +lemma fpr_bind_sn: ∀L1,L2,V1,V2. ⦃L1, V1⦄ ➡ ⦃L2, V2⦄ → ∀T1,T2. T1 ➡ T2 → + ∀a,I. ⦃L1, ⓑ{a,I}V1.T1⦄ ➡ ⦃L2, ⓑ{a,I}V2.T2⦄. +#L1 #L2 #V1 #V2 #H #T1 #T2 #HT12 #a #I +elim (fpr_inv_all … H) /3 width=4/ +qed. + +lemma fpr_bind2_minus: ∀I,L1,L2,V1,T1,U2. ⦃L1, -ⓑ{I}V1.T1⦄ ➡ ⦃L2, U2⦄ → + ∃∃V2,T2. ⦃L1.ⓑ{I}V1, T1⦄ ➡ ⦃L2.ⓑ{I}V2, T2⦄ & + U2 = -ⓑ{I}V2.T2. +#I1 #L1 #L2 #V1 #T1 #U2 #H +elim (fpr_inv_all … H) -H #L #HL1 #H #HL2 +elim (cpr_fwd_bind1_minus … H false) -H /4 width=4/ +qed-. + +(* Advanced forward lemmas **************************************************) + +lemma fpr_fwd_bind2_minus: ∀I,L1,L,V1,T1,T. ⦃L1, -ⓑ{I}V1.T1⦄ ➡ ⦃L, T⦄ → ∀b. + ∃∃V2,T2. ⦃L1, ⓑ{b,I}V1.T1⦄ ➡ ⦃L, ⓑ{b,I}V2.T2⦄ & + T = -ⓑ{I}V2.T2. +#I #L1 #L #V1 #T1 #T #H1 #b +elim (fpr_inv_all … H1) -H1 #L0 #HL10 #HT1 #HL0 +elim (cpr_fwd_bind1_minus … HT1 b) -HT1 /3 width=4/ +qed-. + +lemma fpr_fwd_shift_bind_minus: ∀I1,I2,L1,L2,V1,V2,T1,T2. + ⦃L1, -ⓑ{I1}V1.T1⦄ ➡ ⦃L2, -ⓑ{I2}V2.T2⦄ → + ⦃L1, V1⦄ ➡ ⦃L2, V2⦄ ∧ I1 = I2. +* #I2 #L1 #L2 #V1 #V2 #T1 #T2 #H +elim (fpr_inv_all … H) -H #L #HL1 #H #HL2 +[ elim (cpr_inv_abbr1 … H) -H * + [ #V #V0 #T #HV1 #HV0 #_ #H destruct /4 width=4/ + | #T #_ #_ #H destruct + ] +| elim (cpr_inv_abst1 … H Abst V2) -H + #V #T #HV1 #_ #H destruct /3 width=4/ +] +qed-. + +lemma fpr_fwd_abst2: ∀a,L1,L2,V1,T1,U2. ⦃L1, ⓛ{a}V1.T1⦄ ➡ ⦃L2, U2⦄ → ∀b,I,W. + ∃∃V2,T2. ⦃L1, ⓑ{b,I}W.T1⦄ ➡ ⦃L2, ⓑ{b,I}W.T2⦄ & + U2 = ⓛ{a}V2.T2. +#a #L1 #L2 #V1 #T1 #U2 #H +elim (fpr_inv_all … H) #L #HL1 #H #HL2 #b #I #W +elim (cpr_fwd_abst1 … H b I W) -H /3 width=4/ +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma fpr_inv_pair1: ∀I,K1,L2,V1,T1,T2. ⦃K1.ⓑ{I}V1, T1⦄ ➡ ⦃L2, T2⦄ → + ∃∃K2,V2. ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ & + ⦃K1, -ⓑ{I}V1.T1⦄ ➡ ⦃K2, -ⓑ{I}V2.T2⦄ & + L2 = K2.ⓑ{I}V2. +#I1 #K1 #X #V1 #T1 #T2 #H +elim (fpr_fwd_pair1 … H) -H #I2 #K2 #V2 #HT12 #H destruct +elim (fpr_fwd_shift_bind_minus … HT12) #HV12 #H destruct /2 width=5/ +qed-. + +lemma fpr_inv_pair3: ∀I,L1,K2,V2,T1,T2. ⦃L1, T1⦄ ➡ ⦃K2.ⓑ{I}V2, T2⦄ → + ∃∃K1,V1. ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ & + ⦃K1, -ⓑ{I}V1.T1⦄ ➡ ⦃K2, -ⓑ{I}V2.T2⦄ & + L1 = K1.ⓑ{I}V1. +#I2 #X #K2 #V2 #T1 #T2 #H +elim (fpr_fwd_pair3 … H) -H #I1 #K1 #V1 #HT12 #H destruct +elim (fpr_fwd_shift_bind_minus … HT12) #HV12 #H destruct /2 width=5/ +qed-. + +(* More advanced forward lemmas *********************************************) + +lemma fpr_fwd_pair1_full: ∀I,K1,L2,V1,T1,T2. ⦃K1.ⓑ{I}V1, T1⦄ ➡ ⦃L2, T2⦄ → + ∀b. ∃∃K2,V2. ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ & + ⦃K1, ⓑ{b,I}V1.T1⦄ ➡ ⦃K2, ⓑ{b,I}V2.T2⦄ & + L2 = K2.ⓑ{I}V2. +#I #K1 #L2 #V1 #T1 #T2 #H #b +elim (fpr_inv_pair1 … H) -H #K2 #V2 #HV12 #HT12 #H destruct +elim (fpr_fwd_bind2_minus … HT12 b) -HT12 #W1 #U1 #HTU1 #H destruct /2 width=5/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr_fpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr_fpr.etc new file mode 100644 index 000000000..3f7ac2ceb --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/fpr_fpr.etc @@ -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/reducibility/tpr_tpr.ma". +include "basic_2/reducibility/fpr.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON CLOSURES ******************************) + +(* Main properties **********************************************************) + +theorem fpr_conf: bi_confluent … fpr. +#L0 #L1 #T0 #T1 * #HL01 #HT01 #L2 #T2 * >HL01 #HL12 #HT02 +elim (tpr_conf … HT01 HT02) -L0 -T0 #X #H1 #H2 +elim (tpr_fwd_shift1 … H1) #L #T #HL1 #H destruct /3 width=5/ +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/lfpr_alt.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/lfpr_alt.etc new file mode 100644 index 000000000..95ec60ba2 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/lfpr_alt.etc @@ -0,0 +1,87 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +notation "hvbox( L1 ⊢ ➡ ➡ break term 46 L2 )" + non associative with precedence 45 + for @{ 'PRedSnAlt $L1 $L2 }. + +notation "hvbox( ⦃ term 46 L1 ⦄ ➡ ➡ break ⦃ term 46 L2 ⦄ )" + non associative with precedence 45 + for @{ 'FocalizedPRedAlt $L1 $L2 }. + +include "basic_2/grammar/lenv_px_bi.ma". +include "basic_2/reducibility/fpr_cpr.ma". +include "basic_2/reducibility/lfpr_fpr.ma". + +(* FOCALIZED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS **********************) + +(* alternative definition *) +definition lfpra: relation lenv ≝ lpx_bi fpr. + +interpretation + "focalized parallel reduction (environment) alternative" + 'FocalizedPRedAlt L1 L2 = (lfpra L1 L2). + +(* Basic properties *********************************************************) + +lemma lfpra_refl: reflexive … lfpra. +/2 width=1/ qed. + +lemma fpr_lfpra: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ⦃L1⦄ ➡➡ ⦃L2⦄. +#L1 elim L1 -L1 +[ #L2 #T1 #T2 #H + elim (fpr_inv_atom1 … H) -H #_ #H destruct // +| #L1 #I #V1 #IH #L2 #T1 #T2 #H + elim (fpr_inv_pair1 … H) -H #L #V #HV1 #HL1 #H destruct /3 width=3/ +] +qed. + +(* Basic inversion lemmas ***************************************************) + +lemma lfpra_inv_atom1: ∀L2. ⦃⋆⦄ ➡➡ ⦃L2⦄ → L2 = ⋆. +/2 width=2 by lpx_bi_inv_atom1/ qed-. + +lemma lfpra_inv_pair1: ∀K1,I,V1,L2. ⦃K1. ⓑ{I} V1⦄ ➡➡ ⦃L2⦄ → + ∃∃K2,V2. ⦃K1⦄ ➡➡ ⦃K2⦄ & ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ & + L2 = K2. ⓑ{I} V2. +/2 width=1 by lpx_bi_inv_pair1/ qed-. + +lemma lfpra_inv_atom2: ∀L1. ⦃L1⦄ ➡➡ ⦃⋆⦄ → L1 = ⋆. +/2 width=2 by lpx_bi_inv_atom2/ qed-. + +lemma lfpra_inv_pair2: ∀L1,K2,I,V2. ⦃L1⦄ ➡➡ ⦃K2. ⓑ{I} V2⦄ → + ∃∃K1,V1. ⦃K1⦄ ➡➡ ⦃K2⦄ & ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ & + L1 = K1. ⓑ{I} V1. +/2 width=1 by lpx_bi_inv_pair2/ qed-. + +lemma lfpra_inv_fpr: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → ∀T.⦃L1, T⦄ ➡ ⦃L2, T⦄. +#L1 #L2 * -L1 -L2 // /3 width=1/ +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma lfpra_fwd_length: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → |L1| = |L2|. +/2 width=2 by lpx_bi_fwd_length/ qed-. + +(* Main properties **********************************************************) + +theorem lfpr_lfpra: ∀L1,L2. ⦃L1⦄ ➡ ⦃L2⦄ → ⦃L1⦄ ➡➡ ⦃L2⦄. +#L1 #L2 #H +lapply (lfpr_inv_fpr … H (⋆0)) -H /2 width=3/ +qed. + +theorem lfpra_lfpr: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → ⦃L1⦄ ➡ ⦃L2⦄. +#L1 #L2 #H +lapply (lfpra_inv_fpr … H (⋆0)) -H /2 width=3/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/lfpr_fpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/lfpr_fpr.etc new file mode 100644 index 000000000..f798d566b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/fpr/lfpr_fpr.etc @@ -0,0 +1,31 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/lfpr.ma". +include "basic_2/reducibility/cfpr_cpr.ma". + +(* FOCALIZED PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ***********************) + +(* Properties on context-free parallel reduction for closures ***************) + +lemma fpr_lfpr: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ⦃L1⦄ ➡ ⦃L2⦄. +#L1 #L2 #T1 #T2 #H +elim (fpr_inv_all … H) -H /2 width=3/ +qed. + +(* Inversion lemmas on context-free parallel reduction for closures *********) + +lemma lfpr_inv_fpr: ∀L1,L2. ⦃L1⦄ ➡ ⦃L2⦄ → ∀T. ⦃L1, T⦄ ➡ ⦃L2, T⦄. +#L1 #L2 * /2 width=4/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr.ma deleted file mode 100644 index e1e3c4386..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr.ma +++ /dev/null @@ -1,55 +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/reducibility/cpr.ma". -include "basic_2/reducibility/fpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) - -definition cfpr: lenv → bi_relation lenv term ≝ - λL,L1,T1,L2,T2. |L1| = |L2| ∧ L ⊢ L1 @@ T1 ➡ L2 @@ T2. - -interpretation - "context-sensitive parallel reduction (closure)" - 'FocalizedPRed L L1 T1 L2 T2 = (cfpr L L1 T1 L2 T2). - -(* Basic properties *********************************************************) - -lemma cfpr_refl: ∀L. bi_reflexive … (cfpr L). -/2 width=1/ qed. - -lemma fpr_cfpr: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ⋆ ⊢ ⦃L1, T1⦄ ➡ ⦃L2, T2⦄. -#L1 #L2 #T1 #T2 * /3 width=1/ -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma cfpr_inv_atom1: ∀L,L2,T1,T2. L ⊢ ⦃⋆, T1⦄ ➡ ⦃L2, T2⦄ → L ⊢ T1 ➡ T2 ∧ L2 = ⋆. -#L #L2 #T1 #T2 * #H >(length_inv_zero_sn … H) /2 width=1/ -qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma fpr_inv_pair1_sn: ∀I,K1,L2,V1,T1,T2. ⦃⋆.ⓑ{I}V1@@K1, T1⦄ ➡ ⦃L2, T2⦄ → - ∃∃K2,V2. V1 ➡ V2 & - ⋆.ⓑ{I}V2 ⊢ ⦃K1, T1⦄ ➡ ⦃K2, T2⦄ & - L2 = ⋆.ⓑ{I}V2@@K2. -#I1 #K1 #L2 #V1 #T1 #T2 * >append_length #H -elim (length_inv_pos_sn_append … H) -H #I2 #K2 #V2 #HK12 #H destruct ->shift_append_assoc >shift_append_assoc normalize in ⊢ (%→?); #H -elim (tpr_inv_bind1 … H) -H * -[ #V0 #T #T0 #HV10 #HT1 #HT0 #H destruct /5 width=5/ -| #T0 #_ #_ #H destruct -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_aaa.ma deleted file mode 100644 index 72802e499..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_aaa.ma +++ /dev/null @@ -1,27 +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/reducibility/cpr_aaa.ma". -include "basic_2/reducibility/cfpr_cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) - -(* Properties about atomic arity assignment on terms ************************) - -lemma aaa_fpr_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → - ∀L2,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → L2 ⊢ T2 ⁝ A. -#L1 #T1 #A #HT1 #L2 #T2 #H -elim (fpr_inv_all … H) -H -/4 width=5 by aaa_cpr_conf, aaa_ltpr_conf, aaa_ltpss_sn_conf/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_cfpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_cfpr.ma deleted file mode 100644 index f442be28b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_cfpr.ma +++ /dev/null @@ -1,26 +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/reducibility/cpr_cpr.ma". -include "basic_2/reducibility/cfpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) - -(* Main properties **********************************************************) - -theorem cfpr_conf: ∀L. bi_confluent … (cfpr L). -#L #L0 #L1 #T0 #T1 * #HL01 #HT01 #L2 #T2 * >HL01 #HL12 #HT02 -elim (cpr_conf … HT01 HT02) -L0 -T0 #X #H1 #H2 -elim (cpr_fwd_shift1 … H1) #L0 #T0 #HL10 #H destruct /3 width=5/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_cpr.ma deleted file mode 100644 index a1e20a025..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_cpr.ma +++ /dev/null @@ -1,64 +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/reducibility/cpr_tpss.ma". -include "basic_2/reducibility/cpr_cpr.ma". -include "basic_2/reducibility/cfpr_ltpss.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) - -(* Advanced properties ******************************************************) - -lemma fpr_all: ∀L1,L. L1 ➡ L → ∀L2,T1,T2. L ⊢ T1 ➡ T2 → - L ⊢ ▶* [0, |L|] L2 → ⦃L1, T1⦄ ➡ ⦃L2, T2⦄. -#L1 #L #H elim H -L1 -L -[ #L2 #T1 #T2 #HT12 #HL2 - lapply (ltpss_sn_inv_atom1 … HL2) -HL2 #H destruct - lapply (cpr_inv_atom … HT12) -HT12 /2 width=1/ -| #I #L1 #L #V1 #V #_ #HV1 #IH #X #T1 #T2 #HT12 #H - elim (ltpss_sn_inv_tpss21 … H ?) -H // append_length >append_length #H - lapply (injective_plus_r … H) -H #H - @(ex3_intro … (⋆.ⓑ{I}V@@Y)) append_length H -H >commutative_plus /3 width=1/ -] -qed-. - -lemma fpr_inv_all: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → - ∃∃L. L1 ➡ L & L ⊢ T1 ➡ T2 & L ⊢ ▶* [0, |L|] L2. -#L1 #L2 #T1 #T2 #H -lapply (fpr_cfpr … H) -H #H -elim (cfpr_inv_all … H) -H /2 width=4/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_ltpss.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_ltpss.ma deleted file mode 100644 index 66e102e63..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cfpr_ltpss.ma +++ /dev/null @@ -1,39 +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/reducibility/cpr_lift.ma". -include "basic_2/reducibility/cpr_ltpss_sn.ma". -include "basic_2/reducibility/cfpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON CLOSURES *************************) - -(* Advanced inversion lemmas ************************************************) - -lemma cfpr_inv_pair1: ∀I,L,K1,L2,V1,T1,T2. L ⊢ ⦃⋆.ⓑ{I}V1@@K1, T1⦄ ➡ ⦃L2, T2⦄ → - ∃∃K2,V,V2. V1 ➡ V & L ⊢ V ▶* [0, |L|] V2 & - L.ⓑ{I}V ⊢ ⦃K1, T1⦄ ➡ ⦃K2, T2⦄ & - L2 = ⋆.ⓑ{I}V2@@K2. -* #L #K1 #L2 #V1 #T1 #T2 * >append_length #H -elim (length_inv_pos_sn_append … H) -H #I2 #K2 #V2 #HK12 #H destruct ->shift_append_assoc >shift_append_assoc normalize in ⊢ (??%%→?); #H -[ elim (cpr_inv_abbr1 … H) -H * - [ #V #V0 #T0 #HV1 #HV0 #HT10 #H destruct /3 width=7/ - | #T0 #_ #_ #H destruct - ] -| elim (cpr_inv_abst1 … H Abst V2) -H - #V #T * #V0 #HV10 #HV0 #HT1 #H destruct - lapply (ltpss_sn_cpr_trans (L.ⓛV0) … 0 (|L|+1) … HT1) -HT1 /2 width=1/ #HT12 - /3 width=7/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cif.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cif.ma deleted file mode 100644 index 0a536a914..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cif.ma +++ /dev/null @@ -1,71 +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/reducibility/crf.ma". - -(* CONTEXT-SENSITIVE IRREDUCIBLE TERMS **************************************) - -definition cif: lenv → predicate term ≝ λL,T. L ⊢ 𝐑⦃T⦄ → ⊥. - -interpretation "context-sensitive irreducibility (term)" - 'NotReducible L T = (cif L T). - -(* Basic inversion lemmas ***************************************************) - -lemma cif_inv_delta: ∀L,K,V,i. ⇩[0, i] L ≡ K.ⓓV → L ⊢ 𝐈⦃#i⦄ → ⊥. -/3 width=3/ qed-. - -lemma cif_inv_ri2: ∀I,L,V,T. ri2 I → L ⊢ 𝐈⦃②{I}V.T⦄ → ⊥. -/3 width=1/ qed-. - -lemma cif_inv_ib2: ∀a,I,L,V,T. ib2 a I → L ⊢ 𝐈⦃ⓑ{a,I}V.T⦄ → - L ⊢ 𝐈⦃V⦄ ∧ L.ⓑ{I}V ⊢ 𝐈⦃T⦄. -/4 width=1/ qed-. - -lemma cif_inv_bind: ∀a,I,L,V,T. L ⊢ 𝐈⦃ⓑ{a,I}V.T⦄ → - ∧∧ L ⊢ 𝐈⦃V⦄ & L.ⓑ{I}V ⊢ 𝐈⦃T⦄ & ib2 a I. -#a * [ elim a -a ] -[ #L #V #T #H elim (H ?) -H /3 width=1/ -|*: #L #V #T #H elim (cif_inv_ib2 … H) -H /2 width=1/ /3 width=1/ -] -qed-. - -lemma cif_inv_appl: ∀L,V,T. L ⊢ 𝐈⦃ⓐV.T⦄ → ∧∧ L ⊢ 𝐈⦃V⦄ & L ⊢ 𝐈⦃T⦄ & 𝐒⦃T⦄. -#L #V #T #HVT @and3_intro /3 width=1/ -generalize in match HVT; -HVT elim T -T // -* // #a * #U #T #_ #_ #H elim (H ?) -H /2 width=1/ -qed-. - -lemma cif_inv_flat: ∀I,L,V,T. L ⊢ 𝐈⦃ⓕ{I}V.T⦄ → - ∧∧ L ⊢ 𝐈⦃V⦄ & L ⊢ 𝐈⦃T⦄ & 𝐒⦃T⦄ & I = Appl. -* #L #V #T #H -[ elim (cif_inv_appl … H) -H /2 width=1/ -| elim (cif_inv_ri2 … H) -H /2 width=1/ -] -qed-. - -(* Basic properties *********************************************************) - -lemma tif_atom: ∀I. ⋆ ⊢ 𝐈⦃⓪{I}⦄. -/2 width=2 by trf_inv_atom/ qed. - -lemma cif_ib2: ∀a,I,L,V,T. ib2 a I → L ⊢ 𝐈⦃V⦄ → L.ⓑ{I}V ⊢ 𝐈⦃T⦄ → L ⊢ 𝐈⦃ⓑ{a,I}V.T⦄. -#a #I #L #V #T #HI #HV #HT #H -elim (crf_inv_ib2 … HI H) -HI -H /2 width=1/ -qed. - -lemma cif_appl: ∀L,V,T. L ⊢ 𝐈⦃V⦄ → L ⊢ 𝐈⦃T⦄ → 𝐒⦃T⦄ → L ⊢ 𝐈⦃ⓐV.T⦄. -#L #V #T #HV #HT #H1 #H2 -elim (crf_inv_appl … H2) -H2 /2 width=1/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cif_append.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cif_append.ma deleted file mode 100644 index 45fd178cf..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cif_append.ma +++ /dev/null @@ -1,34 +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/reducibility/crf_append.ma". -include "basic_2/reducibility/cif.ma". - -(* CONTEXT-SENSITIVE IRREDUCIBLE TERMS **************************************) - -(* Advanved properties ******************************************************) - -lemma cif_labst_last: ∀L,T,W. L ⊢ 𝐈⦃T⦄ → ⋆.ⓛW @@ L ⊢ 𝐈⦃T⦄. -/3 width=2 by crf_inv_labst_last/ qed. - -lemma cif_tif: ∀T,W. ⋆ ⊢ 𝐈⦃T⦄ → ⋆.ⓛW ⊢ 𝐈⦃T⦄. -/3 width=2 by crf_inv_trf/ qed. - -(* Advanced inversion lemmas ************************************************) - -lemma cif_inv_labst_last: ∀L,T,W. ⋆.ⓛW @@ L ⊢ 𝐈⦃T⦄ → L ⊢ 𝐈⦃T⦄. -/3 width=1/ qed-. - -lemma cif_inv_tif: ∀T,W. ⋆.ⓛW ⊢ 𝐈⦃T⦄ → ⋆ ⊢ 𝐈⦃T⦄. -/3 width=1/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf.ma deleted file mode 100644 index 02bbcf87a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf.ma +++ /dev/null @@ -1,66 +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/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) - -definition cnf: lenv → predicate term ≝ λL. NF … (cpr L) (eq …). - -interpretation - "context-sensitive normality (term)" - 'Normal L T = (cnf L T). - -(* Basic inversion lemmas ***************************************************) - -lemma cnf_inv_appl: ∀L,V,T. L ⊢ 𝐍⦃ⓐV.T⦄ → ∧∧ L ⊢ 𝐍⦃V⦄ & L ⊢ 𝐍⦃T⦄ & 𝐒⦃T⦄. -#L #V1 #T1 #HVT1 @and3_intro -[ #V2 #HV2 lapply (HVT1 (ⓐV2.T1) ?) -HVT1 /2 width=1/ -HV2 #H destruct // -| #T2 #HT2 lapply (HVT1 (ⓐV1.T2) ?) -HVT1 /2 width=1/ -HT2 #H destruct // -| generalize in match HVT1; -HVT1 elim T1 -T1 * // #a * #W1 #U1 #_ #_ #H - [ elim (lift_total V1 0 1) #V2 #HV12 - lapply (H (ⓓ{a}W1.ⓐV2.U1) ?) -H /3 width=3/ -HV12 #H destruct - | lapply (H (ⓓ{a}V1.U1) ?) -H /3 width=1/ #H destruct -] -qed-. - -lemma cnf_inv_zeta: ∀L,V,T. L ⊢ 𝐍⦃+ⓓV.T⦄ → ⊥. -#L #V #T #H elim (is_lift_dec T 0 1) -[ * #U #HTU - lapply (H U ?) -H /3 width=3 by cpr_tpr, tpr_zeta/ #H destruct (**) (* auto too slow without trace *) - elim (lift_inv_pair_xy_y … HTU) -| #HT - elim (tps_full (⋆) V T (⋆. ⓓV) 0 ?) // #T2 #T1 #HT2 #HT12 - lapply (H (+ⓓV.T2) ?) -H /3 width=3 by cpr_tpr, tpr_delta/ -HT2 #H destruct /3 width=2/ (**) (* auto too slow without trace *) -] -qed. - -lemma cnf_inv_tau: ∀L,V,T. L ⊢ 𝐍⦃ⓝV.T⦄ → ⊥. -#L #V #T #H lapply (H T ?) -H /2 width=1/ #H -@discr_tpair_xy_y // -qed-. - -(* Basic properties *********************************************************) - -(* Basic_1: was: nf2_sort *) -lemma cnf_sort: ∀L,k. L ⊢ 𝐍⦃⋆k⦄. -#L #k #X #H ->(cpr_inv_sort1 … H) // -qed. - -(* Basic_1: was: nf2_dec *) -axiom cnf_dec: ∀L,T1. L ⊢ 𝐍⦃T1⦄ ∨ - ∃∃T2. L ⊢ T1 ➡ T2 & (T1 = T2 → ⊥). - -(* Basic_1: removed theorems 1: nf2_abst_shift *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf_cif.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf_cif.ma deleted file mode 100644 index 5a1ae5a5b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf_cif.ma +++ /dev/null @@ -1,106 +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/reducibility/cif.ma". -include "basic_2/reducibility/cnf_lift.ma". - -(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) - -(* Main properties **********************************************************) - -lemma tps_cif_eq: ∀L,T1,T2,d,e. L ⊢ T1 ▶[d, e] T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e -[ // -| #L #K #V #W #i #d #e #_ #_ #HLK #_ #H -d -e - elim (cif_inv_delta … HLK ?) // -| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H - elim (cif_inv_bind … H) -H #HV1 #HT1 * #H destruct - lapply (IHV12 … HV1) -IHV12 -HV1 #H destruct /3 width=1/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #H - elim (cif_inv_flat … H) -H #HV1 #HT1 #_ #_ /3 width=1/ -] -qed. - -lemma tpss_cif_eq: ∀L,T1,T2,d,e. L ⊢ T1 ▶*[d, e] T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2. -#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT1 #HT1 -lapply (IHT1 HT1) -IHT1 #H destruct /2 width=5/ -qed. - -lemma tpr_cif_eq: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ 𝐈⦃T1⦄ → T1 = T2. -#T1 #T2 #H elim H -T1 -T2 -[ // -| * #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #L #H - [ elim (cif_inv_appl … H) -H #HV1 #HT1 #_ - >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 // - | elim (cif_inv_ri2 … H) /2 width=1/ - ] -| #a #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #L #H - elim (cif_inv_appl … H) -H #_ #_ #H - elim (simple_inv_bind … H) -| #a * #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV1 #IHT1 #L #H - [ lapply (tps_lsubr_trans … HT2 (L.ⓓV2) ?) -HT2 /2 width=1/ #HT2 - elim (cif_inv_bind … H) -H #HV1 #HT1 * #H destruct - lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct - lapply (IHT1 … HT1) -IHT1 #H destruct - lapply (tps_cif_eq … HT2 ?) -HT2 // - | <(tps_inv_refl_SO2 … HT2 ?) -HT2 // - elim (cif_inv_ib2 … H) -H /2 width=1/ /3 width=2/ - ] -| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #L #H - elim (cif_inv_appl … H) -H #_ #_ #H - elim (simple_inv_bind … H) -| #V1 #T1 #T #T2 #_ #_ #_ #L #H - elim (cif_inv_ri2 … H) /2 width=1/ -| #V1 #T1 #T2 #_ #_ #L #H - elim (cif_inv_ri2 … H) /2 width=1/ -] -qed. - -lemma cpr_cif_eq: ∀L,T1,T2. L ⊢ T1 ➡ T2 → L ⊢ 𝐈⦃T1⦄ → T1 = T2. -#L #T1 #T2 * #T0 #HT10 #HT02 #HT1 -lapply (tpr_cif_eq … HT10 … HT1) -HT10 #H destruct /2 width=5/ -qed. - -theorem cif_cnf: ∀L,T. L ⊢ 𝐈⦃T⦄ → L ⊢ 𝐍⦃T⦄. -/3 width=3/ qed. - -(* Note: this property is unusual *) -lemma cnf_crf_false: ∀L,T. L ⊢ 𝐑⦃T⦄ → L ⊢ 𝐍⦃T⦄ → ⊥. -#L #T #H elim H -L -T -[ #L #K #V #i #HLK #H - elim (cnf_inv_delta … HLK H) -| #L #V #T #_ #IHV #H - elim (cnf_inv_appl … H) -H /2 width=1/ -| #L #V #T #_ #IHT #H - elim (cnf_inv_appl … H) -H /2 width=1/ -| #I #L #V #T * #H1 #H2 destruct - [ elim (cnf_inv_zeta … H2) - | elim (cnf_inv_tau … H2) - ] -|5,6: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct - [1,3: elim (cnf_inv_abbr … H2) -H2 /2 width=1/ - |*: elim (cnf_inv_abst … H2) -H2 /2 width=1/ - ] -| #a #L #V #W #T #H - elim (cnf_inv_appl … H) -H #_ #_ #H - elim (simple_inv_bind … H) -| #a #L #V #W #T #H - elim (cnf_inv_appl … H) -H #_ #_ #H - elim (simple_inv_bind … H) -] -qed. - -theorem cnf_cif: ∀L,T. L ⊢ 𝐍⦃T⦄ → L ⊢ 𝐈⦃T⦄. -/2 width=4/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf_lift.ma deleted file mode 100644 index 0e1a8551f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cnf_lift.ma +++ /dev/null @@ -1,85 +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/reducibility/cpr_lift.ma". -include "basic_2/reducibility/cpr_cpr.ma". -include "basic_2/reducibility/cnf.ma". - -(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) - -(* Advanced inversion lemmas ************************************************) - -lemma cnf_inv_delta: ∀L,K,V,i. ⇩[0, i] L ≡ K.ⓓV → L ⊢ 𝐍⦃#i⦄ → ⊥. -#L #K #V #i #HLK #H -elim (lift_total V 0 (i+1)) #W #HVW -lapply (H W ?) -H [ /3 width=6/ ] -HLK #H destruct -elim (lift_inv_lref2_be … HVW ? ?) -HVW // -qed-. - -lemma cnf_inv_abst: ∀a,L,V,T. L ⊢ 𝐍⦃ⓛ{a}V.T⦄ → L ⊢ 𝐍⦃V⦄ ∧ L.ⓛV ⊢ 𝐍⦃T⦄. -#a #L #V1 #T1 #HVT1 @conj -[ #V2 #HV2 lapply (HVT1 (ⓛ{a}V2.T1) ?) -HVT1 /2 width=2/ -HV2 #H destruct // -| #T2 #HT2 lapply (HVT1 (ⓛ{a}V1.T2) ?) -HVT1 /2 width=2/ -HT2 #H destruct // -] -qed-. - -lemma cnf_inv_abbr: ∀L,V,T. L ⊢ 𝐍⦃-ⓓV.T⦄ → L ⊢ 𝐍⦃V⦄ ∧ L.ⓓV ⊢ 𝐍⦃T⦄. -#L #V1 #T1 #HVT1 @conj -[ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2/ -HV2 #H destruct // -| #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2/ -HT2 #H destruct // -] -qed-. - -(* Advanced properties ******************************************************) - -(* Basic_1: was only: nf2_csort_lref *) -lemma cnf_lref_atom: ∀L,i. ⇩[0, i] L ≡ ⋆ → L ⊢ 𝐍⦃#i⦄. -#L #i #HLK #X #H -elim (cpr_inv_lref1 … H) // * -#K0 #V0 #V1 #HLK0 #_ #_ #_ -lapply (ldrop_mono … HLK … HLK0) -L #H destruct -qed. - -(* Basic_1: was: nf2_lref_abst *) -lemma cnf_lref_abst: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓛV → L ⊢ 𝐍⦃#i⦄. -#L #K #V #i #HLK #X #H -elim (cpr_inv_lref1 … H) // * -#K0 #V0 #V1 #HLK0 #_ #_ #_ -lapply (ldrop_mono … HLK … HLK0) -L #H destruct -qed. - -(* Basic_1: was: nf2_abst *) -lemma cnf_abst: ∀a,I,L,V,W,T. L ⊢ 𝐍⦃W⦄ → L. ⓑ{I} V ⊢ 𝐍⦃T⦄ → L ⊢ 𝐍⦃ⓛ{a}W.T⦄. -#a #I #L #V #W #T #HW #HT #X #H -elim (cpr_inv_abst1 … H I V) -H #W0 #T0 #HW0 #HT0 #H destruct ->(HW … HW0) -W0 >(HT … HT0) -T0 // -qed. - -(* Basic_1: was only: nf2_appl_lref *) -lemma cnf_appl_simple: ∀L,V,T. L ⊢ 𝐍⦃V⦄ → L ⊢ 𝐍⦃T⦄ → 𝐒⦃T⦄ → L ⊢ 𝐍⦃ⓐV.T⦄. -#L #V #T #HV #HT #HS #X #H -elim (cpr_inv_appl1_simple … H ?) -H // #V0 #T0 #HV0 #HT0 #H destruct ->(HV … HV0) -V0 >(HT … HT0) -T0 // -qed. - -(* Relocation properties ****************************************************) - -(* Basic_1: was: nf2_lift *) -lemma cnf_lift: ∀L0,L,T,T0,d,e. - L ⊢ 𝐍⦃T⦄ → ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → L0 ⊢ 𝐍⦃T0⦄. -#L0 #L #T #T0 #d #e #HLT #HL0 #HT0 #X #H -elim (cpr_inv_lift1 … HL0 … HT0 … H) -L0 #T1 #HT10 #HT1 -<(HLT … HT1) in HT0; -L #HT0 ->(lift_mono … HT10 … HT0) -T1 -X // -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr.ma deleted file mode 100644 index d82299c38..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr.ma +++ /dev/null @@ -1,111 +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/unfold/tpss.ma". -include "basic_2/reducibility/tpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Basic_1: includes: pr2_delta1 *) -definition cpr: lenv → relation term ≝ - λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T ▶* [0, |L|] T2. - -interpretation - "context-sensitive parallel reduction (term)" - 'PRed L T1 T2 = (cpr L T1 T2). - -(* Basic properties *********************************************************) - -lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ➡ T2. -/3 width=5/ qed-. - -(* Basic_1: was by definition: pr2_free *) -lemma cpr_tpr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2. -/2 width=3/ qed. - -lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ➡ T2. -/3 width=5/ qed. - -lemma cpr_refl: ∀L,T. L ⊢ T ➡ T. -/2 width=1/ qed. - -(* Note: new property *) -(* Basic_1: was only: pr2_thin_dx *) -lemma cpr_flat: ∀I,L,V1,V2,T1,T2. - L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2. -#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/ -qed. - -lemma cpr_cast: ∀L,V,T1,T2. - L ⊢ T1 ➡ T2 → L ⊢ ⓝV. T1 ➡ T2. -#L #V #T1 #T2 * /3 width=3/ -qed. - -(* Note: it does not hold replacing |L1| with |L2| *) -(* Basic_1: was only: pr2_change *) -lemma cpr_lsubr_trans: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 → - ∀L2. L2 ⊑ [0, |L1|] L1 → L2 ⊢ T1 ➡ T2. -#L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 -lapply (tpss_lsubr_trans … HT2 … HL12) -HT2 -HL12 /3 width=4/ -qed. - -(* Basic inversion lemmas ***************************************************) - -(* Basic_1: was: pr2_gen_csort *) -lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → T1 ➡ T2. -#T1 #T2 * #T #HT normalize #HT2 -<(tpss_inv_refl_O2 … HT2) -HT2 // -qed-. - -(* Basic_1: was: pr2_gen_sort *) -lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k. -#L #T2 #k * #X #H ->(tpr_inv_atom1 … H) -H #H ->(tpss_inv_sort1 … H) -H // -qed-. - -(* Basic_1: was: pr2_gen_cast *) -lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓝV1. T1 ➡ U2 → ( - ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 & - U2 = ⓝV2. T2 - ) ∨ L ⊢ T1 ➡ U2. -#L #V1 #T1 #U2 * #X #H #HU2 -elim (tpr_inv_cast1 … H) -H /3 width=3/ -* #V #T #HV1 #HT1 #H destruct -elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/ -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma cpr_fwd_bind1_minus: ∀I,L,V1,T1,T. L ⊢ -ⓑ{I}V1.T1 ➡ T → ∀b. - ∃∃V2,T2. L ⊢ ⓑ{b,I}V1.T1 ➡ ⓑ{b,I}V2.T2 & - T = -ⓑ{I}V2.T2. -#I #L #V1 #T1 #T * #X #H1 #H2 #b -elim (tpr_fwd_bind1_minus … H1 b) -H1 #V0 #T0 #HT10 #H destruct -elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV02 #HT02 #H destruct /4 width=5/ -qed-. - -lemma cpr_fwd_shift1: ∀L,L1,T1,T. L ⊢ L1 @@ T1 ➡ T → - ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. -#L #L1 #T1 #T * #X #H1 #H2 -elim (tpr_fwd_shift1 … H1) -H1 #L0 #T0 #HL10 #H destruct -elim (tpss_fwd_shift1 … H2) -H2 #L2 #T2 #HL02 #H destruct /2 width=4/ -qed-. - -(* Basic_1: removed theorems 6: - pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans - pr2_gen_ctail pr2_ctail - Basic_1: removed local theorems 3: - pr2_free_free pr2_free_delta pr2_delta_delta -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_aaa.ma deleted file mode 100644 index 7c1273334..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_aaa.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/static/aaa_ltpss_sn.ma". -include "basic_2/reducibility/ltpr_aaa.ma". -include "basic_2/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Properties about atomic arity assignment on terms ************************) - -lemma aaa_cpr_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → ∀T2. L ⊢ T1 ➡ T2 → L ⊢ T2 ⁝ A. -#L #T1 #A #HT1 #T2 * /3 width=5/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_cpr.ma deleted file mode 100644 index d231c849f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_cpr.ma +++ /dev/null @@ -1,63 +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/reducibility/tpr_tpr.ma". -include "basic_2/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Advanced properties ******************************************************) - -lemma cpr_bind_sn: ∀a,I,L,V1,V2,T1,T2. L ⊢ V1 ➡ V2 → T1 ➡ T2 → - L ⊢ ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2. -#a #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 #HT12 -@ex2_intro [2: @(tpr_delta … HV1 HT12) | skip ] /2 width=3/ (* /3 width=5/ is too slow *) -qed. - -(* Basic_1: was only: pr2_gen_cbind *) -lemma cpr_bind_dx: ∀a,I,L,V1,V2,T1,T2. V1 ➡ V2 → L. ⓑ{I} V2 ⊢ T1 ➡ T2 → - L ⊢ ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2. -#a #I #L #V1 #V2 #T1 #T2 #HV12 * #T #HT1 normalize #HT2 -elim (tpss_split_up … HT2 1 ? ?) -HT2 // #T0 (tpr_inv_atom1 … H) -H #H -elim (tpss_inv_lref1 … H) -H /2 width=1/ -* /3 width=6/ -qed-. - -(* Basic_1: was pr2_gen_abbr *) -lemma cpr_inv_abbr1: ∀a,L,V1,T1,U2. L ⊢ ⓓ{a}V1. T1 ➡ U2 → - (∃∃V,V2,T2. V1 ➡ V & L ⊢ V ▶* [O, |L|] V2 & - L. ⓓV ⊢ T1 ➡ T2 & - U2 = ⓓ{a}V2. T2 - ) ∨ - ∃∃T2. L.ⓓV1 ⊢ T1 ➡ T2 & ⇧[0,1] U2 ≡ T2 & a = true. -#a #L #V1 #T1 #Y * #X #H1 #H2 -elim (tpr_inv_abbr1 … H1) -H1 * -[ #V #T #T0 #HV1 #HT1 #HT0 #H destruct - elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT02 #H destruct - lapply (tps_lsubr_trans … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0 - lapply (tps_weak_full … HT0) -HT0 #HT0 - lapply (tpss_lsubr_trans … HT02 (L. ⓓV) ?) -HT02 /2 width=1/ #HT02 - lapply (tpss_weak_full … HT02) -HT02 #HT02 - lapply (tpss_strap2 … HT0 HT02) -T0 /4 width=7/ -| #T2 #HT12 #HXT2 #H destruct - elim (lift_total Y 0 1) #Z #HYZ - lapply (tpss_lift_ge … H2 (L.ⓓV1) … HXT2 … HYZ) -X // /2 width=1/ #H - lapply (cpr_intro … HT12 … H) -T2 /3 width=3/ -] -qed-. - -(* Basic_1: was: pr2_gen_abst *) -lemma cpr_inv_abst1: ∀a,L,V1,T1,U2. L ⊢ ⓛ{a}V1. T1 ➡ U2 → ∀I,W. - ∃∃V2,T2. L ⊢ V1 ➡ V2 & L. ⓑ{I} W ⊢ T1 ➡ T2 & U2 = ⓛ{a}V2. T2. -#a #L #V1 #T1 #Y * #X #H1 #H2 #I #W -elim (tpr_inv_abst1 … H1) -H1 #V #T #HV1 #HT1 #H destruct -elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct -lapply (tpss_lsubr_trans … HT2 (L. ⓑ{I} W) ?) -HT2 /2 width=1/ /4 width=5/ -qed-. - -(* Basic_1: was pr2_gen_appl *) -lemma cpr_inv_appl1: ∀L,V1,U0,U2. L ⊢ ⓐV1. U0 ➡ U2 → - ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U0 ➡ T2 & - U2 = ⓐV2. T2 - | ∃∃a,V2,W,T1,T2. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 & - U0 = ⓛ{a}W. T1 & - U2 = ⓓ{a}V2. T2 - | ∃∃a,V2,V,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & L. ⓓW2 ⊢ T1 ➡ T2 & - ⇧[0,1] V2 ≡ V & - U0 = ⓓ{a}W1. T1 & - U2 = ⓓ{a}W2. ⓐV. T2. -#L #V1 #U0 #Y * #X #H1 #H2 -elim (tpr_inv_appl1 … H1) -H1 * -[ #V #U #HV1 #HU0 #H destruct - elim (tpss_inv_flat1 … H2) -H2 #V2 #U2 #HV2 #HU2 #H destruct /4 width=5/ -| #a #V #W #T0 #T #HV1 #HT0 #H #H1 destruct - elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct - lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=9/ -| #a #V0 #V #W #W0 #T #T0 #HV10 #HW0 #HT0 #HV0 #H #H1 destruct - elim (tpss_inv_bind1 … H2) -H2 #W2 #X #HW02 #HX #HY destruct - elim (tpss_inv_flat1 … HX) -HX #V2 #T2 #HV2 #HT2 #H destruct - elim (tpss_inv_lift1_ge … HV2 … HV0 ?) -V // [3: /2 width=1/ |2: skip ] #V (ltpr_fwd_length … HL12) in HT2; #HT2 -elim (ltpr_tpr_tpss_conf … HL12 … HT2) -L1 /3 width=3/ -qed. - -lemma cpr_ltpr_conf_tpss: ∀L1,L2. L1 ➡ L2 → ∀T1,T2. L1 ⊢ T1 ➡ T2 → - ∀d,e,U1. L1 ⊢ T1 ▶* [d, e] U1 → - ∃∃U2. L2 ⊢ U1 ➡ U2 & L2 ⊢ T2 ➡ U2. -#L1 #L2 #HL12 #T1 #T2 #HT12 #d #e #U1 #HTU1 -elim (cpr_ltpr_conf_eq … HT12 … HL12) -HT12 #T #HT1 #HT2 -elim (cpr_tpss_ltpr … HL12 … HT1 … HTU1) -L1 -HT1 #U2 #HU12 #HTU2 -lapply (tpss_weak_full … HTU2) -HTU2 #HTU2 /3 width=5/ (**) (* /4 width=5/ is too slow *) -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_ltpss_dx.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_ltpss_dx.ma deleted file mode 100644 index 5d90e6865..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_ltpss_dx.ma +++ /dev/null @@ -1,28 +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/unfold/ltpss_dx_ltpss_dx.ma". -include "basic_2/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Properties concerning dx partial unfold on local environments ************) - -lemma ltpss_dx_cpr_conf: ∀L1,T,U1. L1 ⊢ T ➡ U1 → - ∀L2,d,e. L1 ▶* [d, e] L2 → - ∃∃U2. L2 ⊢ T ➡ U2 & L2 ⊢ U1 ▶* [d, e] U2. -#L1 #T #U1 * #U #HTU #HU1 #L2 #d #e #HL12 -lapply (ltpss_dx_fwd_length … HL12) #H >H in HU1; -H #HU1 -elim (ltpss_dx_tpss_conf … HU1 … HL12) -L1 /3 width=3/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_ltpss_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_ltpss_sn.ma deleted file mode 100644 index f946e4cfb..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_ltpss_sn.ma +++ /dev/null @@ -1,35 +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/unfold/ltpss_sn_alt.ma". -include "basic_2/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Properties concerning sn partial unfold on local environments ************) - -lemma ltpss_sn_cpr_trans: ∀L1,L2,d,e. L1 ⊢ ▶* [d, e] L2 → - ∀T1,T2. L2 ⊢ T1 ➡ T2 → L1 ⊢ T1 ➡ T2. -#L1 #L2 #d #e #HL12 #T1 #T2 * -lapply (ltpss_sn_weak_full … HL12) -<(ltpss_sn_fwd_length … HL12) -HL12 /3 width=5/ -qed. - -lemma ltpss_sn_cpr_conf: ∀L1,T,U1. L1 ⊢ T ➡ U1 → - ∀L2,d,e. L1 ⊢ ▶* [d, e] L2 → - ∃∃U2. L2 ⊢ T ➡ U2 & L1 ⊢ U1 ▶* [d, e] U2. -#L1 #T #U1 * #U #HTU #HU1 #L2 #d #e #HL12 -lapply (ltpss_sn_fwd_length … HL12) #H >H in HU1; -H #HU1 -elim (ltpss_sn_tpss_conf … HU1 … HL12) -HU1 -HL12 /3 width=3/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_tpss.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_tpss.ma deleted file mode 100644 index d8c7225d9..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_tpss.ma +++ /dev/null @@ -1,39 +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/reducibility/ltpr_tpss.ma". -include "basic_2/reducibility/cpr.ma". - -(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) - -(* Properties on partial unfold for terms ***********************************) - -lemma cpr_tpss_trans: ∀L,T1,T. L ⊢ T1 ➡ T → - ∀T2,d,e. L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ➡ T2. -#L #T1 #T * #T0 #HT10 #HT0 #T2 #d #e #HT2 -lapply (tpss_weak_full … HT2) -HT2 #HT2 -lapply (tpss_trans_eq … HT0 HT2) -T /2 width=3/ -qed. - -lemma cpr_tps_trans: ∀L,T1,T. L ⊢ T1 ➡ T → - ∀T2,d,e. L ⊢ T ▶ [d, e] T2 → L ⊢ T1 ➡ T2. -/3 width=5/ qed. - -lemma cpr_tpss_conf: ∀L,T0,T1. L ⊢ T0 ➡ T1 → - ∀T2,d,e. L ⊢ T0 ▶* [d, e] T2 → - ∃∃T. L ⊢ T1 ▶* [d, e] T & L ⊢ T2 ➡ T. -#L #T0 #T1 * #U0 #HTU0 #HU0T1 #T2 #d #e #HT02 -elim (tpr_tpss_conf … HTU0 … HT02) -T0 #T0 #HT20 #HUT0 -elim (tpss_conf_eq … HU0T1 … HUT0) -U0 /3 width=5/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/crf.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/crf.ma deleted file mode 100644 index 5015f033c..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/crf.ma +++ /dev/null @@ -1,116 +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/substitution/ldrop.ma". - -(* CONTEXT-SENSITIVE REDUCIBLE TERMS ****************************************) - -(* reducible binary items *) -definition ri2: item2 → Prop ≝ - λI. I = Bind2 true Abbr ∨ I = Flat2 Cast. - -(* irreducible binary binders *) -definition ib2: bool → bind2 → Prop ≝ - λa,I. I = Abst ∨ Bind2 a I = Bind2 false Abbr. - -(* reducible terms *) -inductive crf: lenv → predicate term ≝ -| crf_delta : ∀L,K,V,i. ⇩[0, i] L ≡ K.ⓓV → crf L (#i) -| crf_appl_sn: ∀L,V,T. crf L V → crf L (ⓐV. T) -| crf_appl_dx: ∀L,V,T. crf L T → crf L (ⓐV. T) -| crf_ri2 : ∀I,L,V,T. ri2 I → crf L (②{I}V. T) -| crf_ib2_sn : ∀a,I,L,V,T. ib2 a I → crf L V → crf L (ⓑ{a,I}V. T) -| crf_ib2_dx : ∀a,I,L,V,T. ib2 a I → crf (L.ⓑ{I}V) T → crf L (ⓑ{a,I}V. T) -| crf_beta : ∀a,L,V,W,T. crf L (ⓐV. ⓛ{a}W. T) -| crf_theta : ∀a,L,V,W,T. crf L (ⓐV. ⓓ{a}W. T) -. - -interpretation - "context-sensitive reducibility (term)" - 'Reducible L T = (crf L T). - -(* Basic inversion lemmas ***************************************************) - -fact trf_inv_atom_aux: ∀I,L,T. L ⊢ 𝐑⦃T⦄ → L = ⋆ → T = ⓪{I} → ⊥. -#I #L #T * -L -T -[ #L #K #V #i #HLK #H1 #H2 destruct - lapply (ldrop_inv_atom1 … HLK) -HLK #H destruct -| #L #V #T #_ #_ #H destruct -| #L #V #T #_ #_ #H destruct -| #J #L #V #T #_ #_ #H destruct -| #a #J #L #V #T #_ #_ #_ #H destruct -| #a #J #L #V #T #_ #_ #_ #H destruct -| #a #L #V #W #T #_ #H destruct -| #a #L #V #W #T #_ #H destruct -] -qed. - -lemma trf_inv_atom: ∀I. ⋆ ⊢ 𝐑⦃⓪{I}⦄ → ⊥. -/2 width=6/ qed-. - -fact trf_inv_lref_aux: ∀L,T. L ⊢ 𝐑⦃T⦄ → ∀i. T = #i → ∃∃K,V. ⇩[0, i] L ≡ K.ⓓV. -#L #T * -L -T -[ #L #K #V #j #HLK #i #H destruct /2 width=3/ -| #L #V #T #_ #i #H destruct -| #L #V #T #_ #i #H destruct -| #J #L #V #T #_ #i #H destruct -| #a #J #L #V #T #_ #_ #i #H destruct -| #a #J #L #V #T #_ #_ #i #H destruct -| #a #L #V #W #T #i #H destruct -| #a #L #V #W #T #i #H destruct -] -qed. - -lemma trf_inv_lref: ∀L,i. L ⊢ 𝐑⦃#i⦄ → ∃∃K,V. ⇩[0, i] L ≡ K.ⓓV. -/2 width=3/ qed-. - -fact crf_inv_ib2_aux: ∀a,I,L,W,U,T. ib2 a I → L ⊢ 𝐑⦃T⦄ → T = ⓑ{a,I}W. U → - L ⊢ 𝐑⦃W⦄ ∨ L.ⓑ{I}W ⊢ 𝐑⦃U⦄. -#a #I #L #W #U #T #HI * -T -[ #L #K #V #i #_ #H destruct -| #L #V #T #_ #H destruct -| #L #V #T #_ #H destruct -| #J #L #V #T #H1 #H2 destruct - elim H1 -H1 #H destruct - elim HI -HI #H destruct -| #b #J #L #V #T #_ #HV #H destruct /2 width=1/ -| #b #J #L #V #T #_ #HT #H destruct /2 width=1/ -| #b #L #V #W #T #H destruct -| #b #L #V #W #T #H destruct -] -qed. - -lemma crf_inv_ib2: ∀a,I,L,W,T. ib2 a I → L ⊢ 𝐑⦃ⓑ{a,I}W.T⦄ → - L ⊢ 𝐑⦃W⦄ ∨ L.ⓑ{I}W ⊢ 𝐑⦃T⦄. -/2 width=5/ qed-. - -fact crf_inv_appl_aux: ∀L,W,U,T. L ⊢ 𝐑⦃T⦄ → T = ⓐW. U → - ∨∨ L ⊢ 𝐑⦃W⦄ | L ⊢ 𝐑⦃U⦄ | (𝐒⦃U⦄ → ⊥). -#L #W #U #T * -T -[ #L #K #V #i #_ #H destruct -| #L #V #T #HV #H destruct /2 width=1/ -| #L #V #T #HT #H destruct /2 width=1/ -| #J #L #V #T #H1 #H2 destruct - elim H1 -H1 #H destruct -| #a #I #L #V #T #_ #_ #H destruct -| #a #I #L #V #T #_ #_ #H destruct -| #a #L #V #W0 #T #H destruct - @or3_intro2 #H elim (simple_inv_bind … H) -| #a #L #V #W0 #T #H destruct - @or3_intro2 #H elim (simple_inv_bind … H) -] -qed. - -lemma crf_inv_appl: ∀L,V,T. L ⊢ 𝐑⦃ⓐV.T⦄ → ∨∨ L ⊢ 𝐑⦃V⦄ | L ⊢ 𝐑⦃T⦄ | (𝐒⦃T⦄ → ⊥). -/2 width=3/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/crf_append.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/crf_append.ma deleted file mode 100644 index f50b97e95..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/crf_append.ma +++ /dev/null @@ -1,56 +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/substitution/ldrop_append.ma". -include "basic_2/reducibility/crf.ma". - -(* CONTEXT-SENSITIVE REDUCIBLE TERMS ****************************************) - -(* Advanved properties ******************************************************) - -lemma crf_labst_last: ∀L,T,W. L ⊢ 𝐑⦃T⦄ → ⋆.ⓛW @@ L ⊢ 𝐑⦃T⦄. -#L #T #W #H elim H -L -T /2 width=1/ -#L #K #V #i #HLK -lapply (ldrop_fwd_ldrop2_length … HLK) #Hi -lapply (ldrop_O1_append_sn_le … HLK … (⋆.ⓛW)) -HLK /2 width=2/ -Hi /2 width=3/ -qed. - -lemma crf_trf: ∀T,W. ⋆ ⊢ 𝐑⦃T⦄ → ⋆.ⓛW ⊢ 𝐑⦃T⦄. -#T #W #H lapply (crf_labst_last … W H) // -qed. - -(* Advanced inversion lemmas ************************************************) - -fact crf_inv_labst_last_aux: ∀L1,T,W. L1 ⊢ 𝐑⦃T⦄ → - ∀L2. L1 = ⋆.ⓛW @@ L2 → L2 ⊢ 𝐑⦃T⦄. -#L1 #T #W #H elim H -L1 -T /2 width=1/ /3 width=1/ -[ #L1 #K1 #V1 #i #HLK1 #L2 #H destruct - lapply (ldrop_fwd_ldrop2_length … HLK1) - >append_length >commutative_plus normalize in ⊢ (??% → ?); #H - elim (le_to_or_lt_eq i (|L2|) ?) /2 width=1/ -H #Hi destruct - [ elim (ldrop_O1_lt … Hi) #I2 #K2 #V2 #HLK2 - lapply (ldrop_O1_inv_append1_le … HLK1 … HLK2) -HLK1 /2 width=2/ -Hi - normalize #H destruct /2 width=3/ - | lapply (ldrop_O1_inv_append1_ge … HLK1 ?) -HLK1 // HL01 #HL12 #HT02 -elim (tpr_conf … HT01 HT02) -L0 -T0 #X #H1 #H2 -elim (tpr_fwd_shift1 … H1) #L #T #HL1 #H destruct /3 width=5/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr.ma deleted file mode 100644 index 44745490e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr.ma +++ /dev/null @@ -1,44 +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/unfold/ltpss_sn.ma". -include "basic_2/reducibility/ltpr.ma". - -(* FOCALIZED PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ***********************) - -definition lfpr: relation lenv ≝ - λL1,L2. ∃∃L. L1 ➡ L & L ⊢ ▶* [0, |L|] L2 -. - -interpretation - "focalized parallel reduction (environment)" - 'FocalizedPRed L1 L2 = (lfpr L1 L2). - -(* Basic properties *********************************************************) - -(* Note: lemma 250 *) -lemma lfpr_refl: ∀L. ⦃L⦄ ➡ ⦃L⦄. -/2 width=3/ qed. - -lemma ltpss_sn_lfpr: ∀L1,L2,d,e. L1 ⊢ ▶* [d, e] L2 → ⦃L1⦄ ➡ ⦃L2⦄. -/3 width=5/ qed. - -lemma ltpr_lfpr: ∀L1,L2. L1 ➡ L2 → ⦃L1⦄ ➡ ⦃L2⦄. -/3 width=3/ qed. - -(* Basic inversion lemmas ***************************************************) - -lemma lfpr_inv_atom1: ∀L2. ⦃⋆⦄ ➡ ⦃L2⦄ → L2 = ⋆. -#L2 * #L #HL >(ltpr_inv_atom1 … HL) -HL #HL2 >(ltpss_sn_inv_atom1 … HL2) -HL2 // -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_aaa.ma deleted file mode 100644 index 6f6c49df3..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_aaa.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/static/aaa_ltpss_sn.ma". -include "basic_2/reducibility/ltpr_aaa.ma". -include "basic_2/reducibility/lfpr.ma". - -(* FOCALIZED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS **********************) - -(* Properties about atomic arity assignment on terms ************************) - -lemma aaa_lfpr_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. ⦃L1⦄ ➡ ⦃L2⦄ → L2 ⊢ T ⁝ A. -#L1 #T #A #HT #L2 * /3 width=5/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_alt.ma deleted file mode 100644 index adff7ad07..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_alt.ma +++ /dev/null @@ -1,79 +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/grammar/lenv_px_bi.ma". -include "basic_2/reducibility/fpr_cpr.ma". -include "basic_2/reducibility/lfpr_fpr.ma". - -(* FOCALIZED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS **********************) - -(* alternative definition *) -definition lfpra: relation lenv ≝ lpx_bi fpr. - -interpretation - "focalized parallel reduction (environment) alternative" - 'FocalizedPRedAlt L1 L2 = (lfpra L1 L2). - -(* Basic properties *********************************************************) - -lemma lfpra_refl: reflexive … lfpra. -/2 width=1/ qed. - -lemma fpr_lfpra: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ⦃L1⦄ ➡➡ ⦃L2⦄. -#L1 elim L1 -L1 -[ #L2 #T1 #T2 #H - elim (fpr_inv_atom1 … H) -H #_ #H destruct // -| #L1 #I #V1 #IH #L2 #T1 #T2 #H - elim (fpr_inv_pair1 … H) -H #L #V #HV1 #HL1 #H destruct /3 width=3/ -] -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma lfpra_inv_atom1: ∀L2. ⦃⋆⦄ ➡➡ ⦃L2⦄ → L2 = ⋆. -/2 width=2 by lpx_bi_inv_atom1/ qed-. - -lemma lfpra_inv_pair1: ∀K1,I,V1,L2. ⦃K1. ⓑ{I} V1⦄ ➡➡ ⦃L2⦄ → - ∃∃K2,V2. ⦃K1⦄ ➡➡ ⦃K2⦄ & ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ & - L2 = K2. ⓑ{I} V2. -/2 width=1 by lpx_bi_inv_pair1/ qed-. - -lemma lfpra_inv_atom2: ∀L1. ⦃L1⦄ ➡➡ ⦃⋆⦄ → L1 = ⋆. -/2 width=2 by lpx_bi_inv_atom2/ qed-. - -lemma lfpra_inv_pair2: ∀L1,K2,I,V2. ⦃L1⦄ ➡➡ ⦃K2. ⓑ{I} V2⦄ → - ∃∃K1,V1. ⦃K1⦄ ➡➡ ⦃K2⦄ & ⦃K1, V1⦄ ➡ ⦃K2, V2⦄ & - L1 = K1. ⓑ{I} V1. -/2 width=1 by lpx_bi_inv_pair2/ qed-. - -lemma lfpra_inv_fpr: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → ∀T.⦃L1, T⦄ ➡ ⦃L2, T⦄. -#L1 #L2 * -L1 -L2 // /3 width=1/ -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lfpra_fwd_length: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → |L1| = |L2|. -/2 width=2 by lpx_bi_fwd_length/ qed-. - -(* Main properties **********************************************************) - -theorem lfpr_lfpra: ∀L1,L2. ⦃L1⦄ ➡ ⦃L2⦄ → ⦃L1⦄ ➡➡ ⦃L2⦄. -#L1 #L2 #H -lapply (lfpr_inv_fpr … H (⋆0)) -H /2 width=3/ -qed. - -theorem lfpra_lfpr: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → ⦃L1⦄ ➡ ⦃L2⦄. -#L1 #L2 #H -lapply (lfpra_inv_fpr … H (⋆0)) -H /2 width=3/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_cpr.ma deleted file mode 100644 index 8a01f263a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_cpr.ma +++ /dev/null @@ -1,52 +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/unfold/ltpss_sn_ltpss_sn.ma". -include "basic_2/reducibility/ltpr_ldrop.ma". -include "basic_2/reducibility/cpr.ma". -include "basic_2/reducibility/lfpr.ma". - -(* FOCALIZED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS **********************) - -(* Advanced properties ******************************************************) - -lemma lfpr_pair_cpr: ∀L1,L2. ⦃L1⦄ ➡ ⦃L2⦄ → ∀V1,V2. L2 ⊢ V1 ➡ V2 → - ∀I. ⦃L1. ⓑ{I} V1⦄ ➡ ⦃L2. ⓑ{I} V2⦄. -#L1 #L2 * #L #HL1 #HL2 #V1 #V2 * -<(ltpss_sn_fwd_length … HL2) #V #HV1 #HV2 #I -lapply (ltpss_sn_tpss_trans_eq … HV2 … HL2) -HV2 #V2 -@(ex2_intro … (L.ⓑ{I}V)) /2 width=1/ (**) (* explicit constructor *) -qed. - -(* Properties on supclosure *************************************************) -(* -lamma fsub_cpr_trans: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ➡ U2 → - ∃∃L,U1. ⦃L1⦄ ➡ ⦃L⦄ & L ⊢ T1 ➡ U1 & ⦃L, U1⦄ ⊃ ⦃L2, U2⦄. -#L1 #L2 #T1 #T2 #HT12 #U2 * #T #H1 #H2 -elim (fsub_tpr_trans … HT12 … H1) -T2 #L #U #HL1 #HT1U #HUT -elim (fsup_tpss_trans_full … HUT … H2) -T -HUT -H2 #L #U #HL1 #HT1U #HUT - - - - - - - #H elim H -L1 -L2 -T1 -T2 [1,2,3,4,5: /3 width=5/ ] -#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 -elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HK2 -elim (lift_total T d e) #U #HTU -elim (ldrop_ltpr_trans … HLK1 … HK1) -HLK1 -HK1 #L #HL1 #HLK -lapply (tpr_lift … HT1 … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ -qed-. -*) \ No newline at end of file diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_fpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_fpr.ma deleted file mode 100644 index f798d566b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_fpr.ma +++ /dev/null @@ -1,31 +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/reducibility/lfpr.ma". -include "basic_2/reducibility/cfpr_cpr.ma". - -(* FOCALIZED PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ***********************) - -(* Properties on context-free parallel reduction for closures ***************) - -lemma fpr_lfpr: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ⦃L1⦄ ➡ ⦃L2⦄. -#L1 #L2 #T1 #T2 #H -elim (fpr_inv_all … H) -H /2 width=3/ -qed. - -(* Inversion lemmas on context-free parallel reduction for closures *********) - -lemma lfpr_inv_fpr: ∀L1,L2. ⦃L1⦄ ➡ ⦃L2⦄ → ∀T. ⦃L1, T⦄ ➡ ⦃L2, T⦄. -#L1 #L2 * /2 width=4/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_lfpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_lfpr.ma deleted file mode 100644 index 7031e792b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/lfpr_lfpr.ma +++ /dev/null @@ -1,39 +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/reducibility/ltpr_ltpss_sn.ma". -include "basic_2/reducibility/ltpr_ltpr.ma". -include "basic_2/reducibility/lfpr.ma". - -(* FOCALIZED PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ***********************) - -(* Main properties **********************************************************) - -theorem lfpr_conf: ∀L0,L1,L2. ⦃L0⦄ ➡ ⦃L1⦄ → ⦃L0⦄ ➡ ⦃L2⦄ → - ∃∃L. ⦃L1⦄ ➡ ⦃L⦄ & ⦃L2⦄ ➡ ⦃L⦄. -#K0 #L1 #L2 * #K1 #HK01 #HKL1 * #K2 #HK02 #HKL2 -lapply (ltpr_fwd_length … HK01) #H ->(ltpr_fwd_length … HK02) in H; #H -elim (ltpr_conf … HK01 … HK02) -K0 #K #HK1 #HK2 -lapply (ltpss_sn_fwd_length … HKL1) #H1 -lapply (ltpss_sn_fwd_length … HKL2) #H2 ->H1 in HKL1 H; -H1 #HKL1 ->H2 in HKL2; -H2 #HKL2 #H -elim (ltpr_ltpss_sn_conf … HKL1 … HK1) -K1 #K1 #HK1 #HLK1 -elim (ltpr_ltpss_sn_conf … HKL2 … HK2) -K2 #K2 #HK2 #HLK2 -elim (ltpss_sn_conf … HK1 … HK2) -K #K #HK1 #HK2 -lapply (ltpr_fwd_length … HLK1) #H1 -lapply (ltpr_fwd_length … HLK2) #H2 -/3 width=5/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr.ma deleted file mode 100644 index a910ea7df..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr.ma +++ /dev/null @@ -1,67 +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/grammar/lenv_px.ma". -include "basic_2/reducibility/tpr.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -definition ltpr: relation lenv ≝ lpx tpr. - -interpretation - "context-free parallel reduction (environment)" - 'PRed L1 L2 = (ltpr L1 L2). - -(* Basic properties *********************************************************) - -lemma ltpr_refl: reflexive … ltpr. -/2 width=1/ qed. - -lemma ltpr_append: ∀K1,K2. K1 ➡ K2 → ∀L1,L2:lenv. L1 ➡ L2 → K1 @@ L1 ➡ K2 @@ L2. -/2 width=1/ qed. - -(* Basic inversion lemmas ***************************************************) - -(* Basic_1: was: wcpr0_gen_sort *) -lemma ltpr_inv_atom1: ∀L2. ⋆ ➡ L2 → L2 = ⋆. -/2 width=2 by lpx_inv_atom1/ qed-. - -(* Basic_1: was: wcpr0_gen_head *) -lemma ltpr_inv_pair1: ∀K1,I,V1,L2. K1. ⓑ{I} V1 ➡ L2 → - ∃∃K2,V2. K1 ➡ K2 & V1 ➡ V2 & L2 = K2. ⓑ{I} V2. -/2 width=1 by lpx_inv_pair1/ qed-. - -lemma ltpr_inv_atom2: ∀L1. L1 ➡ ⋆ → L1 = ⋆. -/2 width=2 by lpx_inv_atom2/ qed-. - -lemma ltpr_inv_pair2: ∀L1,K2,I,V2. L1 ➡ K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 ➡ K2 & V1 ➡ V2 & L1 = K1. ⓑ{I} V1. -/2 width=1 by lpx_inv_pair2/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma ltpr_fwd_length: ∀L1,L2. L1 ➡ L2 → |L1| = |L2|. -/2 width=2 by lpx_fwd_length/ qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma ltpr_inv_append1: ∀K1,L1. ∀L:lenv. K1 @@ L1 ➡ L → - ∃∃K2,L2. K1 ➡ K2 & L1 ➡ L2 & L = K2 @@ L2. -/2 width=1 by lpx_inv_append1/ qed-. - -lemma ltpr_inv_append2: ∀L:lenv. ∀K2,L2. L ➡ K2 @@ L2 → - ∃∃K1,L1. K1 ➡ K2 & L1 ➡ L2 & L = K1 @@ L1. -/2 width=1 by lpx_inv_append2/ qed-. - -(* Basic_1: removed theorems 2: wcpr0_getl wcpr0_getl_back *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_aaa.ma deleted file mode 100644 index 5208a6405..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_aaa.ma +++ /dev/null @@ -1,85 +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/static/aaa_ltpss_dx.ma". -include "basic_2/static/lsuba_aaa.ma". -include "basic_2/reducibility/ltpr_ldrop.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Properties about atomic arity assignment on terms ************************) - -lemma aaa_ltpr_tpr_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → ∀L2. L1 ➡ L2 → - ∀T2. T1 ➡ T2 → L2 ⊢ T2 ⁝ A. -#L1 #T1 @(f2_ind … fw … L1 T1) -L1 -T1 #n #IH #L1 * * -[1,2,3: - #i #Hn #X #H1 #L2 #HL12 #Y #H2 destruct - >(tpr_inv_atom1 … H2) -Y -|4,5: [ #a ] * #V1 #T1 #Hn #X #H1 #L2 #HL12 #Y #H2 destruct -] -[ >(aaa_inv_sort … H1) -X // -| elim (aaa_inv_lref … H1) #I #K1 #V1 #HLK1 #HA - lapply (ldrop_pair2_fwd_fw … HLK1 (#i)) #HKV1 - elim (ltpr_ldrop_conf … HLK1 … HL12) -HLK1 -HL12 #Y #H #HLK2 - elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct - lapply (IH … HKV1 … HA … HK12 … HV12) -L1 -K1 -V1 /2 width=5/ -| elim (aaa_inv_gref … H1) -| elim (aaa_inv_abbr … H1) -H1 #B #HB #HA - elim (tpr_inv_abbr1 … H2) -H2 * - [ #V2 #T #T2 #HV12 #HT1 #HT2 #H destruct - lapply (tps_lsubr_trans … HT2 (L2.ⓓV2) ?) -HT2 /2 width=1/ #HT2 - lapply (IH … HB … HL12 … HV12) -HB /width=5/ #HB - lapply (IH … HA … (L2.ⓓV2) … HT1) -IH -HA -HT1 /width=5/ -T1 /2 width=1/ -L1 -V1 /3 width=5/ - | -B #T #HT1 #HXT #H destruct - lapply (IH … HA … (L2.ⓓV1) … HT1) /width=5/ -T1 /2 width=1/ -L1 #HA - @(aaa_inv_lift … HA … HXT) /2 width=1/ - ] -| elim (aaa_inv_abst … H1) -H1 #B #A #HB #HA #H destruct - elim (tpr_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HT12 #H destruct - lapply (IH … HB … HL12 … HV12) -HB /width=5/ #HB - lapply (IH … HA … (L2.ⓛV2) … HT12) -IH -HA -HT12 /width=5/ -T1 /2 width=1/ -| elim (aaa_inv_appl … H1) -H1 #B #HB #HA - elim (tpr_inv_appl1 … H2) -H2 * - [ #V2 #T2 #HV12 #HT12 #H destruct - lapply (IH … HB … HL12 … HV12) -HB -HV12 /width=5/ #HB - lapply (IH … HA … HL12 … HT12) -IH -HA -HL12 -HT12 /width=5/ /2 width=3/ - | #b #V2 #W2 #T0 #T2 #HV12 #HT02 #H1 #H2 destruct - elim (aaa_inv_abst … HA) -HA #B0 #A0 #HB0 #HA0 #H destruct - lapply (IH … HB … HL12 … HV12) -HB -HV12 /width=5/ #HB - lapply (IH … HB0 … HL12 W2 ?) -HB0 /width=5/ #HB0 - lapply (IH … HA0 … (L2.ⓛW2) … HT02) -IH -HA0 -HT02 // /2 width=1/ -T0 -L1 -V1 /4 width=7/ - | #b #V0 #V2 #W0 #W2 #T0 #T2 #HV10 #HW02 #HT02 #HV02 #H1 #H2 destruct - elim (aaa_inv_abbr … HA) -HA #B0 #HW0 #HT0 - lapply (IH … HW0 … HL12 … HW02) -HW0 /width=5/ #HW2 - lapply (IH … HB … HL12 … HV10) -HB -HV10 /width=5/ #HV0 - lapply (IH … HT0 … (L2.ⓓW2) … HT02) -IH -HT0 -HT02 // /2 width=1/ -V1 -T0 -L1 -W0 #HT2 - @(aaa_abbr … HW2) -HW2 - @(aaa_appl … HT2) -HT2 /3 width=7/ (**) (* explict constructors, /5 width=7/ is too slow *) - ] -| elim (aaa_inv_cast … H1) -H1 #HV1 #HT1 - elim (tpr_inv_cast1 … H2) -H2 - [ * #V2 #T2 #HV12 #HT12 #H destruct - lapply (IH … HV1 … HL12 … HV12) -HV1 -HV12 /width=5/ #HV2 - lapply (IH … HT1 … HL12 … HT12) -IH -HT1 -HL12 -HT12 /width=5/ -L1 -V1 -T1 /2 width=1/ - | -HV1 #HT1X - lapply (IH … HT1 … HL12 … HT1X) -IH -HT1 -HL12 -HT1X /width=5/ - ] -] -qed. - -lemma aaa_ltpr_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. L1 ➡ L2 → L2 ⊢ T ⁝ A. -/2 width=5/ qed. - -lemma aaa_tpr_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → ∀T2. T1 ➡ T2 → L ⊢ T2 ⁝ A. -/2 width=5/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ldrop.ma deleted file mode 100644 index 02404a11d..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ldrop.ma +++ /dev/null @@ -1,45 +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/substitution/ldrop_lpx.ma". -include "basic_2/substitution/fsup.ma". -include "basic_2/reducibility/tpr_lift.ma". -include "basic_2/reducibility/ltpr.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Properies on local environment slicing ***********************************) - -(* Basic_1: was: wcpr0_drop *) -lemma ltpr_ldrop_conf: dropable_sn ltpr. -/3 width=3 by lpx_deliftable_dropable, tpr_inv_lift1/ qed. - -(* Basic_1: was: wcpr0_drop_back *) -lemma ldrop_ltpr_trans: dedropable_sn ltpr. -/2 width=3/ qed. - -lemma ltpr_ldrop_trans_O1: dropable_dx ltpr. -/2 width=3/ qed. - -(* Properties on supclosure *************************************************) - -lemma fsub_tpr_trans: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. T2 ➡ U2 → - ∃∃L,U1. L1 ➡ L & T1 ➡ U1 & ⦃L, U1⦄ ⊃ ⦃L2, U2⦄. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [1,2,3,4,5: /3 width=5/ ] -#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 -elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HK2 -elim (lift_total T d e) #U #HTU -elim (ldrop_ltpr_trans … HLK1 … HK1) -HLK1 -HK1 #L #HL1 #HLK -lapply (tpr_lift … HT1 … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpr.ma deleted file mode 100644 index 4a27a6e70..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpr.ma +++ /dev/null @@ -1,29 +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/reducibility/tpr_tpr.ma". -include "basic_2/reducibility/ltpr.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Main properties **********************************************************) - -theorem ltpr_conf: ∀L0:lenv. ∀L1. L0 ➡ L1 → ∀L2. L0 ➡ L2 → - ∃∃L. L1 ➡ L & L2 ➡ L. -#L0 #L1 #H elim H -L0 -L1 /2 width=3/ -#I #K0 #K1 #V0 #V1 #_ #HV01 #IHK01 #L2 #H -elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK02 #HV02 #H destruct -elim (IHK01 … HK02) -K0 #K #HK1 #HK2 -elim (tpr_conf … HV01 HV02) -V0 /3 width=5/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpss_dx.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpss_dx.ma deleted file mode 100644 index 00730b732..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpss_dx.ma +++ /dev/null @@ -1,36 +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/reducibility/ltpr_tpss.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Properties concerning dx parallel unfold on local environments ***********) - -lemma ltpr_ltpss_dx_conf: ∀L1,K1,d,e. L1 ▶* [d, e] K1 → ∀L2. L1 ➡ L2 → - ∃∃K2. L2 ▶* [d, e] K2 & K1 ➡ K2. -#L1 #K1 #d #e #H elim H -L1 -K1 -d -e -[ /2 width=3/ -| #L1 #I #V1 #X #H - elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct /3 width=5/ -| #L1 #K1 #I #V1 #W1 #e #_ #HVW1 #IHLK1 #X #H - elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct - elim (IHLK1 … HL12) -L1 #K2 #HLK2 #HK12 - elim (ltpr_tpr_tpss_conf … HK12 … HV12 … HVW1) -V1 /3 width=5/ -| #L1 #K1 #I #V1 #W1 #d #e #_ #HVW1 #IHLK1 #X #H - elim (ltpr_inv_pair1 … H) -H #L2 #V2 #HL12 #HV12 #H destruct - elim (IHLK1 … HL12) -L1 #K2 #HLK2 #HK12 - elim (ltpr_tpr_tpss_conf … HK12 … HV12 … HVW1) -V1 /3 width=5/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpss_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpss_sn.ma deleted file mode 100644 index 79ffb7c9e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_ltpss_sn.ma +++ /dev/null @@ -1,31 +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/unfold/ltpss_sn_alt.ma". -include "basic_2/reducibility/ltpr_ltpss_dx.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Properties on sn parallel unfold on local environments *******************) - -(* Note: this can also be proved like ltpr_ltpss_dx_conf *) -lemma ltpr_ltpss_sn_conf: ∀L1,K1,d,e. L1 ⊢ ▶* [d, e] K1 → ∀L2. L1 ➡ L2 → - ∃∃K2. L2 ⊢ ▶* [d, e] K2 & K1 ➡ K2. -#L1 #K1 #d #e #H -lapply (ltpss_sn_ltpssa … H) -H #H -@(ltpssa_ind … H) -K1 /2 width=3/ -#K #K1 #_ #HK1 #IHK #L2 #HL12 -elim (IHK … HL12) -L1 #K2 #HLK2 #HK2 -elim (ltpr_ltpss_dx_conf … HK1 … HK2) -K /3 width=3/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_tps.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_tps.ma deleted file mode 100644 index 7f08be62b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_tps.ma +++ /dev/null @@ -1,57 +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/reducibility/ltpr_ldrop.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Properties concerning parallel substitution on terms *********************) - -(* Basic_1: was: pr0_subst1_fwd *) -lemma ltpr_tps_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶ [d, e] T2 → ∀L2. L1 ➡ L2 → - ∃∃T. L2 ⊢ T1 ▶ [d, e] T & T2 ➡ T. -#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e -[ /2 width=3/ -| #L1 #K1 #V1 #W1 #i #d #e #Hdi #Hide #HLK1 #HVW1 #L2 #HL12 - elim (ltpr_ldrop_conf … HLK1 … HL12) -L1 #X #H #HLK2 - elim (ltpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct -K1 - elim (lift_total V2 0 (i+1)) #W2 #HVW2 - lapply (tpr_lift … HV12 … HVW1 … HVW2) -V1 /3 width=4/ -| #L1 #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12 - elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2 - elim (IHT12 (L2.ⓑ{I}V) ?) /2 width=1/ -L1 /3 width=5/ -| #L1 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #HL12 - elim (IHV12 … HL12) -IHV12 - elim (IHT12 … HL12) -L1 /3 width=5/ -] -qed-. - -(* Basic_1: was: pr0_subst1_back *) -lemma ltpr_tps_trans: ∀L2,T1,T2,d,e. L2 ⊢ T1 ▶ [d, e] T2 → ∀L1. L1 ➡ L2 → - ∃∃T. L1 ⊢ T1 ▶ [d, e] T & T ➡ T2. -#L2 #T1 #T2 #d #e #H elim H -L2 -T1 -T2 -d -e -[ /2 width=3/ -| #L2 #K2 #V2 #W2 #i #d #e #Hdi #Hide #HLK2 #HVW2 #L1 #HL12 - elim (ltpr_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H - elim (ltpr_inv_pair2 … H) -H #K1 #V1 #HK12 #HV12 #H destruct -K2 - elim (lift_total V1 0 (i+1)) #W1 #HVW1 - lapply (tpr_lift … HV12 … HVW1 … HVW2) -V2 /3 width=4/ -| #L2 #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12 - elim (IHV12 … HL12) -IHV12 #V #HV1 #HV2 - elim (IHT12 (L1.ⓑ{I}V) ?) /2 width=1/ -L2 /3 width=5/ -| #L2 #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L1 #HL12 - elim (IHV12 … HL12) -IHV12 - elim (IHT12 … HL12) -L2 /3 width=5/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_tpss.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_tpss.ma deleted file mode 100644 index 71586fa36..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/ltpr_tpss.ma +++ /dev/null @@ -1,91 +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/unfold/ltpss_dx_ltpss_dx.ma". -include "basic_2/reducibility/ltpr_tps.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ********************) - -(* Properties on partial unfold for terms ***********************************) - -(* Basic_1: was: pr0_subst1 *) -lemma ltpr_tpr_tps_conf: ∀T1,T2. T1 ➡ T2 → - ∀L1,d,e,U1. L1 ⊢ T1 ▶ [d, e] U1 → - ∀L2. L1 ➡ L2 → - ∃∃U2. U1 ➡ U2 & L2 ⊢ T2 ▶* [d, e] U2. -#T1 #T2 #H elim H -T1 -T2 -[ #I #L1 #d #e #U1 #H #L2 #HL12 - elim (ltpr_tps_conf … H … HL12) -L1 /3 width=3/ -| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_flat1 … H) -H #W1 #U1 #HVW1 #HTU1 #H destruct - elim (IHV12 … HVW1 … HL12) -V1 - elim (IHT12 … HTU1 … HL12) -T1 -HL12 /3 width=5/ -| #a #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct - elim (tps_inv_bind1 … HY) -HY #WW #TT1 #_ #HTT1 #H destruct - elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 - elim (IHT12 … HTT1 (L2. ⓛWW) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 - lapply (tpss_lsubr_trans … HTT2 (L2. ⓓVV2) ?) -HTT2 /3 width=5/ -| #a #I #V1 #V2 #T1 #T #T2 #HV12 #_ #HT2 #IHV12 #IHT1 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_bind1 … H) -H #W1 #U1 #HVW1 #HTU1 #H destruct - elim (IHV12 … HVW1 … HL12) -V1 #W2 #HW12 #HVW2 - elim (IHT1 … HTU1 (L2. ⓑ{I} W2) ?) -T1 /2 width=1/ -HL12 #U #HU1 #HTU - elim (tpss_strip_neq … HTU … HT2 ?) -T /2 width=1/ #U2 #HU2 #HTU2 - lapply (tps_lsubr_trans … HU2 (L2. ⓑ{I} V2) ?) -HU2 /2 width=1/ #HU2 - elim (ltpss_dx_tps_conf … HU2 (L2. ⓑ{I} W2) (d + 1) e ?) -HU2 /2 width=1/ #U3 #HU3 #HU23 - lapply (tps_lsubr_trans … HU3 (⋆. ⓑ{I} W2) ?) -HU3 /2 width=1/ #HU3 - lapply (tpss_lsubr_trans … HU23 (L2. ⓑ{I} W2) ?) -HU23 /2 width=1/ #HU23 - lapply (tpss_trans_eq … HTU2 … HU23) -U2 /3 width=5/ -| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_flat1 … H) -H #VV1 #Y #HVV1 #HY #HX destruct - elim (tps_inv_bind1 … HY) -HY #WW1 #TT1 #HWW1 #HTT1 #H destruct - elim (IHV12 … HVV1 … HL12) -V1 #VV2 #HVV12 #HVV2 - elim (IHW12 … HWW1 … HL12) -W1 #WW2 #HWW12 #HWW2 - elim (IHT12 … HTT1 (L2. ⓓWW2) ?) -T1 /2 width=1/ -HL12 #TT2 #HTT12 #HTT2 - elim (lift_total VV2 0 1) #VV #H2VV - lapply (tpss_lift_ge … HVV2 (L2. ⓓWW2) … HV2 … H2VV) -V2 /2 width=1/ #HVV - @ex2_intro [2: @tpr_theta |1: skip |3: @tpss_bind [2: @tpss_flat ] ] /width=11/ (**) (* /4 width=11/ is too slow *) -| #V #T1 #T #T2 #_ #HT2 #IHT1 #L1 #d #e #X #H #L2 #HL12 - elim (tps_inv_bind1 … H) -H #W #U1 #_ #HTU1 #H destruct -V - elim (IHT1 … HTU1 (L2.ⓓW) ?) -T1 /2 width=1/ -HL12 #U #HU1 #HTU - elim (tpss_inv_lift1_ge … HTU L2 … HT2 ?) -T shift_append_assoc normalize #H - elim (tpr_inv_bind1 … H) -H * - [ #V0 #T0 #X0 #_ #HT10 #H0 #H destruct - elim (IH … HT10) -IH -T1 #L #T #HL1 #H destruct - elim (tps_fwd_shift1 … H0) -T #L2 #T2 #HL2 #H destruct - >append_length >HL1 >HL2 -L1 -L - @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] // /2 width=3/ (**) (* explicit constructor *) - | #T #_ #_ #H destruct - ] -] -qed-. - -(* Basic_1: removed theorems 3: - pr0_subst0_back pr0_subst0_fwd pr0_subst0 -*) -(* Basic_1: removed local theorems: 1: pr0_delta_tau *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_delift.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_delift.ma deleted file mode 100644 index 3fff8f25f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_delift.ma +++ /dev/null @@ -1,27 +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/unfold/delift.ma". -include "basic_2/reducibility/ltpr_tpss.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) - -(* Properties on inverse basic term relocation ******************************) - -lemma tpr_delift_conf: ∀U1,U2. U1 ➡ U2 → ∀L,T1,d,e. L ⊢ ▼*[d, e] U1 ≡ T1 → - ∃∃T2. T1 ➡ T2 & L ⊢ ▼*[d, e] U2 ≡ T2. -#U1 #U2 #HU12 #L #T1 #d #e * #W1 #HUW1 #HTW1 -elim (tpr_tpss_conf … HU12 … HUW1) -U1 #U1 #HWU1 #HU21 -elim (tpr_inv_lift1 … HWU1 … HTW1) -W1 /3 width=5/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_lift.ma deleted file mode 100644 index 8d6b0363a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_lift.ma +++ /dev/null @@ -1,121 +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/substitution/tps_lift.ma". -include "basic_2/reducibility/tpr.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) - -(* Relocation properties ****************************************************) - -(* Basic_1: was: pr0_lift *) -lemma tpr_lift: t_liftable tpr. -#T1 #T2 #H elim H -T1 -T2 -[ * #i #U1 #d #e #HU1 #U2 #HU2 - lapply (lift_mono … HU1 … HU2) -HU1 #H destruct - [ lapply (lift_inv_sort1 … HU2) -HU2 #H destruct // - | lapply (lift_inv_lref1 … HU2) * * #Hid #H destruct // - | lapply (lift_inv_gref1 … HU2) -HU2 #H destruct // - ] -| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #X1 #d #e #HX1 #X2 #HX2 - elim (lift_inv_flat1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct - elim (lift_inv_flat1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct /3 width=4/ -| #a #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #X1 #d #e #HX1 #X2 #HX2 - elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct - elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct - elim (lift_inv_bind1 … HX2) -HX2 #V3 #T3 #HV23 #HT23 #HX2 destruct /3 width=4/ -| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV12 #IHT1 #X1 #d #e #HX1 #X2 #HX2 - elim (lift_inv_bind1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct - elim (lift_inv_bind1 … HX2) -HX2 #W2 #U0 #HVW2 #HTU0 #HX2 destruct - elim (lift_total T (d + 1) e) #U #HTU - @tpr_delta - [4: @(tps_lift_le … HT2 … HTU HTU0 ?) /2 width=1/ |1: skip |2: /2 width=4/ |3: /2 width=4/ ] (**) (*/3. is too slow *) -| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #X1 #d #e #HX1 #X2 #HX2 - elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct - elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct - elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct - elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct - elim (lift_trans_ge … HV2 … HV3 ?) -V // /3 width=4/ -| #V #T1 #T #T2 #_ #HT2 #IHT1 #X #d #e #H #U2 #HTU2 - elim (lift_inv_bind1 … H) -H #V3 #T3 #_ #HT13 #H destruct -V - elim (lift_conf_O1 … HTU2 … HT2) -T2 /3 width=4/ -| #V #T1 #T2 #_ #IHT12 #X #d #e #HX #T #HT2 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #_ #HT0 #HX destruct /3 width=4/ -] -qed. - -(* Basic_1: was: pr0_gen_lift *) -lemma tpr_inv_lift1: t_deliftable_sn tpr. -#T1 #T2 #H elim H -T1 -T2 -[ * #i #X #d #e #HX - [ lapply (lift_inv_sort2 … HX) -HX #H destruct /2 width=3/ - | lapply (lift_inv_lref2 … HX) -HX * * #Hid #H destruct /3 width=3/ - | lapply (lift_inv_gref2 … HX) -HX #H destruct /2 width=3/ - ] -| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #X #d #e #HX - elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct - elim (IHV12 … HV01) -V1 - elim (IHT12 … HT01) -T1 /3 width=5/ -| #a #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #X #d #e #HX - elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct - elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct - elim (IHV12 … HV01) -V1 - elim (IHT12 … HT01) -T1 /3 width=5/ -| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV12 #IHT1 #X #d #e #HX - elim (lift_inv_bind2 … HX) -HX #W1 #U1 #HWV1 #HUT1 #HX destruct - elim (IHV12 … HWV1) -V1 #W2 #HWV2 #HW12 - elim (IHT1 … HUT1) -T1 #U #HUT #HU1 - elim (tps_inv_lift1_le … HT2 … HUT ?) -T // [3: /2 width=5/ |2: skip ] #U2 #HU2 #HUT2 - @ex2_intro [2: /2 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /3 width=5/ is slow *) -| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #X #d #e #HX - elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct - elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct - elim (IHV12 … HV01) -V1 #V3 #HV32 #HV03 - elim (IHW12 … HW01) -W1 #W3 #HW32 #HW03 - elim (IHT12 … HT01) -T1 #T3 #HT32 #HT03 - elim (lift_trans_le … HV32 … HV2 ?) -V2 // #V2 #HV32 #HV2 - @ex2_intro [2: /3 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /4 width=5/ is slow *) -| #V #T1 #T #T2 #_ #HT2 #IHT1 #X #d #e #HX - elim (lift_inv_bind2 … HX) -HX #V3 #T3 #_ #HT31 #H destruct - elim (IHT1 … HT31) -T1 #T1 #HT1 #HT31 - elim (lift_div_le … HT2 … HT1 ?) -T // /3 width=5/ -| #V #T1 #T2 #_ #IHT12 #X #d #e #HX - elim (lift_inv_flat2 … HX) -HX #V0 #T0 #_ #HT01 #H destruct - elim (IHT12 … HT01) -T1 /3 width=3/ -] -qed-. - -(* Advanced inversion lemmas ************************************************) - -(* Basic_1: was pr0_gen_abst *) -lemma tpr_inv_abst1: ∀a,V1,T1,U2. ⓛ{a}V1. T1 ➡ U2 → - ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛ{a}V2. T2. -#a #V1 #T1 #U2 #H elim (tpr_inv_bind1 … H) -H * -[ #V2 #T #T2 #HV12 #HT1 #HT2 - lapply (tps_inv_refl_SO2 … HT2 ???) -HT2 // /2 width=5/ -| #T2 #_ #_ #_ #H destruct -] -qed-. - -(* Advanced forward lemmas **************************************************) - -lemma tpr_fwd_abst1: ∀a,V1,T1,U2. ⓛ{a}V1.T1 ➡ U2 → ∀b,I,W. - ∃∃V2,T2. ⓑ{b,I}W.T1 ➡ ⓑ{b,I}W.T2 & - U2 = ⓛ{a}V2.T2. -#a #V1 #T1 #U2 #H #b #I #W elim (tpr_inv_bind1 … H) -H * -[ #V2 #T #T2 #HV12 #HT1 #HT2 - lapply (tps_inv_refl_SO2 … HT2 ???) -HT2 // /3 width=4/ -| #T2 #_ #_ #_ #H destruct -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_tpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_tpr.ma deleted file mode 100644 index a733345bf..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/tpr_tpr.ma +++ /dev/null @@ -1,261 +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/reducibility/ltpr_tpss.ma". - -(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) - -(* Confluence lemmas ********************************************************) - -fact tpr_conf_atom_atom: ∀I. ∃∃X. ⓪{I} ➡ X & ⓪{I} ➡ X. -/2 width=3/ qed. - -fact tpr_conf_flat_flat: - ∀I,V0,V1,T0,T1,V2,T2. ( - ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → - ∃∃T0. ⓕ{I} V1. T1 ➡ T0 & ⓕ{I} V2. T2 ➡ T0. -#I #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02 -elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /3 width=5/ -qed. - -fact tpr_conf_flat_beta: - ∀a,V0,V1,T1,V2,W0,U0,T2. ( - ∀X0:term. ♯{X0} < ♯{V0} + (♯{W0} + ♯{U0} + 1) + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → - U0 ➡ T2 → ⓛ{a}W0. U0 ➡ T1 → - ∃∃X. ⓐV1. T1 ➡ X & ⓓ{a}V2. T2 ➡ X. -#a #V0 #V1 #T1 #V2 #W0 #U0 #T2 #IH #HV01 #HV02 #HT02 #H -elim (tpr_inv_abst1 … H) -H #W1 #U1 #HW01 #HU01 #H destruct -elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 -elim (IH … HT02 … HU01) -HT02 -HU01 -IH /2 width=1/ /3 width=5/ -qed. - -(* Basic-1: was: - pr0_cong_upsilon_refl pr0_cong_upsilon_zeta - pr0_cong_upsilon_cong pr0_cong_upsilon_delta -*) -fact tpr_conf_flat_theta: - ∀a,V0,V1,T1,V2,V,W0,W2,U0,U2. ( - ∀X0:term. ♯{X0} < ♯{V0} + (♯{W0} + ♯{U0} + 1) + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → ⇧[O,1] V2 ≡ V → - W0 ➡ W2 → U0 ➡ U2 → ⓓ{a}W0. U0 ➡ T1 → - ∃∃X. ⓐV1. T1 ➡ X & ⓓ{a}W2. ⓐV. U2 ➡ X. -#a #V0 #V1 #T1 #V2 #V #W0 #W2 #U0 #U2 #IH #HV01 #HV02 #HV2 #HW02 #HU02 #H -elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #VV #HVV1 #HVV2 -elim (lift_total VV 0 1) #VVV #HVV -lapply (tpr_lift … HVV2 … HV2 … HVV) #HVVV -elim (tpr_inv_abbr1 … H) -H * -(* case 1: delta *) -[ -HV2 -HVV2 #WW2 #UU2 #UU #HWW2 #HUU02 #HUU2 #H destruct - elim (IH … HW02 … HWW2) -HW02 -HWW2 /2 width=1/ #W #HW02 #HWW2 - elim (IH … HU02 … HUU02) -HU02 -HUU02 -IH /2 width=1/ #U #HU2 #HUUU2 - elim (tpr_tps_conf_bind … HWW2 HUUU2 … HUU2) -UU2 #UUU #HUUU2 #HUUU1 - @ex2_intro - [2: @tpr_theta [6: @HVV |7: @HWW2 |8: @HUUU2 |1,2,3,4: skip | // ] - |1:skip - |3: @tpr_delta [3: @tpr_flat |1: skip ] /2 width=5/ - ] (**) (* /5 width=14/ is too slow *) -(* case 3: zeta *) -| -HV2 -HW02 -HVV2 #U1 #HU01 #HTU1 - elim (IH … HU01 … HU02) -HU01 -HU02 -IH // -U0 #U #HU1 #HU2 - elim (tpr_inv_lift1 … HU1 … HTU1) -U1 #UU #HUU #HT1UU #H destruct - @(ex2_intro … (ⓐVV.UU)) /2 width=1/ /3 width=5/ (**) (* /4 width=9/ is too slow *) -] -qed. - -fact tpr_conf_flat_cast: - ∀X2,V0,V1,T0,T1. ( - ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → T0 ➡ T1 → T0 ➡ X2 → - ∃∃X. ⓝV1. T1 ➡ X & X2 ➡ X. -#X2 #V0 #V1 #T0 #T1 #IH #_ #HT01 #HT02 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /3 width=3/ -qed. - -fact tpr_conf_beta_beta: - ∀a. ∀W0:term. ∀V0,V1,T0,T1,V2,T2. ( - ∀X0:term. ♯{X0} < ♯{V0} + (♯{W0} + ♯{T0} + 1) + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → - ∃∃X. ⓓ{a}V1. T1 ➡X & ⓓ{a}V2. T2 ➡ X. -#a #W0 #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02 -elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ -elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ /3 width=5/ -qed. - -(* Basic_1: was: pr0_cong_delta pr0_delta_delta *) -fact tpr_conf_delta_delta: - ∀a,I1,V0,V1,T0,T1,TT1,V2,T2,TT2. ( - ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → - ⋆. ⓑ{I1} V1 ⊢ T1 ▶ [O, 1] TT1 → - ⋆. ⓑ{I1} V2 ⊢ T2 ▶ [O, 1] TT2 → - ∃∃X. ⓑ{a,I1} V1. TT1 ➡ X & ⓑ{a,I1} V2. TT2 ➡ X. -#a #I1 #V0 #V1 #T0 #T1 #TT1 #V2 #T2 #TT2 #IH #HV01 #HV02 #HT01 #HT02 #HTT1 #HTT2 -elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH // #T #HT1 #HT2 -elim (tpr_tps_conf_bind … HV1 HT1 … HTT1) -T1 #U1 #TTU1 #HTU1 -elim (tpr_tps_conf_bind … HV2 HT2 … HTT2) -T2 #U2 #TTU2 #HTU2 -elim (tps_conf_eq … HTU1 … HTU2) -T #U #HU1 #HU2 -@ex2_intro [2,3: @tpr_delta |1: skip ] /width=10/ (**) (* /3 width=10/ is too slow *) -qed. - -fact tpr_conf_delta_zeta: - ∀X2,V0,V1,T0,T1,TT1,T2. ( - ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → T0 ➡ T1 → ⋆. ⓓV1 ⊢ T1 ▶ [O,1] TT1 → - T0 ➡ T2 → ⇧[O, 1] X2 ≡ T2 → - ∃∃X. +ⓓV1. TT1 ➡ X & X2 ➡ X. -#X2 #V0 #V1 #T0 #T1 #TT1 #T2 #IH #_ #HT01 #HTT1 #HT02 #HXT2 -elim (IH … HT01 … HT02) -IH -HT01 -HT02 // -V0 -T0 #T #HT1 #HT2 -elim (tpr_tps_conf_bind ? ? V1 … HT1 HTT1) -T1 // #TT #HTT1 #HTT -elim (tpr_inv_lift1 … HT2 … HXT2) -T2 #X #HXT #HX2 -lapply (tps_inv_lift1_eq … HTT … HXT) -HTT #H destruct /3 width=3/ -qed. - -(* Basic_1: was: pr0_upsilon_upsilon *) -fact tpr_conf_theta_theta: - ∀a,VV1,V0,V1,W0,W1,T0,T1,V2,VV2,W2,T2. ( - ∀X0:term. ♯{X0} < ♯{V0} + (♯{W0} + ♯{T0} + 1) + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - V0 ➡ V1 → V0 ➡ V2 → W0 ➡ W1 → W0 ➡ W2 → T0 ➡ T1 → T0 ➡ T2 → - ⇧[O, 1] V1 ≡ VV1 → ⇧[O, 1] V2 ≡ VV2 → - ∃∃X. ⓓ{a}W1. ⓐVV1. T1 ➡ X & ⓓ{a}W2. ⓐVV2. T2 ➡ X. -#a #VV1 #V0 #V1 #W0 #W1 #T0 #T1 #V2 #VV2 #W2 #T2 #IH #HV01 #HV02 #HW01 #HW02 #HT01 #HT02 #HVV1 #HVV2 -elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 -elim (IH … HW01 … HW02) -HW01 -HW02 /2 width=1/ #W #HW1 #HW2 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ #T #HT1 #HT2 -elim (lift_total V 0 1) #VV #HVV -lapply (tpr_lift … HV1 … HVV1 … HVV) -V1 #HVV1 -lapply (tpr_lift … HV2 … HVV2 … HVV) -V2 -HVV #HVV2 -@ex2_intro [2,3: @tpr_bind |1:skip ] /2 width=5/ (**) (* /4 width=5/ is too slow *) -qed. - -fact tpr_conf_zeta_zeta: - ∀V0:term. ∀X2,TT0,T0,T1,TT2. ( - ∀X0:term. ♯{X0} < ♯{V0} + ♯{TT0} + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - TT0 ➡ T0 → ⇧[O, 1] T1 ≡ T0 → - TT0 ➡ TT2 → ⇧[O, 1] X2 ≡ TT2 → - ∃∃X. T1 ➡ X & X2 ➡ X. -#V0 #X2 #TT0 #T0 #T1 #TT2 #IH #HTT0 #HT10 #HTT02 #HXTT2 -elim (IH … HTT0 … HTT02) -IH -HTT0 -HTT02 // -V0 -TT0 #T #HT0 #HTT2 -elim (tpr_inv_lift1 … HT0 … HT10) -T0 #T0 #HT0 #HT10 -elim (tpr_inv_lift1 … HTT2 … HXTT2) -TT2 #TT2 #HTT2 #HXTT2 -lapply (lift_inj … HTT2 … HT0) -HTT2 #H destruct /2 width=3/ -qed. - -fact tpr_conf_tau_tau: - ∀V0,T0:term. ∀X2,T1. ( - ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → - ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → - ∃∃X. X1 ➡ X & X2 ➡ X - ) → - T0 ➡ T1 → T0 ➡ X2 → - ∃∃X. T1 ➡ X & X2 ➡ X. -#V0 #T0 #X2 #T1 #IH #HT01 #HT02 -elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /2 width=3/ -qed. - -(* Confluence ***************************************************************) - -(* Basic_1: was: pr0_confluence *) -theorem tpr_conf: ∀T0:term. ∀T1,T2. T0 ➡ T1 → T0 ➡ T2 → - ∃∃T. T1 ➡ T & T2 ➡ T. -#T0 @(f_ind … tw … T0) -T0 #n #IH * -[ #I #_ #X1 #X2 #H1 #H2 -n - >(tpr_inv_atom1 … H1) -X1 - >(tpr_inv_atom1 … H2) -X2 -(* case 1: atom, atom *) - // -| * [ #a ] #I #V0 #T0 #Hn #X1 #X2 #H1 #H2 - [ elim (tpr_inv_bind1 … H1) -H1 * - [ #V1 #T1 #U1 #HV01 #HT01 #HTU1 #H1 - | #T1 #HT01 #HXT1 #H11 #H12 - ] - elim (tpr_inv_bind1 … H2) -H2 * - [1,3: #V2 #T2 #U2 #HV02 #HT02 #HTU2 #H2 - |2,4: #T2 #HT02 #HXT2 #H21 #H22 - ] destruct -(* case 2: delta, delta *) - [ /3 width=11 by tpr_conf_delta_delta/ (**) (* /3 width=11/ is too slow *) -(* case 3: zeta, delta (repeated) *) - | @ex2_commute /3 width=10 by tpr_conf_delta_zeta/ -(* case 4: delta, zeta *) - | /3 width=10 by tpr_conf_delta_zeta/ (**) (* /3 width=10/ is too slow *) -(* case 5: zeta, zeta *) - | /3 width=9 by tpr_conf_zeta_zeta/ (**) (* /3 width=9/ is too slow *) - ] - | elim (tpr_inv_flat1 … H1) -H1 * - [ #V1 #T1 #HV01 #HT01 #H1 - | #b1 #V1 #W1 #U1 #T1 #HV01 #HUT1 #H11 #H12 #H13 - | #b1 #V1 #Y1 #W1 #Z1 #U1 #T1 #HV01 #HWZ1 #HUT1 #HVY1 #H11 #H12 #H13 - | #HX1 #H1 - ] - elim (tpr_inv_flat1 … H2) -H2 * - [1,5,9,13: #V2 #T2 #HV02 #HT02 #H2 - |2,6,10,14: #b2 #V2 #W2 #U2 #T2 #HV02 #HUT2 #H21 #H22 #H23 - |3,7,11,15: #b2 #V2 #Y2 #W2 #Z2 #U2 #T2 #HV02 #HWZ2 #HUT2 #HVY2 #H21 #H22 #H23 - |4,8,12,16: #HX2 #H2 - ] destruct -(* case 6: flat, flat *) - [ /3 width=7 by tpr_conf_flat_flat/ (**) (* /3 width=7/ is too slow *) -(* case 7: beta, flat (repeated) *) - | @ex2_commute /3 width=8 by tpr_conf_flat_beta/ -(* case 8: theta, flat (repeated) *) - | @ex2_commute /3 width=11 by tpr_conf_flat_theta/ -(* case 9: tau, flat (repeated) *) - | @ex2_commute /3 width=6 by tpr_conf_flat_cast/ -(* case 10: flat, beta *) - | /3 width=8 by tpr_conf_flat_beta/ (**) (* /3 width=8/ is too slow *) -(* case 11: beta, beta *) - | /3 width=8 by tpr_conf_beta_beta/ (**) (* /3 width=8/ is too slow *) -(* case 12: flat, theta *) - | /3 width=11 by tpr_conf_flat_theta/ (**) (* /3 width=11/ is too slow *) -(* case 13: theta, theta *) - | /3 width=14 by tpr_conf_theta_theta/ (**) (* /3 width=14/ is too slow *) -(* case 14: flat, tau *) - | /3 width=6 by tpr_conf_flat_cast/ (**) (* /3 width=6/ is too slow *) -(* case 15: tau, tau *) - | /3 width=5 by tpr_conf_tau_tau/ - ] - ] -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reducibility/xpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reducibility/xpr.ma deleted file mode 100644 index 4efbd3545..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reducibility/xpr.ma +++ /dev/null @@ -1,66 +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/static/ssta.ma". -include "basic_2/reducibility/cpr.ma". - -lemma arith1: ∀x,y,z,w. z < y → x + (y + w) - z = x + y - z + w. -#x #y #z #w #H (cpr_inv_sort1 … H) // +qed. + +(* Basic_1: was: nf2_abst *) +lemma cnf_abst: ∀a,I,L,V,W,T. L ⊢ 𝐍⦃W⦄ → L. ⓑ{I} V ⊢ 𝐍⦃T⦄ → L ⊢ 𝐍⦃ⓛ{a}W.T⦄. +#a #I #L #V #W #T #HW #HT #X #H +elim (cpr_fwd_abst1 … H I V) -H #W0 #T0 #HW0 #HT0 #H destruct +>(HW … HW0) -W0 >(HT … HT0) -T0 // +qed. + +(* Basic_1: was only: nf2_appl_lref *) +lemma cnf_appl_simple: ∀L,V,T. L ⊢ 𝐍⦃V⦄ → L ⊢ 𝐍⦃T⦄ → 𝐒⦃T⦄ → L ⊢ 𝐍⦃ⓐV.T⦄. +#L #V #T #HV #HT #HS #X #H +elim (cpr_inv_appl1_simple … H ?) -H // #V0 #T0 #HV0 #HT0 #H destruct +>(HV … HV0) -V0 >(HT … HT0) -T0 // +qed. + +(* Basic_1: was: nf2_dec *) +axiom cnf_dec: ∀L,T1. L ⊢ 𝐍⦃T1⦄ ∨ + ∃∃T2. L ⊢ T1 ➡ T2 & (T1 = T2 → ⊥). + +(* Basic_1: removed theorems 1: nf2_abst_shift *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_cif.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_cif.ma new file mode 100644 index 000000000..e87cdb962 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_cif.ma @@ -0,0 +1,28 @@ +(**************************************************************************) +(* ___ *) +(* ||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/cpr_cif.ma". +include "basic_2/reduction/cnf_crf.ma". + +(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) + +(* Main properties on context-sensitive irreducible terms *******************) + +theorem cif_cnf: ∀L,T. L ⊢ 𝐈⦃T⦄ → L ⊢ 𝐍⦃T⦄. +/2 width=3 by cpr_fwd_cif/ qed. + +(* Main inversion lemmas on context-sensitive irreducible terms *************) + +theorem cnf_inv_cif: ∀L,T. L ⊢ 𝐍⦃T⦄ → L ⊢ 𝐈⦃T⦄. +/2 width=4 by cnf_inv_crf/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_crf.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_crf.ma new file mode 100644 index 000000000..0afff1ab0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_crf.ma @@ -0,0 +1,46 @@ +(**************************************************************************) +(* ___ *) +(* ||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/crf.ma". +include "basic_2/reduction/cnf.ma". + +(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) + +(* Advanced inversion lemmas on context-sensitive reducible terms ***********) + +(* Note: this property is unusual *) +lemma cnf_inv_crf: ∀L,T. L ⊢ 𝐑⦃T⦄ → L ⊢ 𝐍⦃T⦄ → ⊥. +#L #T #H elim H -L -T +[ #L #K #V #i #HLK #H + elim (cnf_inv_delta … HLK H) +| #L #V #T #_ #IHV #H + elim (cnf_inv_appl … H) -H /2 width=1/ +| #L #V #T #_ #IHT #H + elim (cnf_inv_appl … H) -H /2 width=1/ +| #I #L #V #T * #H1 #H2 destruct + [ elim (cnf_inv_zeta … H2) + | elim (cnf_inv_tau … H2) + ] +|5,6: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct + [1,3: elim (cnf_inv_abbr … H2) -H2 /2 width=1/ + |*: elim (cnf_inv_abst … H2) -H2 /2 width=1/ + ] +| #a #L #V #W #T #H + elim (cnf_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +| #a #L #V #W #T #H + elim (cnf_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_lift.ma new file mode 100644 index 000000000..acbb020ce --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnf_lift.ma @@ -0,0 +1,47 @@ +(**************************************************************************) +(* ___ *) +(* ||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/cpr_lift.ma". +include "basic_2/reduction/cnf.ma". + +(* CONTEXT-SENSITIVE NORMAL TERMS *******************************************) + +(* Advanced properties ******************************************************) + +(* Basic_1: was only: nf2_csort_lref *) +lemma cnf_lref_atom: ∀L,i. ⇩[0, i] L ≡ ⋆ → L ⊢ 𝐍⦃#i⦄. +#L #i #HL #X #H +elim (cpr_inv_lref1 … H) -H // * +#K #V1 #V2 #HLK #_ #_ +lapply (ldrop_mono … HL … HLK) -L #H destruct +qed. + +(* Basic_1: was: nf2_lref_abst *) +lemma cnf_lref_abst: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓛV → L ⊢ 𝐍⦃#i⦄. +#L #K #V #i #HLK #X #H +elim (cpr_inv_lref1 … H) -H // * +#K0 #V1 #V2 #HLK0 #_ #_ +lapply (ldrop_mono … HLK … HLK0) -L #H destruct +qed. + +(* Relocation properties ****************************************************) + +(* Basic_1: was: nf2_lift *) +lemma cnf_lift: ∀L0,L,T,T0,d,e. + L ⊢ 𝐍⦃T⦄ → ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → L0 ⊢ 𝐍⦃T0⦄. +#L0 #L #T #T0 #d #e #HLT #HL0 #HT0 #X #H +elim (cpr_inv_lift1 … H … HL0 … HT0) -L0 #T1 #HT10 #HT1 +<(HLT … HT1) in HT0; -L #HT0 +>(lift_mono … HT10 … HT0) -T1 -X // +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma new file mode 100644 index 000000000..e6ad7b2ba --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma @@ -0,0 +1,339 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/cpqs.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) + +(* Basic_1: includes: pr0_delta1 pr2_delta1 pr2_thin_dx pr2_head_1 *) +(* Note: cpr_flat: does not hold in basic_1 *) +inductive cpr: lenv → relation term ≝ +| cpr_atom : ∀I,L. cpr L (⓪{I}) (⓪{I}) +| cpr_delta: ∀L,K,V,V2,W2,i. + ⇩[0, i] L ≡ K. ⓓV → cpr K V V2 → + ⇧[0, i + 1] V2 ≡ W2 → cpr L (#i) W2 +| cpr_bind : ∀a,I,L,V1,V2,T1,T2. + cpr L V1 V2 → cpr (L. ⓑ{I} V1) T1 T2 → + cpr L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) +| cpr_flat : ∀I,L,V1,V2,T1,T2. + cpr L V1 V2 → cpr L T1 T2 → + cpr L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) +| cpr_zeta : ∀L,V,T1,T,T2. cpr (L.ⓓV) T1 T → + ⇧[0, 1] T2 ≡ T → cpr L (+ⓓV. T1) T2 +| cpr_tau : ∀L,V,T1,T2. cpr L T1 T2 → cpr L (ⓝV. T1) T2 +| cpr_beta : ∀a,L,V1,V2,W,T1,T2. + cpr L V1 V2 → cpr (L.ⓛW) T1 T2 → + cpr L (ⓐV1. ⓛ{a}W. T1) (ⓓ{a}V2. T2) +| cpr_theta: ∀a,L,V1,V,V2,W1,W2,T1,T2. + cpr L V1 V → ⇧[0, 1] V ≡ V2 → cpr L W1 W2 → cpr (L.ⓓW1) T1 T2 → + cpr L (ⓐV1. ⓓ{a}W1. T1) (ⓓ{a}W2. ⓐV2. T2) +. + +interpretation "context-sensitive parallel reduction (term)" + 'PRed L T1 T2 = (cpr L T1 T2). + +(* Basic properties *********************************************************) + +(* Note: it does not hold replacing |L1| with |L2| *) +lemma cpr_lsubr_trans: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 → + ∀L2. L2 ⊑ [0, |L1|] L1 → L2 ⊢ T1 ➡ T2. +#L1 #T1 #T2 #H elim H -L1 -T1 -T2 +[ // +| #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 + lapply (ldrop_fwd_ldrop2_length … HLK1) #Hi + lapply (ldrop_fwd_O1_length … HLK1) #H2i + elim (ldrop_lsubr_ldrop2_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // -Hi + shift_append_assoc normalize #H + elim (cpr_inv_bind1 … H) -H * + [ #V0 #T0 #_ #HT10 #H destruct + elim (IH … HT10) -IH -HT10 #L2 #T2 #HL12 #H destruct + >append_length >HL12 -HL12 + @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] // /2 width=3/ (**) (* explicit constructor *) + | #T #_ #_ #H destruct + ] +] +qed-. + +(* Basic_1: removed theorems 12: + pr0_subst0_back pr0_subst0_fwd pr0_subst0 + pr2_head_2 pr2_cflat clear_pr2_trans + pr2_gen_csort pr2_gen_cflat pr2_gen_cbind + pr2_subst1 + pr2_gen_ctail pr2_ctail +*) +(* Basic_1: removed local theorems 4: + pr0_delta_tau pr0_cong_delta + pr2_free_free pr2_free_delta +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cif.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cif.ma new file mode 100644 index 000000000..2366a0865 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cif.ma @@ -0,0 +1,52 @@ +(**************************************************************************) +(* ___ *) +(* ||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/cif.ma". +include "basic_2/reduction/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) + +(* Advanced forward lemmas on context-sensitive irreducible terms ***********) + +lemma cpr_fwd_cif: ∀L,T1,T2. L ⊢ T1 ➡ T2 → L ⊢ 𝐈⦃T1⦄ → T2 = T1. +#L #T1 #T2 #H elim H -L -T1 -T2 +[ // +| #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #H + elim (cif_inv_delta … HLK ?) // +| #a * #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H + [ elim (cif_inv_bind … H) -H #HV1 #HT1 * #H destruct + lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct + lapply (IHT1 … HT1) -IHT1 #H destruct // + | elim (cif_inv_ib2 … H) -H /2 width=1/ /3 width=2/ + ] +| * #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H + [ elim (cif_inv_appl … H) -H #HV1 #HT1 #_ + >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 // + | elim (cif_inv_ri2 … H) /2 width=1/ + ] +| #L #V1 #T1 #T #T2 #_ #_ #_ #H + elim (cif_inv_ri2 … H) /2 width=1/ +| #L #V1 #T1 #T2 #_ #_ #H + elim (cif_inv_ri2 … H) /2 width=1/ +| #a #L #V1 #V2 #W #T1 #T2 #_ #_ #_ #_ #H + elim (cif_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +| #a #L #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #_ #H + elim (cif_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +] +qed-. + +lemma cpss_fwd_cif_eq: ∀L,T1,T2. L ⊢ T1 ▶* T2 → L ⊢ 𝐈⦃T1⦄ → T2 = T1. +/3 width=3 by cpr_fwd_cif, cpss_cpr/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_lift.ma new file mode 100644 index 000000000..8c3d56e31 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_lift.ma @@ -0,0 +1,106 @@ +(**************************************************************************) +(* ___ *) +(* ||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/relocation/ldrop_ldrop.ma". +include "basic_2/reduction/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) + +(* Relocation properties ****************************************************) + +(* Basic_1: includes: pr0_lift pr2_lift *) +lemma cpr_lift: l_liftable cpr. +#K #T1 #T2 #H elim H -K -T1 -T2 +[ #I #K #L #d #e #_ #U1 #H1 #U2 #H2 + >(lift_mono … H1 … H2) -H1 -H2 // +| #K #KV #V #V2 #W2 #i #HKV #HV2 #HVW2 #IHV2 #L #d #e #HLK #U1 #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HVW2 … HWU2) -W2 // plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=6/ + ] +| #a #I #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #d #e #HLK #U1 #H1 #U2 #H2 + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /4 width=5/ +| #I #K #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #L #d #e #HLK #U1 #H1 #U2 #H2 + elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 destruct + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ +| #K #V #T1 #T #T2 #_ #HT2 #IHT1 #L #d #e #HLK #U1 #H #U2 #HTU2 + elim (lift_inv_bind1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct + elim (lift_conf_O1 … HTU2 … HT2) -T2 /4 width=5/ +| #K #V #T1 #T2 #_ #IHT12 #L #d #e #HLK #U1 #H #U2 #HTU2 + elim (lift_inv_flat1 … H) -H #VV1 #TT1 #HVV1 #HTT1 #H destruct /3 width=5/ +| #a #K #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #L #d #e #HLK #X1 #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct + elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct + elim (lift_inv_bind1 … HX2) -HX2 #V3 #T3 #HV23 #HT23 #HX2 destruct /4 width=5/ +| #a #K #V1 #V #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV1 #IHW12 #IHT12 #L #d #e #HLK #X1 #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct + elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct + elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct + elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct + elim (lift_trans_ge … HV2 … HV3 ?) -V2 // /4 width=5/ +] +qed. + +(* Basic_1: includes: pr0_gen_lift pr2_gen_lift *) +lemma cpr_inv_lift1: l_deliftable_sn cpr. +#L #U1 #U2 #H elim H -L -U1 -U2 +[ * #L #i #K #d #e #_ #T1 #H + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ + ] +| #L #LV #V #V2 #W2 #i #HLV #HV2 #HVW2 #IHV2 #K #d #e #HLK #T1 #H + elim (lift_inv_lref2 … H) -H * #Hid #H destruct + [ elim (ldrop_conf_lt … HLK … HLV) -L // #L #U #HKL #HLV #HUV + elim (IHV2 … HLV … HUV) -V #U2 #HUV2 #HU2 + elim (lift_trans_le … HUV2 … HVW2) -V2 // >minus_plus plus_minus // append_length >commutative_plus normalize in ⊢ (??% → ?); #H + elim (le_to_or_lt_eq i (|L2|) ?) /2 width=1/ -H #Hi destruct + [ elim (ldrop_O1_lt … Hi) #I2 #K2 #V2 #HLK2 + lapply (ldrop_O1_inv_append1_le … HLK1 … HLK2) -HLK1 /2 width=2/ -Hi + normalize #H destruct /2 width=3/ + | lapply (ldrop_O1_inv_append1_ge … HLK1 ?) -HLK1 // (cpr_inv_sort1 … H) -H // +| #I #L1 #K1 #V1 #B #i #HLK1 #_ #IHV1 #X #H #L2 #HL12 + elim (cpr_inv_lref1 … H) -H + [ #H destruct + elim (lpr_ldrop_conf … HLK1 … HL12) -L1 #X #H #HLK2 + elim (lpr_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct /3 width=6/ + | * #Y #Z #V2 #H #HV12 #HV2 + lapply (ldrop_mono … H … HLK1) -H #H destruct + elim (lpr_ldrop_conf … HLK1 … HL12) -L1 #Z #H #HLK2 + elim (lpr_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct + lapply (ldrop_fwd_ldrop2 … HLK2) -V0 /3 width=7/ + ] +| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12 + elim (cpr_inv_abbr1 … H) -H * + [ #V2 #T2 #HV12 #HT12 #H destruct /4 width=2/ + | #T2 #HT12 #HT2 #H destruct -IHV1 + @(aaa_inv_lift (L2.ⓓV1) … HT2) -HT2 /2 width=1/ /3 width=1/ + ] +| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12 + elim (cpr_inv_abst1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=1/ +| #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12 + elim (cpr_inv_appl1 … H) -H * + [ #V2 #T2 #HV12 #HT12 #H destruct /3 width=3/ + | #b #V2 #W1 #U1 #U2 #HV12 #HU12 #H1 #H2 destruct + lapply (IHV1 … HV12 … HL12) -IHV1 -HV12 #HV2 + lapply (IHT1 (ⓛ{b}W1.U2) … HL12) -IHT1 /2 width=1/ -L1 #H + elim (aaa_inv_abst … H) -H #B0 #A0 #HW1 #HU2 #H destruct + lapply (lsuba_aaa_trans … HU2 (L2.ⓓV2) ?) -HU2 /2 width=2/ /2 width=3/ + | #b #V #V2 #W1 #W2 #U1 #U2 #HV1 #HV2 #HW12 #HU12 #H1 #H2 destruct + lapply (aaa_lift L2 … B … (L2.ⓓW2) … HV2) -HV2 /2 width=1/ #HV2 + lapply (IHT1 (ⓓ{b}W2.U2) … HL12) -IHT1 /2 width=1/ -L1 #H + elim (aaa_inv_abbr … H) -H /3 width=3/ + ] +| #L1 #V1 #T1 #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12 + elim (cpr_inv_cast1 … H) -H + [ * #V2 #T2 #HV12 #HT12 #H destruct /3 width=1/ + | -IHV1 /2 width=1/ + ] +] +qed-. + +lemma aaa_cpr_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → ∀T2. L ⊢ T1 ➡ T2 → L ⊢ T2 ⁝ A. +/2 width=5 by aaa_cpr_lpr_conf/ qed-. + +lemma aaa_lpr_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. L1 ⊢ ➡ L2 → L2 ⊢ T ⁝ A. +/2 width=5 by aaa_cpr_lpr_conf/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_cpr.ma new file mode 100644 index 000000000..1d83a1314 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_cpr.ma @@ -0,0 +1,355 @@ +(**************************************************************************) +(* ___ *) +(* ||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/relocation/fsup.ma". +include "basic_2/reduction/lpr_ldrop.ma". + +(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************) + +(* Main properties on context-sensitive parallel reduction for terms ********) + +fact cpr_conf_lpr_atom_atom: + ∀I,L1,L2. ∃∃T. L1 ⊢ ⓪{I} ➡ T & L2 ⊢ ⓪{I} ➡ T. +/2 width=3/ qed-. + +fact cpr_conf_lpr_atom_delta: + ∀L0,i. ( + ∀L,T.♯{L, T} < ♯{L0, #i} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀K0,V0. ⇩[O, i] L0 ≡ K0.ⓓV0 → + ∀V2. K0 ⊢ V0 ➡ V2 → ∀T2. ⇧[O, i + 1] V2 ≡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ #i ➡ T & L2 ⊢ T2 ➡ T. +#L0 #i #IH #K0 #V0 #HLK0 #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02 +elim (lpr_ldrop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1 +elim (lpr_inv_pair1 … H1) -H1 #K1 #V1 #HK01 #HV01 #H destruct +elim (lpr_ldrop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2 +elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct +lapply (ldrop_fwd_ldrop2 … HLK2) -W2 #HLK2 +lapply (ldrop_pair2_fwd_fw … HLK0 (#i)) -HLK0 #HLK0 +elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2 +elim (lift_total V 0 (i+1)) #T #HVT +lapply (cpr_lift … HV2 … HLK2 … HVT2 … HVT) -K2 -V2 /3 width=6/ +qed-. + +(* Basic_1: includes: pr0_delta_delta pr2_delta_delta *) +fact cpr_conf_lpr_delta_delta: + ∀L0,i. ( + ∀L,T.♯{L, T} < ♯{L0, #i} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀K0,V0. ⇩[O, i] L0 ≡ K0.ⓓV0 → + ∀V1. K0 ⊢ V0 ➡ V1 → ∀T1. ⇧[O, i + 1] V1 ≡ T1 → + ∀KX,VX. ⇩[O, i] L0 ≡ KX.ⓓVX → + ∀V2. KX ⊢ VX ➡ V2 → ∀T2. ⇧[O, i + 1] V2 ≡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ T1 ➡ T & L2 ⊢ T2 ➡ T. +#L0 #i #IH #K0 #V0 #HLK0 #V1 #HV01 #T1 #HVT1 +#KX #VX #H #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02 +lapply (ldrop_mono … H … HLK0) -H #H destruct +elim (lpr_ldrop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1 +elim (lpr_inv_pair1 … H1) -H1 #K1 #W1 #HK01 #_ #H destruct +lapply (ldrop_fwd_ldrop2 … HLK1) -W1 #HLK1 +elim (lpr_ldrop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2 +elim (lpr_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct +lapply (ldrop_fwd_ldrop2 … HLK2) -W2 #HLK2 +lapply (ldrop_pair2_fwd_fw … HLK0 (#i)) -HLK0 #HLK0 +elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2 +elim (lift_total V 0 (i+1)) #T #HVT +lapply (cpr_lift … HV1 … HLK1 … HVT1 … HVT) -K1 -V1 +lapply (cpr_lift … HV2 … HLK2 … HVT2 … HVT) -K2 -V2 -V /2 width=3/ +qed-. + +fact cpr_conf_lpr_bind_bind: + ∀a,I,L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓑ{a,I}V0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0.ⓑ{I}V0 ⊢ T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ➡ V2 → ∀T2. L0.ⓑ{I}V0 ⊢ T0 ➡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ ⓑ{a,I}V1.T1 ➡ T & L2 ⊢ ⓑ{a,I}V2.T2 ➡ T. +#a #I #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01 +#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HV01 … HV02 … HL01 … HL02) // +elim (IH … HT01 … HT02 (L1.ⓑ{I}V1) … (L2.ⓑ{I}V2)) -IH // /2 width=1/ /3 width=5/ +qed-. + +fact cpr_conf_lpr_bind_zeta: + ∀L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,+ⓓV0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0.ⓓV0 ⊢ T0 ➡ T1 → + ∀T2. L0.ⓓV0 ⊢ T0 ➡ T2 → ∀X2. ⇧[O, 1] X2 ≡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ +ⓓV1.T1 ➡ T & L2 ⊢ X2 ➡ T. +#L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01 +#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02 +elim (IH … HT01 … HT02 (L1.ⓓV1) … (L2.ⓓV1)) -IH -HT01 -HT02 // /2 width=1/ -L0 -V0 -T0 #T #HT1 #HT2 +elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=1/ /3 width=3/ +qed-. + +fact cpr_conf_lpr_zeta_zeta: + ∀L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,+ⓓV0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀T1. L0.ⓓV0 ⊢ T0 ➡ T1 → ∀X1. ⇧[O, 1] X1 ≡ T1 → + ∀T2. L0.ⓓV0 ⊢ T0 ➡ T2 → ∀X2. ⇧[O, 1] X2 ≡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ X1 ➡ T & L2 ⊢ X2 ➡ T. +#L0 #V0 #T0 #IH #T1 #HT01 #X1 #HXT1 +#T2 #HT02 #X2 #HXT2 #L1 #HL01 #L2 #HL02 +elim (IH … HT01 … HT02 (L1.ⓓV0) … (L2.ⓓV0)) -IH -HT01 -HT02 // /2 width=1/ -L0 -T0 #T #HT1 #HT2 +elim (cpr_inv_lift1 … HT1 L1 … HXT1) -T1 /2 width=1/ #T1 #HT1 #HXT1 +elim (cpr_inv_lift1 … HT2 L2 … HXT2) -T2 /2 width=1/ #T2 #HT2 #HXT2 +lapply (lift_inj … HT2 … HT1) -T #H destruct /2 width=3/ +qed-. + +fact cpr_conf_lpr_flat_flat: + ∀I,L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓕ{I}V0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0 ⊢ T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ➡ V2 → ∀T2. L0 ⊢ T0 ➡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ ⓕ{I}V1.T1 ➡ T & L2 ⊢ ⓕ{I}V2.T2 ➡ T. +#I #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01 +#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HV01 … HV02 … HL01 … HL02) // +elim (IH … HT01 … HT02 … HL01 … HL02) // /3 width=5/ +qed-. + +fact cpr_conf_lpr_flat_tau: + ∀L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓝV0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1,T1. L0 ⊢ T0 ➡ T1 → ∀T2. L0 ⊢ T0 ➡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ ⓝV1.T1 ➡ T & L2 ⊢ T2 ➡ T. +#L0 #V0 #T0 #IH #V1 #T1 #HT01 +#T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /3 width=3/ +qed-. + +fact cpr_conf_lpr_tau_tau: + ∀L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓝV0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀T1. L0 ⊢ T0 ➡ T1 → ∀T2. L0 ⊢ T0 ➡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ T1 ➡ T & L2 ⊢ T2 ➡ T. +#L0 #V0 #T0 #IH #T1 #HT01 +#T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /2 width=3/ +qed-. + +fact cpr_conf_lpr_flat_beta: + ∀a,L0,V0,W0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓐV0.ⓛ{a}W0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0 ⊢ ⓛ{a}W0.T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ➡ V2 → ∀T2. L0.ⓛW0 ⊢ T0 ➡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ ⓐV1.T1 ➡ T & L2 ⊢ ⓓ{a}V2.T2 ➡ T. +#a #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H +#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (cpr_inv_abst1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct +elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (IH … HT01 … HT02 (L1.ⓛW1) … (L2.ⓛW1)) /2 width=1/ -L0 -V0 -W0 -T0 #T #HT1 #HT2 +lapply (cpr_lsubr_trans … HT2 (L2.ⓓV2) ?) -HT2 /2 width=1/ /3 width=5/ +qed-. + +(* Basic-1: includes: + pr0_cong_upsilon_refl pr0_cong_upsilon_zeta + pr0_cong_upsilon_cong pr0_cong_upsilon_delta +*) +fact cpr_conf_lpr_flat_theta: + ∀a,L0,V0,W0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓐV0.ⓓ{a}W0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0 ⊢ ⓓ{a}W0.T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ➡ V2 → ∀U2. ⇧[O, 1] V2 ≡ U2 → + ∀W2. L0 ⊢ W0 ➡ W2 → ∀T2. L0.ⓓW0 ⊢ T0 ➡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ ⓐV1.T1 ➡ T & L2 ⊢ ⓓ{a}W2.ⓐU2.T2 ➡ T. +#a #L0 #V0 #W0 #T0 #IH #V1 #HV01 #X #H +#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (lift_total V 0 1) #U #HVU +lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=1/ #HU2 +elim (cpr_inv_abbr1 … H) -H * +[ #W1 #T1 #HW01 #HT01 #H destruct + elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1/ + elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1/ -L0 -V0 -W0 -T0 /4 width=7/ +| #T1 #HT01 #HXT1 #H destruct + elim (IH … HT01 … HT02 (L1.ⓓW2) … (L2.ⓓW2)) /2 width=1/ -L0 -V0 -W0 -T0 #T #HT1 #HT2 + elim (cpr_inv_lift1 … HT1 L1 … HXT1) -HXT1 /2 width=1/ #Y #HYT #HXY + @(ex2_intro … (ⓐV.Y)) /2 width=1/ /3 width=5/ (**) (* auto /4 width=9/ is too slow *) +] +qed-. + +fact cpr_conf_lpr_beta_beta: + ∀a,L0,V0,W0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓐV0.ⓛ{a}W0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0.ⓛW0 ⊢ T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ➡ V2 → ∀T2. L0.ⓛW0 ⊢ T0 ➡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ ⓓ{a}V1.T1 ➡ T & L2 ⊢ ⓓ{a}V2.T2 ➡ T. +#a #L0 #V0 #W0 #T0 #IH #V1 #HV01 #T1 #HT01 +#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (IH … HT01 … HT02 (L1.ⓛW0) … (L2.ⓛW0)) /2 width=1/ -L0 -V0 -T0 #T #HT1 #HT2 +lapply (cpr_lsubr_trans … HT1 (L1.ⓓV1) ?) -HT1 /2 width=1/ +lapply (cpr_lsubr_trans … HT2 (L2.ⓓV2) ?) -HT2 /2 width=1/ /3 width=5/ +qed-. + +(* Basic_1: was: pr0_upsilon_upsilon *) +fact cpr_conf_lpr_theta_theta: + ∀a,L0,V0,W0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓐV0.ⓓ{a}W0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀U1. ⇧[O, 1] V1 ≡ U1 → + ∀W1. L0 ⊢ W0 ➡ W1 → ∀T1. L0.ⓓW0 ⊢ T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ➡ V2 → ∀U2. ⇧[O, 1] V2 ≡ U2 → + ∀W2. L0 ⊢ W0 ➡ W2 → ∀T2. L0.ⓓW0 ⊢ T0 ➡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ➡ L2 → + ∃∃T. L1 ⊢ ⓓ{a}W1.ⓐU1.T1 ➡ T & L2 ⊢ ⓓ{a}W2.ⓐU2.T2 ➡ T. +#a #L0 #V0 #W0 #T0 #IH #V1 #HV01 #U1 #HVU1 #W1 #HW01 #T1 #HT01 +#V2 #HV02 #U2 #HVU2 #W2 #HW02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1/ +elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1/ -L0 -V0 -W0 -T0 +elim (lift_total V 0 1) #U #HVU +lapply (cpr_lift … HV1 (L1.ⓓW1) … HVU1 … HVU) -HVU1 /2 width=1/ +lapply (cpr_lift … HV2 (L2.ⓓW2) … HVU2 … HVU) -HVU2 /2 width=1/ /4 width=7/ +qed-. + +theorem cpr_conf_lpr: lpx_sn_confluent cpr cpr. +#L0 #T0 @(f2_ind … fw … L0 T0) -L0 -T0 #n #IH #L0 * [|*] +[ #I0 #Hn #T1 #H1 #T2 #H2 #L1 #HL01 #L2 #HL02 destruct + elim (cpr_inv_atom1 … H1) -H1 + elim (cpr_inv_atom1 … H2) -H2 + [ #H2 #H1 destruct + /2 width=1 by cpr_conf_lpr_atom_atom/ + | * #K0 #V0 #V2 #i2 #HLK0 #HV02 #HVT2 #H2 #H1 destruct + /3 width=10 by cpr_conf_lpr_atom_delta/ + | #H2 * #K0 #V0 #V1 #i1 #HLK0 #HV01 #HVT1 #H1 destruct + /4 width=10 by ex2_commute, cpr_conf_lpr_atom_delta/ + | * #X #Y #V2 #z #H #HV02 #HVT2 #H2 + * #K0 #V0 #V1 #i #HLK0 #HV01 #HVT1 #H1 destruct + /3 width=17 by cpr_conf_lpr_delta_delta/ + ] +| #a #I #V0 #T0 #Hn #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct + elim (cpr_inv_bind1 … H1) -H1 * + [ #V1 #T1 #HV01 #HT01 #H1 + | #T1 #HT01 #HXT1 #H11 #H12 + ] + elim (cpr_inv_bind1 … H2) -H2 * + [1,3: #V2 #T2 #HV02 #HT02 #H2 + |2,4: #T2 #HT02 #HXT2 #H21 #H22 + ] destruct + [ /3 width=10 by cpr_conf_lpr_bind_bind/ + | /4 width=11 by ex2_commute, cpr_conf_lpr_bind_zeta/ + | /3 width=11 by cpr_conf_lpr_bind_zeta/ + | /3 width=12 by cpr_conf_lpr_zeta_zeta/ + ] +| #I #V0 #T0 #Hn #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct + elim (cpr_inv_flat1 … H1) -H1 * + [ #V1 #T1 #HV01 #HT01 #H1 + | #HX1 #H1 + | #a1 #V1 #Y1 #Z1 #T1 #HV01 #HZT1 #H11 #H12 #H13 + | #a1 #V1 #U1 #Y1 #W1 #Z1 #T1 #HV01 #HVU1 #HYW1 #HZT1 #H11 #H12 #H13 + ] + elim (cpr_inv_flat1 … H2) -H2 * + [1,5,9,13: #V2 #T2 #HV02 #HT02 #H2 + |2,6,10,14: #HX2 #H2 + |3,7,11,15: #a2 #V2 #Y2 #Z2 #T2 #HV02 #HZT2 #H21 #H22 #H23 + |4,8,12,16: #a2 #V2 #U2 #Y2 #W2 #Z2 #T2 #HV02 #HVU2 #HYW2 #HZT2 #H21 #H22 #H23 + ] destruct + [ /3 width=10 by cpr_conf_lpr_flat_flat/ + | /4 width=8 by ex2_commute, cpr_conf_lpr_flat_tau/ + | /4 width=11 by ex2_commute, cpr_conf_lpr_flat_beta/ + | /4 width=14 by ex2_commute, cpr_conf_lpr_flat_theta/ + | /3 width=8 by cpr_conf_lpr_flat_tau/ + | /3 width=7 by cpr_conf_lpr_tau_tau/ + | /3 width=11 by cpr_conf_lpr_flat_beta/ + | /3 width=11 by cpr_conf_lpr_beta_beta/ + | /3 width=14 by cpr_conf_lpr_flat_theta/ + | /3 width=17 by cpr_conf_lpr_theta_theta/ + ] +] +qed-. + +(* Basic_1: includes: pr0_confluence pr2_confluence *) +theorem cpr_conf: ∀L. confluent … (cpr L). +/2 width=6 by cpr_conf_lpr/ qed-. + +(* Properties on context-sensitive parallel reduction for terms *************) + +lemma lpr_cpr_conf_dx: ∀L0,T0,T1. L0 ⊢ T0 ➡ T1 → ∀L1. L0 ⊢ ➡ L1 → + ∃∃T. L1 ⊢ T0 ➡ T & L1 ⊢ T1 ➡ T. +#L0 #T0 #T1 #HT01 #L1 #HL01 +elim (cpr_conf_lpr … HT01 T0 … HL01 … HL01) // -L0 /2 width=3/ +qed-. + +lemma lpr_cpr_conf_sn: ∀L0,T0,T1. L0 ⊢ T0 ➡ T1 → ∀L1. L0 ⊢ ➡ L1 → + ∃∃T. L1 ⊢ T0 ➡ T & L0 ⊢ T1 ➡ T. +#L0 #T0 #T1 #HT01 #L1 #HL01 +elim (cpr_conf_lpr … HT01 T0 … L0 … HL01) // -HT01 -HL01 /2 width=3/ +qed-. + +lemma fsup_cpr_trans: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ➡ U2 → + ∃∃L,U1. L1 ⊢ ➡ L & L ⊢ T1 ➡ U1 & ⦃L, U1⦄ ⊃ ⦃L2, U2⦄. +#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [1,2,3,4,5: /3 width=5/ ] +#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 +elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 +elim (lift_total T d e) #U #HTU +elim (ldrop_lpr_trans … HLK1 … HK1) -HLK1 -HK1 #L2 #HL12 #HL2K +lapply (cpr_lift … HT1 … HL2K … HTU1 … HTU) -HT1 -HTU1 /3 width=11/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_cpss.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_cpss.ma new file mode 100644 index 000000000..c5cb57721 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_cpss.ma @@ -0,0 +1,270 @@ +(**************************************************************************) +(* ___ *) +(* ||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/substitution/lpss_ldrop.ma". +include "basic_2/reduction/lpr_ldrop.ma". + +(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************) + +(* Properties on context-sensitive parallel substitution for terms **********) + +fact cpr_cpss_conf_lpr_lpss_atom_atom: + ∀I,L1,L2. ∃∃T. L1 ⊢ ⓪{I} ▶* T & L2 ⊢ ⓪{I} ➡ T. +/2 width=3/ qed-. + +fact cpr_cpss_conf_lpr_lpss_atom_delta: + ∀L0,i. ( + ∀L,T.♯{L, T} < ♯{L0, #i} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀K0,V0. ⇩[O, i] L0 ≡ K0.ⓓV0 → + ∀V2. K0 ⊢ V0 ▶* V2 → ∀T2. ⇧[O, i + 1] V2 ≡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ #i ▶* T & L2 ⊢ T2 ➡ T. +#L0 #i #IH #K0 #V0 #HLK0 #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02 +elim (lpr_ldrop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1 +elim (lpr_inv_pair1 … H1) -H1 #K1 #V1 #HK01 #HV01 #H destruct +elim (lpss_ldrop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2 +elim (lpss_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct +lapply (ldrop_fwd_ldrop2 … HLK2) -W2 #HLK2 +lapply (ldrop_pair2_fwd_fw … HLK0 (#i)) -HLK0 #HLK0 +elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2 +elim (lift_total V 0 (i+1)) #T #HVT +lapply (cpr_lift … HV2 … HLK2 … HVT2 … HVT) -K2 -V2 /3 width=6/ +qed-. + +fact cpr_cpss_conf_lpr_lpss_delta_atom: + ∀L0,i. ( + ∀L,T.♯{L, T} < ♯{L0, #i} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀K0,V0. ⇩[O, i] L0 ≡ K0.ⓓV0 → + ∀V1. K0 ⊢ V0 ➡ V1 → ∀T1. ⇧[O, i + 1] V1 ≡ T1 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ T1 ▶* T & L2 ⊢ #i ➡ T. +#L0 #i #IH #K0 #V0 #HLK0 #V1 #HV01 #T1 #HVT1 #L1 #HL01 #L2 #HL02 +elim (lpss_ldrop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2 +elim (lpss_inv_pair1 … H2) -H2 #K2 #V2 #HK02 #HV02 #H destruct +elim (lpr_ldrop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1 +elim (lpr_inv_pair1 … H1) -H1 #K1 #W1 #HK01 #_ #H destruct +lapply (ldrop_fwd_ldrop2 … HLK1) -W1 #HLK1 +lapply (ldrop_pair2_fwd_fw … HLK0 (#i)) -HLK0 #HLK0 +elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2 +elim (lift_total V 0 (i+1)) #T #HVT +lapply (cpss_lift … HV1 … HLK1 … HVT1 … HVT) -K1 -V1 /3 width=9/ +qed-. + +fact cpr_cpss_conf_lpr_lpss_delta_delta: + ∀L0,i. ( + ∀L,T.♯{L, T} < ♯{L0, #i} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀K0,V0. ⇩[O, i] L0 ≡ K0.ⓓV0 → + ∀V1. K0 ⊢ V0 ➡ V1 → ∀T1. ⇧[O, i + 1] V1 ≡ T1 → + ∀KX,VX. ⇩[O, i] L0 ≡ KX.ⓓVX → + ∀V2. KX ⊢ VX ▶* V2 → ∀T2. ⇧[O, i + 1] V2 ≡ T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ T1 ▶* T & L2 ⊢ T2 ➡ T. +#L0 #i #IH #K0 #V0 #HLK0 #V1 #HV01 #T1 #HVT1 +#KX #VX #H #V2 #HV02 #T2 #HVT2 #L1 #HL01 #L2 #HL02 +lapply (ldrop_mono … H … HLK0) -H #H destruct +elim (lpr_ldrop_conf … HLK0 … HL01) -HL01 #X1 #H1 #HLK1 +elim (lpr_inv_pair1 … H1) -H1 #K1 #W1 #HK01 #_ #H destruct +lapply (ldrop_fwd_ldrop2 … HLK1) -W1 #HLK1 +elim (lpss_ldrop_conf … HLK0 … HL02) -HL02 #X2 #H2 #HLK2 +elim (lpss_inv_pair1 … H2) -H2 #K2 #W2 #HK02 #_ #H destruct +lapply (ldrop_fwd_ldrop2 … HLK2) -W2 #HLK2 +lapply (ldrop_pair2_fwd_fw … HLK0 (#i)) -HLK0 #HLK0 +elim (IH … HLK0 … HV01 … HV02 … HK01 … HK02) -L0 -K0 -V0 #V #HV1 #HV2 +elim (lift_total V 0 (i+1)) #T #HVT +lapply (cpss_lift … HV1 … HLK1 … HVT1 … HVT) -K1 -V1 +lapply (cpr_lift … HV2 … HLK2 … HVT2 … HVT) -K2 -V2 -V /2 width=3/ +qed-. + +fact cpr_cpss_conf_lpr_lpss_bind_bind: + ∀a,I,L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓑ{a,I}V0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0.ⓑ{I}V0 ⊢ T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ▶* V2 → ∀T2. L0.ⓑ{I}V0 ⊢ T0 ▶* T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ ⓑ{a,I}V1.T1 ▶* T & L2 ⊢ ⓑ{a,I}V2.T2 ➡ T. +#a #I #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01 +#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HV01 … HV02 … HL01 … HL02) // +elim (IH … HT01 … HT02 (L1.ⓑ{I}V1) … (L2.ⓑ{I}V2)) -IH // /2 width=1/ /3 width=5/ +qed-. + +fact cpr_cpss_conf_lpr_lpss_bind_zeta: + ∀L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,+ⓓV0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀T1. L0.ⓓV0 ⊢ T0 ➡ T1 → ∀X1. ⇧[O, 1] X1 ≡ T1 → + ∀V2. L0 ⊢ V0 ▶* V2 → ∀T2. L0.ⓓV0 ⊢ T0 ▶* T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ X1 ▶* T & L2 ⊢ +ⓓV2.T2 ➡ T. +#L0 #V0 #T0 #IH #T1 #HT01 #X1 #HXT1 +#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HT01 … HT02 (L1.ⓓV2) … (L2.ⓓV2)) -IH -HT01 -HT02 // /2 width=1/ /3 width=1/ -L0 -V0 -T0 #T #HT1 #HT2 +elim (cpss_inv_lift1 … HT1 L1 … HXT1) -T1 /2 width=1/ /3 width=9/ +qed-. + +fact cpr_cpss_conf_lpr_lpss_flat_flat: + ∀I,L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓕ{I}V0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0 ⊢ T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ▶* V2 → ∀T2. L0 ⊢ T0 ▶* T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ ⓕ{I}V1.T1 ▶* T & L2 ⊢ ⓕ{I}V2.T2 ➡ T. +#I #L0 #V0 #T0 #IH #V1 #HV01 #T1 #HT01 +#V2 #HV02 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HV01 … HV02 … HL01 … HL02) // +elim (IH … HT01 … HT02 … HL01 … HL02) // /3 width=5/ +qed-. + +fact cpr_cpss_conf_lpr_lpss_flat_tau: + ∀L0,V0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓝV0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀T1. L0 ⊢ T0 ➡ T1 → ∀V2,T2. L0 ⊢ T0 ▶* T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ T1 ▶* T & L2 ⊢ ⓝV2.T2 ➡ T. +#L0 #V0 #T0 #IH #T1 #HT01 +#V2 #T2 #HT02 #L1 #HL01 #L2 #HL02 +elim (IH … HT01 … HT02 … HL01 … HL02) // -L0 -V0 -T0 /3 width=3/ +qed-. + +fact cpr_cpss_conf_lpr_lpss_flat_beta: + ∀a,L0,V0,W0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓐV0.ⓛ{a}W0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0.ⓛW0 ⊢ T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ▶* V2 → ∀T2. L0 ⊢ ⓛ{a}W0.T0 ▶* T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ ⓓ{a}V1.T1 ▶* T & L2 ⊢ ⓐV2.T2 ➡ T. +#a #L0 #V0 #W0 #T0 #IH #V1 #HV01 #T1 #HT01 +#V2 #HV02 #X #H #L1 #HL01 #L2 #HL02 +elim (cpss_inv_bind1 … H) -H #W2 #T2 #HW02 #HT02 #H destruct +elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (IH … HT01 … HT02 (L1.ⓛW2) … (L2.ⓛW2)) /2 width=1/ /3 width=1/ -L0 -V0 -W0 -T0 #T #HT1 #HT2 +lapply (cpss_lsubr_trans … HT1 (L1.ⓓV1) ?) -HT1 /2 width=1/ /3 width=5/ +qed-. + +fact cpr_cpss_conf_lpr_lpss_flat_theta: + ∀a,L0,V0,W0,T0. ( + ∀L,T.♯{L,T} < ♯{L0,ⓐV0.ⓓ{a}W0.T0} → + ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ▶* T2 → + ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ▶* L2 → + ∃∃T0. L1 ⊢ T1 ▶* T0 & L2 ⊢ T2 ➡ T0 + ) → + ∀V1. L0 ⊢ V0 ➡ V1 → ∀U1. ⇧[O, 1] V1 ≡ U1 → + ∀W1. L0 ⊢ W0 ➡ W1 → ∀T1. L0.ⓓW0 ⊢ T0 ➡ T1 → + ∀V2. L0 ⊢ V0 ▶* V2 → ∀T2. L0 ⊢ ⓓ{a}W0.T0 ▶* T2 → + ∀L1. L0 ⊢ ➡ L1 → ∀L2. L0 ⊢ ▶* L2 → + ∃∃T. L1 ⊢ ⓓ{a}W1.ⓐU1.T1 ▶* T & L2 ⊢ ⓐV2.T2 ➡ T. +#a #L0 #V0 #W0 #T0 #IH #V1 #HV01 #U1 #HVU1 #W1 #HW01 #T1 #HT01 +#V2 #HV02 #X #H #L1 #HL01 #L2 #HL02 +elim (IH … HV01 … HV02 … HL01 … HL02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (lift_total V 0 1) #U #HVU +lapply (cpss_lift … HV1 (L1.ⓓW1) … HVU1 … HVU) -HVU1 /2 width=1/ #HU1 +elim (cpss_inv_bind1 … H) -H #W2 #T2 #HW02 #HT02 #H destruct +elim (IH … HW01 … HW02 … HL01 … HL02) /2 width=1/ +elim (IH … HT01 … HT02 (L1.ⓓW1) … (L2.ⓓW2)) /2 width=1/ -L0 -V0 -W0 -T0 +/4 width=9 by ex2_intro, cpr_theta, cpss_bind, cpss_flat/ (**) (* auto too slow without trace *) +qed-. + +lemma cpr_cpss_conf_lpr_lpss: lpx_sn_confluent cpr cpss. +#L0 #T0 @(f2_ind … fw … L0 T0) -L0 -T0 #n #IH #L0 * [|*] +[ #I0 #Hn #T1 #H1 #T2 #H2 #L1 #HL01 #L2 #HL02 destruct + elim (cpr_inv_atom1 … H1) -H1 + elim (cpss_inv_atom1 … H2) -H2 + [ #H2 #H1 destruct + /2 width=1 by cpr_cpss_conf_lpr_lpss_atom_atom/ + | * #K0 #V0 #V2 #i2 #HLK0 #HV02 #HVT2 #H2 #H1 destruct + /3 width=10 by cpr_cpss_conf_lpr_lpss_atom_delta/ + | #H2 * #K0 #V0 #V1 #i1 #HLK0 #HV01 #HVT1 #H1 destruct + /3 width=10 by cpr_cpss_conf_lpr_lpss_delta_atom/ + | * #X #Y #V2 #z #H #HV02 #HVT2 #H2 + * #K0 #V0 #V1 #i #HLK0 #HV01 #HVT1 #H1 destruct + /3 width=17 by cpr_cpss_conf_lpr_lpss_delta_delta/ + ] +| #a #I #V0 #T0 #Hn #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct + elim (cpss_inv_bind1 … H2) -H2 #V2 #T2 #HV02 #HT02 #H2 + elim (cpr_inv_bind1 … H1) -H1 * + [ #V1 #T1 #HV01 #HT01 #H1 + | #T1 #HT01 #HXT1 #H11 #H12 + ] destruct + [ /3 width=10 by cpr_cpss_conf_lpr_lpss_bind_bind/ + | /3 width=11 by cpr_cpss_conf_lpr_lpss_bind_zeta/ + ] +| #I #V0 #T0 #Hn #X1 #H1 #X2 #H2 #L1 #HL01 #L2 #HL02 destruct + elim (cpss_inv_flat1 … H2) -H2 #V2 #T2 #HV02 #HT02 #H2 + elim (cpr_inv_flat1 … H1) -H1 * + [ #V1 #T1 #HV01 #HT01 #H1 + | #HX1 #H1 + | #a1 #V1 #Y1 #Z1 #T1 #HV01 #HZT1 #H11 #H12 #H13 + | #a1 #V1 #U1 #Y1 #W1 #Z1 #T1 #HV01 #HVU1 #HYW1 #HZT1 #H11 #H12 #H13 + ] destruct + [ /3 width=10 by cpr_cpss_conf_lpr_lpss_flat_flat/ + | /3 width=8 by cpr_cpss_conf_lpr_lpss_flat_tau/ + | /3 width=11 by cpr_cpss_conf_lpr_lpss_flat_beta/ + | /3 width=14 by cpr_cpss_conf_lpr_lpss_flat_theta/ + ] +] +qed-. + +(* Basic_1: includes: pr0_subst1 *) +lemma cpr_cpss_conf: ∀L. confluent2 … (cpr L) (cpss L). +/2 width=6 by cpr_cpss_conf_lpr_lpss/ qed-. + +lemma cpr_lpss_conf_dx: ∀L0,T0,T1. L0 ⊢ T0 ➡ T1 → ∀L1. L0 ⊢ ▶* L1 → + ∃∃T. L1 ⊢ T1 ▶* T & L1 ⊢ T0 ➡ T. +#L0 #T0 #T1 #HT01 #L1 #HL01 +elim (cpr_cpss_conf_lpr_lpss … HT01 T0 … L1 … HL01) // /2 width=1/ -L0 /2 width=3/ +qed-. + +lemma cpr_lpss_conf_sn: ∀L0,T0,T1. L0 ⊢ T0 ➡ T1 → ∀L1. L0 ⊢ ▶* L1 → + ∃∃T. L0 ⊢ T1 ▶* T & L1 ⊢ T0 ➡ T. +#L0 #T0 #T1 #HT01 #L1 #HL01 +elim (cpr_cpss_conf_lpr_lpss … HT01 T0 … L0 … HL01) // -HT01 -HL01 /2 width=3/ +qed-. + +(* Basic_1: includes: pr0_subst1_fwd *) +lemma lpr_cpss_conf: ∀L0,T0,T1. L0 ⊢ T0 ▶* T1 → ∀L1. L0 ⊢ ➡ L1 → + ∃∃T. L1 ⊢ T0 ▶* T & L0 ⊢ T1 ➡ T. +#L0 #T0 #T1 #HT01 #L1 #HL01 +elim (cpr_cpss_conf_lpr_lpss ?? T0 … HT01 … HL01 L0) // -HT01 -HL01 /2 width=3/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_ldrop.ma new file mode 100644 index 000000000..2ea40eeb4 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_ldrop.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||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/relocation/ldrop_lpx_sn.ma". +include "basic_2/reduction/cpr_lift.ma". +include "basic_2/reduction/lpr.ma". + +(* SN PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS *****************************) + +(* Properies on local environment slicing ***********************************) + +(* Basic_1: includes: wcpr0_drop *) +lemma lpr_ldrop_conf: dropable_sn lpr. +/3 width=5 by lpx_sn_deliftable_dropable, cpr_inv_lift1/ qed-. + +(* Basic_1: includes: wcpr0_drop_back *) +lemma ldrop_lpr_trans: dedropable_sn lpr. +/3 width=9 by lpx_sn_liftable_dedropable, cpr_lift/ qed-. + +lemma lpr_ldrop_trans_O1: dropable_dx lpr. +/2 width=3 by lpx_sn_dropable/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpr.ma new file mode 100644 index 000000000..ab77692dd --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpr.ma @@ -0,0 +1,23 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lpr_cpr.ma". + +(* SN PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ******************************) + +(* Main properties **********************************************************) + +theorem lpr_conf: confluent … lpr. +/3 width=6 by lpx_sn_conf, cpr_conf_lpr/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpss.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpss.ma new file mode 100644 index 000000000..34aaa4e12 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpss.ma @@ -0,0 +1,23 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lpr_cpss.ma". + +(* SN PARALLEL REDUCTION ON LOCAL ENVIRONMENTS ******************************) + +(* Properties on sn parallel substitution on local environments *************) + +lemma lpr_lpss_conf: confluent2 … lpr lpss. +/3 width=6 by lpx_sn_conf, cpr_cpss_conf_lpr_lpss/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/xpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/xpr.ma new file mode 100644 index 000000000..8cec0b050 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/xpr.ma @@ -0,0 +1,66 @@ +(**************************************************************************) +(* ___ *) +(* ||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/static/ssta.ma". +include "basic_2/reduction/cpr.ma". + +lemma arith1: ∀x,y,z,w. z < y → x + (y + w) - z = x + y - z + w. +#x #y #z #w #H