From: Ferruccio Guidi Date: Mon, 11 Mar 2013 19:29:23 +0000 (+0000) Subject: more results on lenv refinement for stratified native validity X-Git-Tag: make_still_working~1219 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=b01109cef1d0cf392539cffa35751d9d716296ac;p=helm.git more results on lenv refinement for stratified native validity --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpcs.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpcs.ma new file mode 100644 index 000000000..87a72e0fd --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpcs.ma @@ -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/equivalence/cpcs_cpcs.ma". +include "basic_2/dynamic/lsubsv.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************) + +(* Properties on context-sensitive parallel equivalence for terms ***********) + +lemma lsubsv_cpcs_trans: ∀h,g,L1,L2. h ⊢ L1 ⊩:⊑[g] L2 → + ∀T1,T2. L2 ⊢ T1 ⬌* T2 → L1 ⊢ T1 ⬌* T2. +/3 width=5 by lsubsv_fwd_lsubs2, cpcs_lsubs_trans/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_snv.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_snv.ma index f085cc3b5..18ca504d4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_snv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_snv.ma @@ -13,9 +13,9 @@ (**************************************************************************) include "basic_2/dynamic/lsubsv_ldrop.ma". +include "basic_2/dynamic/lsubsv_cpcs.ma". (* include "basic_2/dynamic/lsubsv_ssta.ma". -include "basic_2/dynamic/lsubsv_cpcs.ma". *) (* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************) @@ -26,28 +26,31 @@ axiom lsubsv_xprs_trans: ∀h,g,L1,L2. h ⊢ L1 ⊩:⊑[g] L2 → /3 width=3 by lsubsv_fwd_lsubss, lsubss_xprs_trans/ qed-. *) -axiom lsubsv_snv_trans: ∀h,g,L2,T. ⦃h, L2⦄ ⊩ T :[g] → +lemma lsubsv_snv_trans: ∀h,g,L2,T. ⦃h, L2⦄ ⊩ T :[g] → ∀L1. h ⊢ L1 ⊩:⊑[g] L2 → ⦃h, L1⦄ ⊩ T :[g]. -(* #h #g #L2 #T #H elim H -L2 -T // -[ #I2 #L2 #K2 #V2 #i #HLK2 #_ #IHV2 #L1 #HL12 +[ #I2 #L2 #K2 #W2 #i #HLK2 #_ #IHW2 #L1 #HL12 elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 - elim (lsubsv_inv_pair2 … H) -H * #K1 [ | -IHV2 ] + elim (lsubsv_inv_pair2 … H) -H * #K1 [ | -IHW2 ] [ #HK12 #H destruct /3 width=5/ - | #V1 #l #HV1 #_ #_ #_ #H destruct /2 width=5/ + | #W1 #V1 #l #HV1 #_ #_ #_ #_ #H #_ destruct /2 width=5/ ] | #a #I #L2 #V #T #_ #_ #IHV #IHT #L1 #HL12 /4 width=1/ | #a #L2 #V #W #W0 #T #U #l #_ #_ #HVW #HW0 #HTU #IHV #IHT #L1 #HL12 lapply (IHV … HL12) -IHV #HV lapply (IHT … HL12) -IHT #HT - lapply (lsubsv_ssta_trans … HVW … HL12) -HVW #HVW lapply (lsubsv_cprs_trans … HL12 … HW0) -HW0 #HW0 +(* + lapply (lsubsv_ssta_trans … HVW … HL12) -HVW #HVW lapply (lsubsv_xprs_trans … HL12 … HTU) -HL12 -HTU /2 width=8/ -| #L2 #W #T #U #l #_ #_ #HTU #HWU #IHW #IHT #L1 #HL12 +*) +| #L2 #W #T #U #l #_ #_ #HTU #HUW #IHW #IHT #L1 #HL12 lapply (IHW … HL12) -IHW #HW lapply (IHT … HL12) -IHT #HT + lapply (lsubsv_cpcs_trans … HL12 … HUW) -HUW #HUW +(* lapply (lsubsv_ssta_trans … HTU … HL12) -HTU #HTU - lapply (lsubsv_cpcs_trans … HL12 … HWU) -HL12 -HWU /2 width=4/ +-HL12 -HWU /2 width=4/ ] qed-. *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ssta.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ssta.ma new file mode 100644 index 000000000..f6ef16c40 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ssta.ma @@ -0,0 +1,50 @@ +(**************************************************************************) +(* ___ *) +(* ||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/equivalence/cpcs_cpcs.ma". +include "basic_2/dynamic/lsubsv_ldrop.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************) + +(* Properties on stratified native type assignment **************************) + +lemma lsubsv_ssta_trans: ∀h,g,L2,T,U2,l. ⦃h, L2⦄ ⊢ T •[g,l] U2 → + ∀L1. h ⊢ L1 ⊩:⊑[g] L2 → + ∃∃U1. L1 ⊢ U2 ⬌* U1 & ⦃h, L1⦄ ⊢ T •[g,l] U1. +#h #g #L2 #T #U2 #l #H elim H -L2 -T -U2 -l +[ #L2 #k #l #Hkl #L1 #_ /3 width=3/ +| #L2 #K2 #V2 #W2 #U2 #i #l #HLK2 #_ #HWU2 #IHVW2 #L1 #HL12 + elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 + elim (lsubsv_inv_pair2 … H) -H * + [ #K1 #HK12 #H destruct + elim (IHVW2 … HK12) -K2 #W1 #HW21 #H + elim (lift_total W1 0 (i+1)) #U1 #HWU1 + lapply (ldrop_fwd_ldrop2 … HLK1) /3 width=9/ + | #K1 #V1 #W1 #l0 #_ #_ #_ #_ #_ #_ #H destruct + ] +| #L2 #K2 #W2 #V2 #U2 #i #l #HLK2 #_ #HWU2 #IHWV2 #L1 #HL12 + elim (lsubsv_ldrop_O1_trans … HL12 … HLK2) -L2 #X #H #HLK1 + elim (lsubsv_inv_pair2 … H) -H * + [ #K1 #HK12 #H destruct + elim (IHWV2 … HK12) -K2 #V1 #_ #H -V2 /3 width=6/ + | #K1 #W1 #V1 #l0 #HV1 #HVW1 #HW21 #HW2 #HK12 #H #_ destruct + elim (IHWV2 … HK12) -K2 #V #_ #H + elim (lift_total W1 0 (i+1)) #U1 #HWU1 + lapply (ldrop_fwd_ldrop2 … HLK1) #HLK + @(ex2_intro … U1) + [ @(cpcs_lift … HLK … HWU2 … HWU1 HW21) + | @(ssta_ldef … HLK1 … HWU1) + (* + /3 width=9/ + *) \ No newline at end of file diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ltpr.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ltpr.ma index ba47b1e47..74732eab2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ltpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ltpr.ma @@ -73,5 +73,7 @@ fact snv_ltpr_tpr_aux: ∀h,g,L0,T0. lapply (IH1 … HT20 … (L2.ⓛW20) … HT202) [1,2: /2 width=1/ ] -IH1 -HT20 -HT202 #HT2 elim (IH3 … HVW1 … HL12 … HV12) // [2: /2 width=1/ ] -HV1 -HVW1 -HV12 #W200 #HVW200 #H lapply (cpcs_trans … HW210 … H) -W10 #HW200 +(* lapply (lsubsv_snv_trans … HT2 (L2.ⓓV2) ?) -L1 -HT2 /2 width=1/ /2 width=4/ - | \ No newline at end of file + | +*) \ No newline at end of file diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsubsv/lsubsv_cpcs.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsubsv/lsubsv_cpcs.etc deleted file mode 100644 index 87a72e0fd..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/lsubsv/lsubsv_cpcs.etc +++ /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/equivalence/cpcs_cpcs.ma". -include "basic_2/dynamic/lsubsv.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************) - -(* Properties on context-sensitive parallel equivalence for terms ***********) - -lemma lsubsv_cpcs_trans: ∀h,g,L1,L2. h ⊢ L1 ⊩:⊑[g] L2 → - ∀T1,T2. L2 ⊢ T1 ⬌* T2 → L1 ⊢ T1 ⬌* T2. -/3 width=5 by lsubsv_fwd_lsubs2, cpcs_lsubs_trans/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsubsv/lsubsv_ssta.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsubsv/lsubsv_ssta.etc deleted file mode 100644 index 1e5b5fddd..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/lsubsv/lsubsv_ssta.etc +++ /dev/null @@ -1,24 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/computation/xprs_lsubss.ma". -include "basic_2/dynamic/lsubsv.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************) - -(* Properties on stratified native type assignment **************************) - -axiom lsubsv_ssta_trans: ∀h,g,L2,T,U2,l. ⦃h, L2⦄ ⊢ T •[g,l] U2 → - ∀L1. h ⊢ L1 ⊩:⊑[g] L2 → - ∃∃U1. L1 ⊢ U2 ⬌* U1 & ⦃h, L1⦄ ⊢ T •[g,l] U1.