From: Ferruccio Guidi Date: Wed, 3 Sep 2014 14:59:43 +0000 (+0000) Subject: - the PARTIAL COMMIT continues, we issue the "reduction" component X-Git-Tag: make_still_working~852 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=472cb969d9a01a6d24eabc39ba20d1dc6adf1b04;hp=5117f4af452934db436f22326f35d90f757bdf8a;p=helm.git - the PARTIAL COMMIT continues, we issue the "reduction" component --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_aaa.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_aaa.etc new file mode 100644 index 000000000..dccb10f9b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_aaa.etc @@ -0,0 +1,24 @@ +(**************************************************************************) +(* ___ *) +(* ||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/da_sta.ma". +include "basic_2/static/sta_aaa.ma". + +(* DEGREE ASSIGNMENT FOR TERMS **********************************************) + +(* Properties on atomic arity assignment for terms **************************) + +lemma aaa_da: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l. +#h #g #G #L #T #A #H elim (aaa_sta h … H) -A /2 width=2 by sta_da/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_aaa.ma deleted file mode 100644 index dccb10f9b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_aaa.ma +++ /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/static/da_sta.ma". -include "basic_2/static/sta_aaa.ma". - -(* DEGREE ASSIGNMENT FOR TERMS **********************************************) - -(* Properties on atomic arity assignment for terms **************************) - -lemma aaa_da: ∀h,g,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l. -#h #g #G #L #T #A #H elim (aaa_sta h … H) -A /2 width=2 by sta_da/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_sta.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_sta.etc new file mode 100644 index 000000000..4f45a8165 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/da_sta.etc @@ -0,0 +1,86 @@ +(**************************************************************************) +(* ___ *) +(* ||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/sta.ma". +include "basic_2/static/da_da.ma". + +(* Properties on static type assignment for terms ***************************) + +lemma da_sta_conf: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → + ∀l. ⦃G, L⦄ ⊢ T ▪[h, g] l → ⦃G, L⦄ ⊢ U ▪[h, g] l-1. +#h #g #G #L #T #U #H elim H -G -L -T -U +[ #G #L #k #l #H + lapply (da_inv_sort … H) -H /3 width=1 by da_sort, deg_next/ +| #G #L #K #V #U #W #i #HLK #_ #HWU #IHVW #l #H + elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0 + lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct + lapply (drop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/ +| #G #L #K #W #V #U #i #HLK #_ #HWU #IHWV #l #H + elim (da_inv_lref … H) -H * #K0 #V0 [| #l0] #HLK0 #HV0 [| #H0 ] + lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct + lapply (drop_fwd_drop2 … HLK) -HLK /3 width=8 by da_lift/ +| #a #I #G #L #V #T #U #_ #IHTU #l #H + lapply (da_inv_bind … H) -H /3 width=1 by da_bind/ +| #G #L #V #T #U #_ #IHTU #l #H + lapply (da_inv_flat … H) -H /3 width=1 by da_flat/ +| #G #L #W #T #U #_ #IHTU #l #H + lapply (da_inv_flat … H) -H /2 width=1 by/ +] +qed-. + +lemma sta_da: ∀h,g,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → + ∃l. ⦃G, L⦄ ⊢ T ▪[h, g] l. +#h #g #G #L #T #U #H elim H -G -L -T -U +[ #G #L #k elim (deg_total h g k) /3 width=2 by da_sort, ex_intro/ +| #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5 by da_ldef, ex_intro/ +| #G #L #K #W #V #W0 #i #HLK #_ #_ * /3 width=5 by da_ldec, ex_intro/ +| #a #I #G #L #V #T #U #_ * /3 width=2 by da_bind, ex_intro/ +| #G #L #V #T #U #_ * /3 width=2 by da_flat, ex_intro/ +| #G #L #W #T #U #_ * /3 width=2 by da_flat, ex_intro/ +] +qed-. + +lemma sta_da_ge: ∀h,G,L,T,U,l0. ⦃G, L⦄ ⊢ T •[h] U → + ∃∃g,l. ⦃G, L⦄ ⊢ T ▪[h, g] l & l0 ≤ l. +#h #G #L #T #U #l0 #H elim H -G -L -T -U +[ /3 width=4 by da_sort, ex2_2_intro/ +| #G #L #K #V #W #W0 #i #HLK #_ #_ * /3 width=5 by da_ldef, ex2_2_intro/ +| #G #L #K #W #V #W0 #i #HLK #_ #_ * /4 width=5 by da_ldec, lt_to_le, le_S_S, ex2_2_intro/ +| #a #I #G #L #V #T #U #_ * /3 width=4 by da_bind, ex2_2_intro/ +| #G #L #V #T #U #_ * /3 width=4 by da_flat, ex2_2_intro/ +| #G #L #W #T #U #_ * /3 width=4 by da_flat, ex2_2_intro/ +] +qed-. + +(* Inversion lrmmas on static type assignment for terms *********************) + +lemma da_inv_sta: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l → + ∃U. ⦃G, L⦄ ⊢ T •[h] U. +#h #g #G #L #T #l #H elim H -G -L -T -l +[ /2 width=2/ +| #G #L #K #V #i #l #HLK #_ * #W #HVW + elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldef, ex_intro/ +| #G #L #K #W #i #l #HLK #_ * #V #HWV + elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldec, ex_intro/ +| #a #I #G #L #V #T #l #_ * /3 width=2 by sta_bind, ex_intro/ +| * #G #L #V #T #l #_ * /3 width=2 by sta_appl, sta_cast, ex_intro/ +] +qed-. + +lemma sta_inv_refl_pos: ∀h,g,G,L,T,l. ⦃G, L⦄ ⊢ T ▪[h, g] l+1 → ⦃G, L⦄ ⊢ T •[h] T → ⊥. +#h #g #G #L #T #l #H1T #HTT +lapply (da_sta_conf … HTT … H1T) -HTT iter_SO // +qed-. + +lemma lstas_inv_gref1: ∀h,G,L,X,p,l. ⦃G, L⦄ ⊢ §p •*[h, l+1] X → ⊥. +#h #G #L #X #p #l #H elim (lstas_inv_step_sn … H) -H +#U #H #HUX elim (sta_inv_gref1 … H) +qed-. + +lemma lstas_inv_bind1: ∀h,a,I,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] X → + ∃∃U. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U & X = ⓑ{a,I}V.U. +#h #a #I #G #L #V #T #X #l #H @(lstas_ind_dx … H) -X -l /2 width=3 by ex2_intro/ +#l #X #X0 #_ #HX0 * #U #HTU #H destruct +elim (sta_inv_bind1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3 by lstar_dx, ex2_intro/ +qed-. + +lemma lstas_inv_appl1: ∀h,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓐV.T •*[h, l] X → + ∃∃U. ⦃G, L⦄ ⊢ T •*[h, l] U & X = ⓐV.U. +#h #G #L #V #T #X #l #H @(lstas_ind_dx … H) -X -l /2 width=3 by ex2_intro/ +#l #X #X0 #_ #HX0 * #U #HTU #H destruct +elim (sta_inv_appl1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3 by lstar_dx, ex2_intro/ +qed-. + +lemma lstas_inv_cast1: ∀h,G,L,W,T,U,l. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l+1] U → ⦃G, L⦄ ⊢ T •*[h, l+1] U. +#h #G #L #W #T #X #l #H elim (lstas_inv_step_sn … H) -H +#U #H #HUX lapply (sta_inv_cast1 … H) -H /2 width=3 by lstar_S/ +qed-. + +(* Basic properties *********************************************************) + +lemma lstas_refl: ∀h,G,L. reflexive … (lstas h G L 0). +// qed. + +lemma sta_lstas: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ T •*[h, 1] U. +/2 width=1 by lstar_step/ qed. + +lemma lstas_step_sn: ∀h,G,L,T1,U1,U2,l. ⦃G, L⦄ ⊢ T1 •[h] U1 → ⦃G, L⦄ ⊢ U1 •*[h, l] U2 → + ⦃G, L⦄ ⊢ T1 •*[h, l+1] U2. +/2 width=3 by lstar_S/ qed. + +lemma lstas_step_dx: ∀h,G,L,T1,T2,U2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 → ⦃G, L⦄ ⊢ T2 •[h] U2 → + ⦃G, L⦄ ⊢ T1 •*[h, l+1] U2. +/2 width=3 by lstar_dx/ qed. + +lemma lstas_split: ∀h,G,L. inv_ltransitive … (lstas h G L). +/2 width=1 by lstar_inv_ltransitive/ qed-. + +lemma lstas_sort: ∀h,G,L,l,k. ⦃G, L⦄ ⊢ ⋆k •*[h, l] ⋆((next h)^l k). +#h #G #L #l @(nat_ind_plus … l) -l // +#l #IHl #k >iter_SO /2 width=3 by sta_sort, lstas_step_dx/ +qed. + +lemma lstas_bind: ∀h,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] ⓑ{a,I}V.U. +#h #I #G #L #V #T #U #l #H @(lstas_ind_dx … H) -U -l /3 width=3 by sta_bind, lstar_O, lstas_step_dx/ +qed. + +lemma lstas_appl: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → + ∀V.⦃G, L⦄ ⊢ ⓐV.T •*[h, l] ⓐV.U. +#h #G #L #T #U #l #H @(lstas_ind_dx … H) -U -l /3 width=3 by sta_appl, lstar_O, lstas_step_dx/ +qed. + +lemma lstas_cast: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l+1] U → + ∀W. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l+1] U. +#h #G #L #T #U #l #H elim (lstas_inv_step_sn … H) -H /3 width=3 by sta_cast, lstas_step_sn/ +qed. + +(* Basic_1: removed theorems 7: + sty1_abbr sty1_appl sty1_bind sty1_cast2 + sty1_correct sty1_lift sty1_trans +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas.ma deleted file mode 100644 index feed03e3f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas.ma +++ /dev/null @@ -1,133 +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/notation/relations/statictypestar_6.ma". -include "basic_2/static/sta.ma". - -(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************) - -definition lstas: ∀h. genv → lenv → nat → relation term ≝ - λh,G,L. lstar … (sta h G L). - -interpretation "nat-iterated static type assignment (term)" - 'StaticTypeStar h G L l T U = (lstas h G L l T U). - -(* Basic eliminators ********************************************************) - -lemma lstas_ind_sn: ∀h,G,L,U2. ∀R:relation2 nat term. - R 0 U2 → ( - ∀l,T,U1. ⦃G, L⦄ ⊢ T •[h] U1 → ⦃G, L⦄ ⊢ U1 •* [h, l] U2 → - R l U1 → R (l+1) T - ) → - ∀l,T. ⦃G, L⦄ ⊢ T •*[h, l] U2 → R l T. -/3 width=5 by lstar_ind_l/ qed-. - -lemma lstas_ind_dx: ∀h,G,L,T. ∀R:relation2 nat term. - R 0 T → ( - ∀l,U1,U2. ⦃G, L⦄ ⊢ T •* [h, l] U1 → ⦃G, L⦄ ⊢ U1 •[h] U2 → - R l U1 → R (l+1) U2 - ) → - ∀l,U. ⦃G, L⦄ ⊢ T •*[h, l] U → R l U. -/3 width=5 by lstar_ind_r/ qed-. - -(* Basic inversion lemmas ***************************************************) - -lemma lstas_inv_O: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, 0] U → T = U. -/2 width=4 by lstar_inv_O/ qed-. - -lemma lstas_inv_SO: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •*[h, 1] U → ⦃G, L⦄ ⊢ T •[h] U. -/2 width=1 by lstar_inv_step/ qed-. - -lemma lstas_inv_step_sn: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 → - ∃∃T. ⦃G, L⦄ ⊢ T1 •[h] T & ⦃G, L⦄ ⊢ T •*[h, l] T2. -/2 width=3 by lstar_inv_S/ qed-. - -lemma lstas_inv_step_dx: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 → - ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, l] T & ⦃G, L⦄ ⊢ T •[h] T2. -/2 width=3 by lstar_inv_S_dx/ qed-. - -lemma lstas_inv_sort1: ∀h,G,L,X,k,l. ⦃G, L⦄ ⊢ ⋆k •*[h, l] X → X = ⋆((next h)^l k). -#h #G #L #X #k #l #H @(lstas_ind_dx … H) -X -l // -#l #X #X0 #_ #H #IHX destruct -lapply (sta_inv_sort1 … H) -H #H destruct ->iter_SO // -qed-. - -lemma lstas_inv_gref1: ∀h,G,L,X,p,l. ⦃G, L⦄ ⊢ §p •*[h, l+1] X → ⊥. -#h #G #L #X #p #l #H elim (lstas_inv_step_sn … H) -H -#U #H #HUX elim (sta_inv_gref1 … H) -qed-. - -lemma lstas_inv_bind1: ∀h,a,I,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] X → - ∃∃U. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U & X = ⓑ{a,I}V.U. -#h #a #I #G #L #V #T #X #l #H @(lstas_ind_dx … H) -X -l /2 width=3 by ex2_intro/ -#l #X #X0 #_ #HX0 * #U #HTU #H destruct -elim (sta_inv_bind1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3 by lstar_dx, ex2_intro/ -qed-. - -lemma lstas_inv_appl1: ∀h,G,L,V,T,X,l. ⦃G, L⦄ ⊢ ⓐV.T •*[h, l] X → - ∃∃U. ⦃G, L⦄ ⊢ T •*[h, l] U & X = ⓐV.U. -#h #G #L #V #T #X #l #H @(lstas_ind_dx … H) -X -l /2 width=3 by ex2_intro/ -#l #X #X0 #_ #HX0 * #U #HTU #H destruct -elim (sta_inv_appl1 … HX0) -HX0 #U0 #HU0 #H destruct /3 width=3 by lstar_dx, ex2_intro/ -qed-. - -lemma lstas_inv_cast1: ∀h,G,L,W,T,U,l. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l+1] U → ⦃G, L⦄ ⊢ T •*[h, l+1] U. -#h #G #L #W #T #X #l #H elim (lstas_inv_step_sn … H) -H -#U #H #HUX lapply (sta_inv_cast1 … H) -H /2 width=3 by lstar_S/ -qed-. - -(* Basic properties *********************************************************) - -lemma lstas_refl: ∀h,G,L. reflexive … (lstas h G L 0). -// qed. - -lemma sta_lstas: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ T •*[h, 1] U. -/2 width=1 by lstar_step/ qed. - -lemma lstas_step_sn: ∀h,G,L,T1,U1,U2,l. ⦃G, L⦄ ⊢ T1 •[h] U1 → ⦃G, L⦄ ⊢ U1 •*[h, l] U2 → - ⦃G, L⦄ ⊢ T1 •*[h, l+1] U2. -/2 width=3 by lstar_S/ qed. - -lemma lstas_step_dx: ∀h,G,L,T1,T2,U2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 → ⦃G, L⦄ ⊢ T2 •[h] U2 → - ⦃G, L⦄ ⊢ T1 •*[h, l+1] U2. -/2 width=3 by lstar_dx/ qed. - -lemma lstas_split: ∀h,G,L. inv_ltransitive … (lstas h G L). -/2 width=1 by lstar_inv_ltransitive/ qed-. - -lemma lstas_sort: ∀h,G,L,l,k. ⦃G, L⦄ ⊢ ⋆k •*[h, l] ⋆((next h)^l k). -#h #G #L #l @(nat_ind_plus … l) -l // -#l #IHl #k >iter_SO /2 width=3 by sta_sort, lstas_step_dx/ -qed. - -lemma lstas_bind: ∀h,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U → - ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T •*[h, l] ⓑ{a,I}V.U. -#h #I #G #L #V #T #U #l #H @(lstas_ind_dx … H) -U -l /3 width=3 by sta_bind, lstar_O, lstas_step_dx/ -qed. - -lemma lstas_appl: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → - ∀V.⦃G, L⦄ ⊢ ⓐV.T •*[h, l] ⓐV.U. -#h #G #L #T #U #l #H @(lstas_ind_dx … H) -U -l /3 width=3 by sta_appl, lstar_O, lstas_step_dx/ -qed. - -lemma lstas_cast: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l+1] U → - ∀W. ⦃G, L⦄ ⊢ ⓝW.T •*[h, l+1] U. -#h #G #L #T #U #l #H elim (lstas_inv_step_sn … H) -H /3 width=3 by sta_cast, lstas_step_sn/ -qed. - -(* Basic_1: removed theorems 7: - sty1_abbr sty1_appl sty1_bind sty1_cast2 - sty1_correct sty1_lift sty1_trans -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_aaa.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_aaa.etc new file mode 100644 index 000000000..7497f44b6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_aaa.etc @@ -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/static/sta_aaa.ma". +include "basic_2/unfold/lstas.ma". + +(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************) + +(* Properties on atomic arity assignment for terms **************************) + +lemma lstas_aaa_conf: ∀h,G,L,l. Conf3 … (aaa G L) (lstas h G L l). +/3 width=6 by sta_aaa_conf, lstar_Conf3/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_aaa.ma deleted file mode 100644 index 7497f44b6..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_aaa.ma +++ /dev/null @@ -1,23 +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/sta_aaa.ma". -include "basic_2/unfold/lstas.ma". - -(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************) - -(* Properties on atomic arity assignment for terms **************************) - -lemma lstas_aaa_conf: ∀h,G,L,l. Conf3 … (aaa G L) (lstas h G L l). -/3 width=6 by sta_aaa_conf, lstar_Conf3/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_alt.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_alt.etc new file mode 100644 index 000000000..0ad7b9bbc --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_alt.etc @@ -0,0 +1,102 @@ +(**************************************************************************) +(* ___ *) +(* ||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/notation/relations/statictypestaralt_6.ma". +include "basic_2/unfold/lstas_lift.ma". + +(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************) + +(* alternative definition of lstas *) +inductive lstasa (h): genv → relation4 lenv nat term term ≝ +| lstasa_O : ∀G,L,T. lstasa h G L 0 T T +| lstasa_sort: ∀G,L,l,k. lstasa h G L l (⋆k) (⋆((next h)^l k)) +| lstasa_ldef: ∀G,L,K,V,W,U,i,l. ⇩[i] L ≡ K.ⓓV → lstasa h G K (l+1) V W → + ⇧[0, i+1] W ≡ U → lstasa h G L (l+1) (#i) U +| lstasa_ldec: ∀G,L,K,W,V,V0,U,i,l. ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W •[h] V0 → + lstasa h G K l W V → ⇧[0, i+1] V ≡ U → lstasa h G L (l+1) (#i) U +| lstasa_bind: ∀a,I,G,L,V,T,U,l. lstasa h G (L.ⓑ{I}V) l T U → + lstasa h G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U) +| lstasa_appl: ∀G,L,V,T,U,l. lstasa h G L l T U → lstasa h G L l (ⓐV.T) (ⓐV.U) +| lstasa_cast: ∀G,L,W,T,U,l. lstasa h G L (l+1) T U → lstasa h G L (l+1) (ⓝW.T) U +. + +interpretation "nat-iterated static type assignment (term) alternative" + 'StaticTypeStarAlt h G L l T U = (lstasa h G L l T U). + +(* Base properties **********************************************************) + +lemma sta_lstasa: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ T ••*[h, 1] U. +#h #G #L #T #U #H elim H -G -L -T -U +/2 width=8 by lstasa_O, lstasa_sort, lstasa_ldef, lstasa_ldec, lstasa_bind, lstasa_appl, lstasa_cast/ +qed. + +lemma lstasa_step_dx: ∀h,G,L,T1,T,l. ⦃G, L⦄ ⊢ T1 ••*[h, l] T → + ∀T2. ⦃G, L⦄ ⊢ T •[h] T2 → ⦃G, L⦄ ⊢ T1 ••*[h, l+1] T2. +#h #G #L #T1 #T #l #H elim H -G -L -T1 -T -l +[ /2 width=1 by sta_lstasa/ +| #G #L #l #k #X #H >(sta_inv_sort1 … H) -X >commutative_plus // +| #G #L #K #V #W #U #i #l #HLK #_ #HWU #IHVW #U2 #HU2 + lapply (drop_fwd_drop2 … HLK) #H + elim (sta_inv_lift1 … HU2 … H … HWU) -H -U /3 width=6 by lstasa_ldef/ +| #G #L #K #W #V #V0 #U #i #l #HLK #HWl0 #_ #HVU #IHWV #U2 #HU2 + lapply (drop_fwd_drop2 … HLK) #H + elim (sta_inv_lift1 … HU2 … H … HVU) -H -U /3 width=8 by lstasa_ldec/ +| #a #I #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H + elim (sta_inv_bind1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_bind/ +| #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H + elim (sta_inv_appl1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_appl/ +| /3 width=1 by lstasa_cast/ +] +qed. + +(* Main properties **********************************************************) + +theorem lstas_lstasa: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ⦃G, L⦄ ⊢ T ••*[h, l] U. +#h #G #L #T #U #l #H @(lstas_ind_dx … H) -U -l /2 width=3 by lstasa_step_dx, lstasa_O/ +qed. + +(* Main inversion lemmas ****************************************************) + +theorem lstasa_inv_lstas: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T ••*[h, l] U → ⦃G, L⦄ ⊢ T •*[h, l] U. +#h #G #L #T #U #l #H elim H -G -L -T -U -l +/2 width=8 by lstas_inv_SO, lstas_ldec, lstas_ldef, lstas_cast, lstas_appl, lstas_bind/ +qed-. + +(* Advanced eliminators *****************************************************) + +lemma lstas_ind_alt: ∀h. ∀R:genv→relation4 lenv nat term term. + (∀G,L,T. R G L O T T) → + (∀G,L,l,k. R G L l (⋆k) (⋆((next h)^l k))) → ( + ∀G,L,K,V,W,U,i,l. + ⇩[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V •*[h, l+1] W → ⇧[O, i+1] W ≡ U → + R G K (l+1) V W → R G L (l+1) (#i) U + ) → ( + ∀G,L,K,W,V,V0,U,i,l. + ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W •[h] V0 → + ⦃G, K⦄ ⊢ W •*[h, l]V → ⇧[O, i+1] V ≡ U → + R G K l W V → R G L (l+1) (#i) U + ) → ( + ∀a,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U → + R G (L.ⓑ{I}V) l T U → R G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U) + ) → ( + ∀G,L,V,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → + R G L l T U → R G L l (ⓐV.T) (ⓐV.U) + ) → ( + ∀G,L,W,T,U,l. ⦃G, L⦄⊢ T •*[h, l+1] U → + R G L (l+1) T U → R G L (l+1) (ⓝW.T) U + ) → + ∀G,L,l,T,U. ⦃G, L⦄ ⊢ T •*[h, l] U → R G L l T U. +#h #R #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #G #L #l #T #U #H +elim (lstas_lstasa … H) /3 width=10 by lstasa_inv_lstas/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_alt.ma deleted file mode 100644 index 0ad7b9bbc..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_alt.ma +++ /dev/null @@ -1,102 +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/notation/relations/statictypestaralt_6.ma". -include "basic_2/unfold/lstas_lift.ma". - -(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************) - -(* alternative definition of lstas *) -inductive lstasa (h): genv → relation4 lenv nat term term ≝ -| lstasa_O : ∀G,L,T. lstasa h G L 0 T T -| lstasa_sort: ∀G,L,l,k. lstasa h G L l (⋆k) (⋆((next h)^l k)) -| lstasa_ldef: ∀G,L,K,V,W,U,i,l. ⇩[i] L ≡ K.ⓓV → lstasa h G K (l+1) V W → - ⇧[0, i+1] W ≡ U → lstasa h G L (l+1) (#i) U -| lstasa_ldec: ∀G,L,K,W,V,V0,U,i,l. ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W •[h] V0 → - lstasa h G K l W V → ⇧[0, i+1] V ≡ U → lstasa h G L (l+1) (#i) U -| lstasa_bind: ∀a,I,G,L,V,T,U,l. lstasa h G (L.ⓑ{I}V) l T U → - lstasa h G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U) -| lstasa_appl: ∀G,L,V,T,U,l. lstasa h G L l T U → lstasa h G L l (ⓐV.T) (ⓐV.U) -| lstasa_cast: ∀G,L,W,T,U,l. lstasa h G L (l+1) T U → lstasa h G L (l+1) (ⓝW.T) U -. - -interpretation "nat-iterated static type assignment (term) alternative" - 'StaticTypeStarAlt h G L l T U = (lstasa h G L l T U). - -(* Base properties **********************************************************) - -lemma sta_lstasa: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ⦃G, L⦄ ⊢ T ••*[h, 1] U. -#h #G #L #T #U #H elim H -G -L -T -U -/2 width=8 by lstasa_O, lstasa_sort, lstasa_ldef, lstasa_ldec, lstasa_bind, lstasa_appl, lstasa_cast/ -qed. - -lemma lstasa_step_dx: ∀h,G,L,T1,T,l. ⦃G, L⦄ ⊢ T1 ••*[h, l] T → - ∀T2. ⦃G, L⦄ ⊢ T •[h] T2 → ⦃G, L⦄ ⊢ T1 ••*[h, l+1] T2. -#h #G #L #T1 #T #l #H elim H -G -L -T1 -T -l -[ /2 width=1 by sta_lstasa/ -| #G #L #l #k #X #H >(sta_inv_sort1 … H) -X >commutative_plus // -| #G #L #K #V #W #U #i #l #HLK #_ #HWU #IHVW #U2 #HU2 - lapply (drop_fwd_drop2 … HLK) #H - elim (sta_inv_lift1 … HU2 … H … HWU) -H -U /3 width=6 by lstasa_ldef/ -| #G #L #K #W #V #V0 #U #i #l #HLK #HWl0 #_ #HVU #IHWV #U2 #HU2 - lapply (drop_fwd_drop2 … HLK) #H - elim (sta_inv_lift1 … HU2 … H … HVU) -H -U /3 width=8 by lstasa_ldec/ -| #a #I #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H - elim (sta_inv_bind1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_bind/ -| #G #L #V #T1 #U1 #l #_ #IHTU1 #X #H - elim (sta_inv_appl1 … H) -H #U #HU1 #H destruct /3 width=1 by lstasa_appl/ -| /3 width=1 by lstasa_cast/ -] -qed. - -(* Main properties **********************************************************) - -theorem lstas_lstasa: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → ⦃G, L⦄ ⊢ T ••*[h, l] U. -#h #G #L #T #U #l #H @(lstas_ind_dx … H) -U -l /2 width=3 by lstasa_step_dx, lstasa_O/ -qed. - -(* Main inversion lemmas ****************************************************) - -theorem lstasa_inv_lstas: ∀h,G,L,T,U,l. ⦃G, L⦄ ⊢ T ••*[h, l] U → ⦃G, L⦄ ⊢ T •*[h, l] U. -#h #G #L #T #U #l #H elim H -G -L -T -U -l -/2 width=8 by lstas_inv_SO, lstas_ldec, lstas_ldef, lstas_cast, lstas_appl, lstas_bind/ -qed-. - -(* Advanced eliminators *****************************************************) - -lemma lstas_ind_alt: ∀h. ∀R:genv→relation4 lenv nat term term. - (∀G,L,T. R G L O T T) → - (∀G,L,l,k. R G L l (⋆k) (⋆((next h)^l k))) → ( - ∀G,L,K,V,W,U,i,l. - ⇩[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V •*[h, l+1] W → ⇧[O, i+1] W ≡ U → - R G K (l+1) V W → R G L (l+1) (#i) U - ) → ( - ∀G,L,K,W,V,V0,U,i,l. - ⇩[i] L ≡ K.ⓛW → ⦃G, K⦄ ⊢ W •[h] V0 → - ⦃G, K⦄ ⊢ W •*[h, l]V → ⇧[O, i+1] V ≡ U → - R G K l W V → R G L (l+1) (#i) U - ) → ( - ∀a,I,G,L,V,T,U,l. ⦃G, L.ⓑ{I}V⦄ ⊢ T •*[h, l] U → - R G (L.ⓑ{I}V) l T U → R G L l (ⓑ{a,I}V.T) (ⓑ{a,I}V.U) - ) → ( - ∀G,L,V,T,U,l. ⦃G, L⦄ ⊢ T •*[h, l] U → - R G L l T U → R G L l (ⓐV.T) (ⓐV.U) - ) → ( - ∀G,L,W,T,U,l. ⦃G, L⦄⊢ T •*[h, l+1] U → - R G L (l+1) T U → R G L (l+1) (ⓝW.T) U - ) → - ∀G,L,l,T,U. ⦃G, L⦄ ⊢ T •*[h, l] U → R G L l T U. -#h #R #IH1 #IH2 #IH3 #IH4 #IH5 #IH6 #IH7 #G #L #l #T #U #H -elim (lstas_lstasa … H) /3 width=10 by lstasa_inv_lstas/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_da.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_da.etc new file mode 100644 index 000000000..8b11d1dc0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_da.etc @@ -0,0 +1,40 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lstas.ma". +include "basic_2/static/da_sta.ma". + +(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************) + +(* Properties on degree assignment for terms ********************************) + +lemma lstas_da_conf: ∀h,g,G,L,T,U,l1. ⦃G, L⦄ ⊢ T •*[h, l1] U → + ∀l2. ⦃G, L⦄ ⊢ T ▪[h, g] l2 → ⦃G, L⦄ ⊢ U ▪[h, g] l2-l1. +#h #g #G #L #T #U #l1 #H @(lstas_ind_dx … H) -U -l1 // +#l1 #U #U0 #_ #HU0 #IHTU #l2 #HT +(plus_minus_m_m … Hl12) in ⊢ (%→?); -Hl12 >commutative_plus #H +elim (lstas_split … H) -H #U #HTU +>(lstas_mono … HTU … HTU1) -T // +qed-. + +(* Advanced properties ******************************************************) + +lemma lstas_sta_conf_pos: ∀h,G,L,T,U1. ⦃G, L⦄ ⊢ T •[h] U1 → + ∀U2,l. ⦃G, L⦄ ⊢ T •*[h, l+1] U2 → ⦃G, L⦄ ⊢ U1 •*[h, l] U2. +#h #G #L #T #U1 #HTU1 #U2 #l #HTU2 +lapply (lstas_conf_le … T U1 1 … HTU2) -HTU2 /2 width=1 by sta_lstas/ +qed-. + +lemma lstas_strip_pos: ∀h,G,L,T1,U1. ⦃G, L⦄ ⊢ T1 •[h] U1 → + ∀T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 → + ∃∃U2. ⦃G, L⦄ ⊢ T2 •[h] U2 & ⦃G, L⦄ ⊢ U1 •*[h, l+1] U2. +#h #G #L #T1 #U1 #HTU1 #T2 #l #HT12 +elim (lstas_fwd_correct … HTU1 … HT12) +lapply (lstas_sta_conf_pos … HTU1 … HT12) -T1 /3 width=5 by lstas_step_dx, ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_lstas.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_lstas.ma deleted file mode 100644 index 58614e3ff..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/lstas_lstas.ma +++ /dev/null @@ -1,51 +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/sta_sta.ma". -include "basic_2/unfold/lstas_lift.ma". - -(* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************) - -(* Main properties **********************************************************) - -theorem lstas_trans: ∀h,G,L. ltransitive … (lstas h G L). -/2 width=3 by lstar_ltransitive/ qed-. - -theorem lstas_mono: ∀h,G,L,l. singlevalued … (lstas h G L l). -/3 width=7 by sta_mono, lstar_singlevalued/ qed-. - -theorem lstas_conf_le: ∀h,G,L,T,U1,l1. ⦃G, L⦄ ⊢ T •*[h, l1] U1 → - ∀U2,l2. l1 ≤ l2 → ⦃G, L⦄ ⊢ T •*[h, l2] U2 → - ⦃G, L⦄ ⊢ U1 •*[h, l2-l1] U2. -#h #G #L #T #U1 #l1 #HTU1 #U2 #l2 #Hl12 ->(plus_minus_m_m … Hl12) in ⊢ (%→?); -Hl12 >commutative_plus #H -elim (lstas_split … H) -H #U #HTU ->(lstas_mono … HTU … HTU1) -T // -qed-. - -(* Advanced properties ******************************************************) - -lemma lstas_sta_conf_pos: ∀h,G,L,T,U1. ⦃G, L⦄ ⊢ T •[h] U1 → - ∀U2,l. ⦃G, L⦄ ⊢ T •*[h, l+1] U2 → ⦃G, L⦄ ⊢ U1 •*[h, l] U2. -#h #G #L #T #U1 #HTU1 #U2 #l #HTU2 -lapply (lstas_conf_le … T U1 1 … HTU2) -HTU2 /2 width=1 by sta_lstas/ -qed-. - -lemma lstas_strip_pos: ∀h,G,L,T1,U1. ⦃G, L⦄ ⊢ T1 •[h] U1 → - ∀T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l+1] T2 → - ∃∃U2. ⦃G, L⦄ ⊢ T2 •[h] U2 & ⦃G, L⦄ ⊢ U1 •*[h, l+1] U2. -#h #G #L #T1 #U1 #HTU1 #T2 #l #HT12 -elim (lstas_fwd_correct … HTU1 … HT12) -lapply (lstas_sta_conf_pos … HTU1 … HT12) -T1 /3 width=5 by lstas_step_dx, ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta.etc new file mode 100644 index 000000000..fc6bf6a2a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta.etc @@ -0,0 +1,143 @@ +(**************************************************************************) +(* ___ *) +(* ||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/notation/relations/statictype_5.ma". +include "basic_2/grammar/genv.ma". +include "basic_2/substitution/drop.ma". +include "basic_2/static/sh.ma". + +(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************) + +(* activate genv *) +inductive sta (h:sh): relation4 genv lenv term term ≝ +| sta_sort: ∀G,L,k. sta h G L (⋆k) (⋆(next h k)) +| sta_ldef: ∀G,L,K,V,W,U,i. ⇩[i] L ≡ K.ⓓV → sta h G K V W → + ⇧[0, i + 1] W ≡ U → sta h G L (#i) U +| sta_ldec: ∀G,L,K,W,V,U,i. ⇩[i] L ≡ K.ⓛW → sta h G K W V → + ⇧[0, i + 1] W ≡ U → sta h G L (#i) U +| sta_bind: ∀a,I,G,L,V,T,U. sta h G (L.ⓑ{I}V) T U → + sta h G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U) +| sta_appl: ∀G,L,V,T,U. sta h G L T U → sta h G L (ⓐV.T) (ⓐV.U) +| sta_cast: ∀G,L,W,T,U. sta h G L T U → sta h G L (ⓝW.T) U +. + +interpretation "static type assignment (term)" + 'StaticType h G L T U = (sta h G L T U). + +(* Basic inversion lemmas ************************************************) + +fact sta_inv_sort1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀k0. T = ⋆k0 → + U = ⋆(next h k0). +#h #G #L #T #U * -G -L -T -U +[ #G #L #k #k0 #H destruct // +| #G #L #K #V #W #U #i #_ #_ #_ #k0 #H destruct +| #G #L #K #W #V #U #i #_ #_ #_ #k0 #H destruct +| #a #I #G #L #V #T #U #_ #k0 #H destruct +| #G #L #V #T #U #_ #k0 #H destruct +| #G #L #W #T #U #_ #k0 #H destruct +qed-. + +(* Basic_1: was: sty0_gen_sort *) +lemma sta_inv_sort1: ∀h,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k •[h] U → U = ⋆(next h k). +/2 width=5 by sta_inv_sort1_aux/ qed-. + +fact sta_inv_lref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀j. T = #j → + (∃∃K,V,W. ⇩[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W & + ⇧[0, j+1] W ≡ U + ) ∨ + (∃∃K,W,V. ⇩[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V & + ⇧[0, j+1] W ≡ U + ). +#h #G #L #T #U * -G -L -T -U +[ #G #L #k #j #H destruct +| #G #L #K #V #W #U #i #HLK #HVW #HWU #j #H destruct /3 width=6 by or_introl, ex3_3_intro/ +| #G #L #K #W #V #U #i #HLK #HWV #HWU #j #H destruct /3 width=6 by or_intror, ex3_3_intro/ +| #a #I #G #L #V #T #U #_ #j #H destruct +| #G #L #V #T #U #_ #j #H destruct +| #G #L #W #T #U #_ #j #H destruct +] +qed-. + +(* Basic_1: was sty0_gen_lref *) +lemma sta_inv_lref1: ∀h,G,L,U,i. ⦃G, L⦄ ⊢ #i •[h] U → + (∃∃K,V,W. ⇩[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W & + ⇧[0, i+1] W ≡ U + ) ∨ + (∃∃K,W,V. ⇩[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V & + ⇧[0, i+1] W ≡ U + ). +/2 width=3 by sta_inv_lref1_aux/ qed-. + +fact sta_inv_gref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀p0. T = §p0 → ⊥. +#h #G #L #T #U * -G -L -T -U +[ #G #L #k #p0 #H destruct +| #G #L #K #V #W #U #i #_ #_ #_ #p0 #H destruct +| #G #L #K #W #V #U #i #_ #_ #_ #p0 #H destruct +| #a #I #G #L #V #T #U #_ #p0 #H destruct +| #G #L #V #T #U #_ #p0 #H destruct +| #G #L #W #T #U #_ #p0 #H destruct +qed-. + +lemma sta_inv_gref1: ∀h,G,L,U,p. ⦃G, L⦄ ⊢ §p •[h] U → ⊥. +/2 width=8 by sta_inv_gref1_aux/ qed-. + +fact sta_inv_bind1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀b,J,X,Y. T = ⓑ{b,J}Y.X → + ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h] Z & U = ⓑ{b,J}Y.Z. +#h #G #L #T #U * -G -L -T -U +[ #G #L #k #b #J #X #Y #H destruct +| #G #L #K #V #W #U #i #_ #_ #_ #b #J #X #Y #H destruct +| #G #L #K #W #V #U #i #_ #_ #_ #b #J #X #Y #H destruct +| #a #I #G #L #V #T #U #HTU #b #J #X #Y #H destruct /2 width=3 by ex2_intro/ +| #G #L #V #T #U #_ #b #J #X #Y #H destruct +| #G #L #W #T #U #_ #b #J #X #Y #H destruct +] +qed-. + +(* Basic_1: was: sty0_gen_bind *) +lemma sta_inv_bind1: ∀h,b,J,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X •[h] U → + ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h] Z & U = ⓑ{b,J}Y.Z. +/2 width=3 by sta_inv_bind1_aux/ qed-. + +fact sta_inv_appl1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀X,Y. T = ⓐY.X → + ∃∃Z. ⦃G, L⦄ ⊢ X •[h] Z & U = ⓐY.Z. +#h #G #L #T #U * -G -L -T -U +[ #G #L #k #X #Y #H destruct +| #G #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct +| #G #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct +| #a #I #G #L #V #T #U #_ #X #Y #H destruct +| #G #L #V #T #U #HTU #X #Y #H destruct /2 width=3 by ex2_intro/ +| #G #L #W #T #U #_ #X #Y #H destruct +] +qed-. + +(* Basic_1: was: sty0_gen_appl *) +lemma sta_inv_appl1: ∀h,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓐY.X •[h] U → + ∃∃Z. ⦃G, L⦄ ⊢ X •[h] Z & U = ⓐY.Z. +/2 width=3 by sta_inv_appl1_aux/ qed-. + +fact sta_inv_cast1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀X,Y. T = ⓝY.X → + ⦃G, L⦄ ⊢ X •[h] U. +#h #G #L #T #U * -G -L -T -U +[ #G #L #k #X #Y #H destruct +| #G #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct +| #G #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct +| #a #I #G #L #V #T #U #_ #X #Y #H destruct +| #G #L #V #T #U #_ #X #Y #H destruct +| #G #L #W #T #U #HTU #X #Y #H destruct // +] +qed-. + +(* Basic_1: was: sty0_gen_cast *) +lemma sta_inv_cast1: ∀h,G,L,X,Y,U. ⦃G, L⦄ ⊢ ⓝY.X •[h] U → ⦃G, L⦄ ⊢ X •[h] U. +/2 width=4 by sta_inv_cast1_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta.ma deleted file mode 100644 index fc6bf6a2a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta.ma +++ /dev/null @@ -1,143 +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/notation/relations/statictype_5.ma". -include "basic_2/grammar/genv.ma". -include "basic_2/substitution/drop.ma". -include "basic_2/static/sh.ma". - -(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************) - -(* activate genv *) -inductive sta (h:sh): relation4 genv lenv term term ≝ -| sta_sort: ∀G,L,k. sta h G L (⋆k) (⋆(next h k)) -| sta_ldef: ∀G,L,K,V,W,U,i. ⇩[i] L ≡ K.ⓓV → sta h G K V W → - ⇧[0, i + 1] W ≡ U → sta h G L (#i) U -| sta_ldec: ∀G,L,K,W,V,U,i. ⇩[i] L ≡ K.ⓛW → sta h G K W V → - ⇧[0, i + 1] W ≡ U → sta h G L (#i) U -| sta_bind: ∀a,I,G,L,V,T,U. sta h G (L.ⓑ{I}V) T U → - sta h G L (ⓑ{a,I}V.T) (ⓑ{a,I}V.U) -| sta_appl: ∀G,L,V,T,U. sta h G L T U → sta h G L (ⓐV.T) (ⓐV.U) -| sta_cast: ∀G,L,W,T,U. sta h G L T U → sta h G L (ⓝW.T) U -. - -interpretation "static type assignment (term)" - 'StaticType h G L T U = (sta h G L T U). - -(* Basic inversion lemmas ************************************************) - -fact sta_inv_sort1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀k0. T = ⋆k0 → - U = ⋆(next h k0). -#h #G #L #T #U * -G -L -T -U -[ #G #L #k #k0 #H destruct // -| #G #L #K #V #W #U #i #_ #_ #_ #k0 #H destruct -| #G #L #K #W #V #U #i #_ #_ #_ #k0 #H destruct -| #a #I #G #L #V #T #U #_ #k0 #H destruct -| #G #L #V #T #U #_ #k0 #H destruct -| #G #L #W #T #U #_ #k0 #H destruct -qed-. - -(* Basic_1: was: sty0_gen_sort *) -lemma sta_inv_sort1: ∀h,G,L,U,k. ⦃G, L⦄ ⊢ ⋆k •[h] U → U = ⋆(next h k). -/2 width=5 by sta_inv_sort1_aux/ qed-. - -fact sta_inv_lref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀j. T = #j → - (∃∃K,V,W. ⇩[j] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W & - ⇧[0, j+1] W ≡ U - ) ∨ - (∃∃K,W,V. ⇩[j] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V & - ⇧[0, j+1] W ≡ U - ). -#h #G #L #T #U * -G -L -T -U -[ #G #L #k #j #H destruct -| #G #L #K #V #W #U #i #HLK #HVW #HWU #j #H destruct /3 width=6 by or_introl, ex3_3_intro/ -| #G #L #K #W #V #U #i #HLK #HWV #HWU #j #H destruct /3 width=6 by or_intror, ex3_3_intro/ -| #a #I #G #L #V #T #U #_ #j #H destruct -| #G #L #V #T #U #_ #j #H destruct -| #G #L #W #T #U #_ #j #H destruct -] -qed-. - -(* Basic_1: was sty0_gen_lref *) -lemma sta_inv_lref1: ∀h,G,L,U,i. ⦃G, L⦄ ⊢ #i •[h] U → - (∃∃K,V,W. ⇩[i] L ≡ K.ⓓV & ⦃G, K⦄ ⊢ V •[h] W & - ⇧[0, i+1] W ≡ U - ) ∨ - (∃∃K,W,V. ⇩[i] L ≡ K.ⓛW & ⦃G, K⦄ ⊢ W •[h] V & - ⇧[0, i+1] W ≡ U - ). -/2 width=3 by sta_inv_lref1_aux/ qed-. - -fact sta_inv_gref1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀p0. T = §p0 → ⊥. -#h #G #L #T #U * -G -L -T -U -[ #G #L #k #p0 #H destruct -| #G #L #K #V #W #U #i #_ #_ #_ #p0 #H destruct -| #G #L #K #W #V #U #i #_ #_ #_ #p0 #H destruct -| #a #I #G #L #V #T #U #_ #p0 #H destruct -| #G #L #V #T #U #_ #p0 #H destruct -| #G #L #W #T #U #_ #p0 #H destruct -qed-. - -lemma sta_inv_gref1: ∀h,G,L,U,p. ⦃G, L⦄ ⊢ §p •[h] U → ⊥. -/2 width=8 by sta_inv_gref1_aux/ qed-. - -fact sta_inv_bind1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀b,J,X,Y. T = ⓑ{b,J}Y.X → - ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h] Z & U = ⓑ{b,J}Y.Z. -#h #G #L #T #U * -G -L -T -U -[ #G #L #k #b #J #X #Y #H destruct -| #G #L #K #V #W #U #i #_ #_ #_ #b #J #X #Y #H destruct -| #G #L #K #W #V #U #i #_ #_ #_ #b #J #X #Y #H destruct -| #a #I #G #L #V #T #U #HTU #b #J #X #Y #H destruct /2 width=3 by ex2_intro/ -| #G #L #V #T #U #_ #b #J #X #Y #H destruct -| #G #L #W #T #U #_ #b #J #X #Y #H destruct -] -qed-. - -(* Basic_1: was: sty0_gen_bind *) -lemma sta_inv_bind1: ∀h,b,J,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓑ{b,J}Y.X •[h] U → - ∃∃Z. ⦃G, L.ⓑ{J}Y⦄ ⊢ X •[h] Z & U = ⓑ{b,J}Y.Z. -/2 width=3 by sta_inv_bind1_aux/ qed-. - -fact sta_inv_appl1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀X,Y. T = ⓐY.X → - ∃∃Z. ⦃G, L⦄ ⊢ X •[h] Z & U = ⓐY.Z. -#h #G #L #T #U * -G -L -T -U -[ #G #L #k #X #Y #H destruct -| #G #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct -| #G #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct -| #a #I #G #L #V #T #U #_ #X #Y #H destruct -| #G #L #V #T #U #HTU #X #Y #H destruct /2 width=3 by ex2_intro/ -| #G #L #W #T #U #_ #X #Y #H destruct -] -qed-. - -(* Basic_1: was: sty0_gen_appl *) -lemma sta_inv_appl1: ∀h,G,L,Y,X,U. ⦃G, L⦄ ⊢ ⓐY.X •[h] U → - ∃∃Z. ⦃G, L⦄ ⊢ X •[h] Z & U = ⓐY.Z. -/2 width=3 by sta_inv_appl1_aux/ qed-. - -fact sta_inv_cast1_aux: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∀X,Y. T = ⓝY.X → - ⦃G, L⦄ ⊢ X •[h] U. -#h #G #L #T #U * -G -L -T -U -[ #G #L #k #X #Y #H destruct -| #G #L #K #V #W #U #i #_ #_ #_ #X #Y #H destruct -| #G #L #K #W #V #U #i #_ #_ #_ #X #Y #H destruct -| #a #I #G #L #V #T #U #_ #X #Y #H destruct -| #G #L #V #T #U #_ #X #Y #H destruct -| #G #L #W #T #U #HTU #X #Y #H destruct // -] -qed-. - -(* Basic_1: was: sty0_gen_cast *) -lemma sta_inv_cast1: ∀h,G,L,X,Y,U. ⦃G, L⦄ ⊢ ⓝY.X •[h] U → ⦃G, L⦄ ⊢ X •[h] U. -/2 width=4 by sta_inv_cast1_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_aaa.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_aaa.etc new file mode 100644 index 000000000..affd2eded --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_aaa.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/static/sta.ma". +include "basic_2/static/aaa_lift.ma". + +(* STATIC TYPE ASSIGNMENT FOR TERMS *****************************************) + +(* Properties on atomic arity assignment for terms **************************) + +lemma aaa_sta: ∀h,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃U. ⦃G, L⦄ ⊢ T •[h] U. +#h #G #L #T #A #H elim H -G -L -T -A +[ /2 width=2 by sta_sort, ex_intro/ +| * #G #L #K [ #V | #W ] #B #i #HLK #_ * [ #W | #V ] #HVW + elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldef, sta_ldec, ex_intro/ +| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_bind, ex_intro/ +| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_bind, ex_intro/ +| #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_appl, ex_intro/ +| #G #L #W #T #A #_ #_ #_ * /3 width=2 by sta_cast, ex_intro/ +] +qed-. + +lemma sta_aaa_conf: ∀h,G,L. Conf3 … (aaa G L) (sta h G L). +#h #G #L #T #A #H elim H -G -L -T -A +[ #G #L #k #U #H + lapply (sta_inv_sort1 … H) -H #H destruct // +| #I #G #L #K #V #B #i #HLK #HV #IHV #U #H + elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HU + lapply (drop_mono … HLK0 … HLK) -HLK0 #H0 destruct + lapply (drop_fwd_drop2 … HLK) -HLK #HLK + @(aaa_lift … HLK … HU) -HU -L /2 width=2 by/ +| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H + elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abbr/ +| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H + elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abst/ +| #G #L #V #T #B #A #HV #_ #_ #IHT #X #H + elim (sta_inv_appl1 … H) -H #U #HTU #H destruct /3 width=3 by aaa_appl/ +| #G #L #V #T #A #_ #_ #IHV #IHT #X #H + lapply (sta_inv_cast1 … H) -H /2 width=2 by/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_aaa.ma deleted file mode 100644 index affd2eded..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_aaa.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/static/sta.ma". -include "basic_2/static/aaa_lift.ma". - -(* STATIC TYPE ASSIGNMENT FOR TERMS *****************************************) - -(* Properties on atomic arity assignment for terms **************************) - -lemma aaa_sta: ∀h,G,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∃U. ⦃G, L⦄ ⊢ T •[h] U. -#h #G #L #T #A #H elim H -G -L -T -A -[ /2 width=2 by sta_sort, ex_intro/ -| * #G #L #K [ #V | #W ] #B #i #HLK #_ * [ #W | #V ] #HVW - elim (lift_total W 0 (i+1)) /3 width=7 by sta_ldef, sta_ldec, ex_intro/ -| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_bind, ex_intro/ -| #a #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_bind, ex_intro/ -| #G #L #V #T #B #A #_ #_ #_ * /3 width=2 by sta_appl, ex_intro/ -| #G #L #W #T #A #_ #_ #_ * /3 width=2 by sta_cast, ex_intro/ -] -qed-. - -lemma sta_aaa_conf: ∀h,G,L. Conf3 … (aaa G L) (sta h G L). -#h #G #L #T #A #H elim H -G -L -T -A -[ #G #L #k #U #H - lapply (sta_inv_sort1 … H) -H #H destruct // -| #I #G #L #K #V #B #i #HLK #HV #IHV #U #H - elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HU - lapply (drop_mono … HLK0 … HLK) -HLK0 #H0 destruct - lapply (drop_fwd_drop2 … HLK) -HLK #HLK - @(aaa_lift … HLK … HU) -HU -L /2 width=2 by/ -| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H - elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abbr/ -| #a #G #L #V #T #B #A #HV #_ #_ #IHT #X #H - elim (sta_inv_bind1 … H) -H #U #HTU #H destruct /3 width=2 by aaa_abst/ -| #G #L #V #T #B #A #HV #_ #_ #IHT #X #H - elim (sta_inv_appl1 … H) -H #U #HTU #H destruct /3 width=3 by aaa_appl/ -| #G #L #V #T #A #_ #_ #IHV #IHT #X #H - lapply (sta_inv_cast1 … H) -H /2 width=2 by/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_lift.etc new file mode 100644 index 000000000..4d342878f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_lift.etc @@ -0,0 +1,113 @@ +(**************************************************************************) +(* ___ *) +(* ||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/drop_drop.ma". +include "basic_2/static/sta.ma". + +(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************) + +(* Properties on relocation *************************************************) + +(* Basic_1: was: sty0_lift *) +lemma sta_lift: ∀h,G. l_liftable (sta h G). +#h #G #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1 +[ #G #L1 #k #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2 + >(lift_inv_sort1 … H1) -X1 + >(lift_inv_sort1 … H2) -X2 // +| #G #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HW1 … HWU2) -W // #W2 #HW12 #HWU2 + elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H + elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct + /3 width=9 by sta_ldef/ + | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2 + lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldef, drop_inv_gen/ + ] +| #G #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HW1 … HWU2) -W // (lift_inv_sort2 … H) -X /2 width=3 by sta_sort, lift_sort, ex2_intro/ +| #G #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #d #e #HL21 #X #H + elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ] + [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12 + elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1 + elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus minus_minus_m_m /3 width=8 by sta_ldef, le_S, ex2_intro/ + ] +| #G #L2 #K2 #W2 #V2 #W #i #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #d #e #HL21 #X #H + elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ] + [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12 + elim (IHWV2 … HK21 … HW12) -K2 #V1 #_ #HWV1 + elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus minus_minus_m_m /3 width=8 by sta_ldec, le_S, ex2_intro/ + ] +| #a #I #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H + elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by sta_bind, drop_skip, lift_bind, ex2_intro/ +| #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H + elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct + elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=5 by sta_appl, lift_flat, ex2_intro/ +| #G #L2 #W2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H + elim (lift_inv_flat2 … H) -H #W1 #T1 #_ #HT12 #H destruct + elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3 by sta_cast, ex2_intro/ +] +qed-. + +(* Advanced forward lemmas **************************************************) + +(* Basic_1: was: sty0_correct *) +lemma sta_fwd_correct: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∃T0. ⦃G, L⦄ ⊢ U •[h] T0. +#h #G #L #T #U #H elim H -G -L -T -U +[ /2 width=2/ +| #G #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0 + lapply (drop_fwd_drop2 … HLK) -HLK #HLK + elim (lift_total V0 0 (i+1)) /3 width=11 by ex_intro, sta_lift/ +| #G #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_ + lapply (drop_fwd_drop2 … HLK) -HLK #HLK + elim (lift_total V 0 (i+1)) /3 width=11 by ex_intro, sta_lift/ +| #a #I #G #L #V #T #U #_ * /3 width=2 by sta_bind, ex_intro/ +| #G #L #V #T #U #_ * #T0 #HUT0 /3 width=2 by sta_appl, ex_intro/ +| #G #L #W #T #U #_ * /2 width=2 by ex_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_lift.ma deleted file mode 100644 index 4d342878f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_lift.ma +++ /dev/null @@ -1,113 +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/drop_drop.ma". -include "basic_2/static/sta.ma". - -(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************) - -(* Properties on relocation *************************************************) - -(* Basic_1: was: sty0_lift *) -lemma sta_lift: ∀h,G. l_liftable (sta h G). -#h #G #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1 -[ #G #L1 #k #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2 - >(lift_inv_sort1 … H1) -X1 - >(lift_inv_sort1 … H2) -X2 // -| #G #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2 - elim (lift_inv_lref1 … H) * #Hid #H destruct - [ elim (lift_trans_ge … HW1 … HWU2) -W // #W2 #HW12 #HWU2 - elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H - elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct - /3 width=9 by sta_ldef/ - | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2 - lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldef, drop_inv_gen/ - ] -| #G #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2 - elim (lift_inv_lref1 … H) * #Hid #H destruct - [ elim (lift_trans_ge … HW1 … HWU2) -W // (lift_inv_sort2 … H) -X /2 width=3 by sta_sort, lift_sort, ex2_intro/ -| #G #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #d #e #HL21 #X #H - elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ] - [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12 - elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1 - elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus minus_minus_m_m /3 width=8 by sta_ldef, le_S, ex2_intro/ - ] -| #G #L2 #K2 #W2 #V2 #W #i #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #d #e #HL21 #X #H - elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ] - [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12 - elim (IHWV2 … HK21 … HW12) -K2 #V1 #_ #HWV1 - elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus minus_minus_m_m /3 width=8 by sta_ldec, le_S, ex2_intro/ - ] -| #a #I #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H - elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by sta_bind, drop_skip, lift_bind, ex2_intro/ -| #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H - elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct - elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=5 by sta_appl, lift_flat, ex2_intro/ -| #G #L2 #W2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H - elim (lift_inv_flat2 … H) -H #W1 #T1 #_ #HT12 #H destruct - elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3 by sta_cast, ex2_intro/ -] -qed-. - -(* Advanced forward lemmas **************************************************) - -(* Basic_1: was: sty0_correct *) -lemma sta_fwd_correct: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∃T0. ⦃G, L⦄ ⊢ U •[h] T0. -#h #G #L #T #U #H elim H -G -L -T -U -[ /2 width=2/ -| #G #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0 - lapply (drop_fwd_drop2 … HLK) -HLK #HLK - elim (lift_total V0 0 (i+1)) /3 width=11 by ex_intro, sta_lift/ -| #G #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_ - lapply (drop_fwd_drop2 … HLK) -HLK #HLK - elim (lift_total V 0 (i+1)) /3 width=11 by ex_intro, sta_lift/ -| #a #I #G #L #V #T #U #_ * /3 width=2 by sta_bind, ex_intro/ -| #G #L #V #T #U #_ * #T0 #HUT0 /3 width=2 by sta_appl, ex_intro/ -| #G #L #W #T #U #_ * /2 width=2 by ex_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_llpx_sn.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_llpx_sn.etc new file mode 100644 index 000000000..3adf28fc4 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_llpx_sn.etc @@ -0,0 +1,43 @@ +(**************************************************************************) +(* ___ *) +(* ||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/multiple/llpx_sn_drop.ma". +include "basic_2/static/sta.ma". + +(* STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS ******************************) + +(* Properties on lazy sn pointwise extensions *******************************) + +lemma sta_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → l_liftable R → + ∀h,G. s_r_confluent1 … (sta h G) (llpx_sn R 0). +#R #H1R #H2R #h #G #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2 +[ /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/ +| #G #Ls #Ks #V1s #W2s #V2s #i #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H + #Kd #V1d #HLKd #HV1s #HV1sd + lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs + lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd + @(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *) +| #G #Ls #Ks #V1s #W1s #V2s #i #HLKs #_ #HV12s #IHVW1s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H + #Kd #V1d #HLKd #HV1s #HV1sd + lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs + lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd + @(llpx_sn_lift_le … HLKs HLKd … HV12s) -HLKs -HLKd -HV12s /2 width=1 by/ (**) (* full auto too slow *) +| #a #I #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H + /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/ +| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H + /3 width=1 by llpx_sn_flat/ +| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H + /3 width=1 by llpx_sn_flat/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_llpx_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_llpx_sn.ma deleted file mode 100644 index 3adf28fc4..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_llpx_sn.ma +++ /dev/null @@ -1,43 +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/multiple/llpx_sn_drop.ma". -include "basic_2/static/sta.ma". - -(* STRATIFIED STATIC TYPE ASSIGNMENT FOR TERMS ******************************) - -(* Properties on lazy sn pointwise extensions *******************************) - -lemma sta_llpx_sn_conf: ∀R. (∀L. reflexive … (R L)) → l_liftable R → - ∀h,G. s_r_confluent1 … (sta h G) (llpx_sn R 0). -#R #H1R #H2R #h #G #Ls #T1 #T2 #H elim H -G -Ls -T1 -T2 -[ /3 width=4 by llpx_sn_fwd_length, llpx_sn_sort/ -| #G #Ls #Ks #V1s #W2s #V2s #i #HLKs #_ #HVW2s #IHV12s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H - #Kd #V1d #HLKd #HV1s #HV1sd - lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs - lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd - @(llpx_sn_lift_le … HLKs HLKd … HVW2s) -HLKs -HLKd -HVW2s /2 width=1 by/ (**) (* full auto too slow *) -| #G #Ls #Ks #V1s #W1s #V2s #i #HLKs #_ #HV12s #IHVW1s #Ld #H elim (llpx_sn_inv_lref_ge_sn … H … HLKs) // -H - #Kd #V1d #HLKd #HV1s #HV1sd - lapply (drop_fwd_drop2 … HLKs) -HLKs #HLKs - lapply (drop_fwd_drop2 … HLKd) -HLKd #HLKd - @(llpx_sn_lift_le … HLKs HLKd … HV12s) -HLKs -HLKd -HV12s /2 width=1 by/ (**) (* full auto too slow *) -| #a #I #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_bind_O … H) -H - /4 width=5 by llpx_sn_bind_repl_SO, llpx_sn_bind/ -| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H - /3 width=1 by llpx_sn_flat/ -| #G #Ls #V #T1 #T2 #_ #IHT12 #Ld #H elim (llpx_sn_inv_flat … H) -H - /3 width=1 by llpx_sn_flat/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_sta.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_sta.etc new file mode 100644 index 000000000..6d54d56d7 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_sta.etc @@ -0,0 +1,43 @@ +(**************************************************************************) +(* ___ *) +(* ||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/drop_drop.ma". +include "basic_2/static/sta.ma". + +(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************) + +(* Main properties **********************************************************) + +(* Note: apparently this was missing in basic_1 *) +theorem sta_mono: ∀h,G,L. singlevalued … (sta h G L). +#h #G #L #T #U1 #H elim H -G -L -T -U1 +[ #G #L #k #X #H >(sta_inv_sort1 … H) -X // +| #G #L #K #V #W #U1 #i #HLK #_ #HWU1 #IHVW #U2 #H + elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HW0U2 + lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct + lapply (IHVW … HVW0) -IHVW -HVW0 #H destruct + >(lift_mono … HWU1 … HW0U2) -W0 -U1 // +| #G #L #K #W #V #U1 #i #HLK #_ #HWU1 #IHWV #U2 #H + elim (sta_inv_lref1 … H) -H * #K0 #W0 #V0 #HLK0 #HWV0 #HV0U2 + lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct + lapply (IHWV … HWV0) -IHWV -HWV0 #H destruct + >(lift_mono … HWU1 … HV0U2) -W -U1 // +| #a #I #G #L #V #T #U1 #_ #IHTU1 #X #H + elim (sta_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/ +| #G #L #V #T #U1 #_ #IHTU1 #X #H + elim (sta_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/ +| #G #L #W #T #U1 #_ #IHTU1 #U2 #H + lapply (sta_inv_cast1 … H) -H /2 width=1 by/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_sta.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_sta.ma deleted file mode 100644 index 6d54d56d7..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/sta_sta.ma +++ /dev/null @@ -1,43 +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/drop_drop.ma". -include "basic_2/static/sta.ma". - -(* STATIC TYPE ASSIGNMENT ON TERMS ******************************************) - -(* Main properties **********************************************************) - -(* Note: apparently this was missing in basic_1 *) -theorem sta_mono: ∀h,G,L. singlevalued … (sta h G L). -#h #G #L #T #U1 #H elim H -G -L -T -U1 -[ #G #L #k #X #H >(sta_inv_sort1 … H) -X // -| #G #L #K #V #W #U1 #i #HLK #_ #HWU1 #IHVW #U2 #H - elim (sta_inv_lref1 … H) -H * #K0 #V0 #W0 #HLK0 #HVW0 #HW0U2 - lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct - lapply (IHVW … HVW0) -IHVW -HVW0 #H destruct - >(lift_mono … HWU1 … HW0U2) -W0 -U1 // -| #G #L #K #W #V #U1 #i #HLK #_ #HWU1 #IHWV #U2 #H - elim (sta_inv_lref1 … H) -H * #K0 #W0 #V0 #HLK0 #HWV0 #HV0U2 - lapply (drop_mono … HLK0 … HLK) -HLK -HLK0 #H destruct - lapply (IHWV … HWV0) -IHWV -HWV0 #H destruct - >(lift_mono … HWU1 … HV0U2) -W -U1 // -| #a #I #G #L #V #T #U1 #_ #IHTU1 #X #H - elim (sta_inv_bind1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/ -| #G #L #V #T #U1 #_ #IHTU1 #X #H - elim (sta_inv_appl1 … H) -H #U2 #HTU2 #H destruct /3 width=1 by eq_f/ -| #G #L #W #T #U1 #_ #IHTU1 #U2 #H - lapply (sta_inv_cast1 … H) -H /2 width=1 by/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictype_5.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictype_5.etc new file mode 100644 index 000000000..9249be4e0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictype_5.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||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 FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • break [ term 46 h ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'StaticType $h $G $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictype_5.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictype_5.ma deleted file mode 100644 index 9249be4e0..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictype_5.ma +++ /dev/null @@ -1,19 +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 *) -(* *) -(**************************************************************************) - -(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • break [ term 46 h ] break term 46 T2 )" - non associative with precedence 45 - for @{ 'StaticType $h $G $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictypestaralt_6.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictypestaralt_6.etc new file mode 100644 index 000000000..a6e46be62 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictypestaralt_6.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||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 FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • • * break [ term 46 h , break term 46 l ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'StaticTypeStarAlt $h $G $L $l $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictypestaralt_6.ma b/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictypestaralt_6.ma deleted file mode 100644 index a6e46be62..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/sta/statictypestaralt_6.ma +++ /dev/null @@ -1,19 +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 *) -(* *) -(**************************************************************************) - -(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 • • * break [ term 46 h , break term 46 l ] break term 46 T2 )" - non associative with precedence 45 - for @{ 'StaticTypeStarAlt $h $G $L $l $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma index 6db05fe53..d2034f9d1 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma @@ -15,6 +15,7 @@ include "basic_2/notation/relations/pred_4.ma". include "basic_2/grammar/genv.ma". include "basic_2/static/lsubr.ma". +include "basic_2/unfold/lstas.ma". (* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) @@ -93,6 +94,20 @@ lemma cpr_delift: ∀G,K,V,T1,L,d. ⇩[d] L ≡ (K.ⓓV) → ] qed-. +fact lstas_cpr_aux: ∀h,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*[h, l] T2 → + l = 0 → ⦃G, L⦄ ⊢ T1 ➡ T2. +#h #G #L #T1 #T2 #l #H elim H -G -L -T1 -T2 -l +/3 width=1 by cpr_eps, cpr_flat, cpr_bind/ +[ #G #L #l #k #H0 destruct normalize // +| #G #L #K #V1 #V2 #W2 #i #l #HLK #_ #HVW2 #IHV12 #H destruct + /3 width=6 by cpr_delta/ +| #G #L #K #V1 #V2 #W2 #i #l #_ #_ #_ #_