From: Ferruccio Guidi Date: Thu, 1 Aug 2013 12:32:09 +0000 (+0000) Subject: partial commit: just the components before "static" ... X-Git-Tag: make_still_working~1123 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=29973426e0227ee48368d1c24dc0c17bf2baef77 partial commit: just the components before "static" ... --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/acp_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/acp_aaa.ma index 77b705f97..168d5dfcc 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/acp_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/acp_aaa.ma @@ -81,11 +81,11 @@ qed. (* Basic_1: was: sc3_arity *) lemma aacr_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L,T,A. L ⊢ T ⁝ A → ⦃L, T⦄ ϵ[RP] 〚A〛. + ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃L, T⦄ ϵ[RP] 〚A〛. /2 width=8/ qed. lemma acp_aaa: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L,T,A. L ⊢ T ⁝ A → RP L T. + ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → RP L T. #RR #RS #RP #H1RP #H2RP #L #T #A #HT lapply (aacr_acr … H1RP H2RP A) #HA @(s1 … HA) /2 width=4/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpds.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpds.ma index 981809df8..b725a809a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpds.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpds.ma @@ -19,7 +19,7 @@ include "basic_2/computation/cprs.ma". (* DECOMPOSED EXTENDED PARALLEL COMPUTATION ON TERMS ************************) definition cpds: ∀h. sd h → lenv → relation term ≝ λh,g,L,T1,T2. - ∃∃T. ⦃h, L⦄ ⊢ T1 •*[g] T & L ⊢ T ➡* T2. + ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, g] T & ⦃G, L⦄ ⊢ T ➡* T2. interpretation "decomposed extended parallel computation (term)" 'DPRedStar h g L T1 T2 = (cpds h g L T1 T2). @@ -29,22 +29,22 @@ interpretation "decomposed extended parallel computation (term)" lemma cpds_refl: ∀h,g,L. reflexive … (cpds h g L). /2 width=3/ qed. -lemma sstas_cpds: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 •*[g] T2 → ⦃h, L⦄ ⊢ T1 •*➡*[g] T2. +lemma sstas_cpds: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2. /2 width=3/ qed. -lemma cprs_cpds: ∀h,g,L,T1,T2. L ⊢ T1 ➡* T2 → ⦃h, L⦄ ⊢ T1 •*➡*[g] T2. +lemma cprs_cpds: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2. /2 width=3/ qed. lemma cpds_strap1: ∀h,g,L,T1,T,T2. - ⦃h, L⦄ ⊢ T1 •*➡*[g] T → L ⊢ T ➡ T2 → ⦃h, L⦄ ⊢ T1 •*➡*[g] T2. + ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2. #h #g #L #T1 #T #T2 * /3 width=5/ qed. lemma cpds_strap2: ∀h,g,L,T1,T,T2,l. - ⦃h, L⦄ ⊢ T1 •[g] ⦃l+1, T⦄ → ⦃h, L⦄ ⊢ T •*➡*[g] T2 → ⦃h, L⦄ ⊢ T1 •*➡*[g] T2. + ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, T⦄ → ⦃G, L⦄ ⊢ T •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2. #h #g #L #T1 #T #T2 #l #HT1 * /3 width=4/ qed. -lemma ssta_cprs_cpds: ∀h,g,L,T1,T,T2,l. ⦃h, L⦄ ⊢ T1 •[g] ⦃l+1, T⦄ → - L ⊢ T ➡* T2 → ⦃h, L⦄ ⊢ T1 •*➡*[g] T2. +lemma ssta_cprs_cpds: ∀h,g,L,T1,T,T2,l. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, T⦄ → + ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2. /3 width=3/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_aaa.ma index ed3b092fc..7ab1b9b0b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_aaa.ma @@ -20,6 +20,6 @@ include "basic_2/computation/cpds.ma". (* Properties on atomic arity assignment for terms **************************) -lemma cpds_aaa: ∀h,g,L,T,A. L ⊢ T ⁝ A → ∀U. ⦃h, L⦄ ⊢ T •*➡*[g] U → L ⊢ U ⁝ A. +lemma cpds_aaa: ∀h,g,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀U. ⦃G, L⦄ ⊢ T •*➡*[h, g] U → ⦃G, L⦄ ⊢ U ⁝ A. #h #g #L #T #A #HT #U * /3 width=5 by sstas_aaa, aaa_cprs_conf/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_cpds.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_cpds.ma index ec3d035f0..f5b7792cd 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_cpds.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_cpds.ma @@ -22,29 +22,29 @@ include "basic_2/computation/cpds.ma". (* Advanced properties ******************************************************) lemma cpds_cprs_trans: ∀h,g,L,T1,T,T2. - ⦃h, L⦄ ⊢ T1 •*➡*[g] T → L ⊢ T ➡* T2 → ⦃h, L⦄ ⊢ T1 •*➡*[g] T2. + ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2. #h #g #L #T1 #T #T2 * #T0 #HT10 #HT0 #HT2 lapply (cprs_trans … HT0 … HT2) -T /2 width=3/ qed-. lemma sstas_cpds_trans: ∀h,g,L,T1,T,T2. - ⦃h, L⦄ ⊢ T1 •*[g] T → ⦃h, L⦄ ⊢ T •*➡*[g] T2 → ⦃h, L⦄ ⊢ T1 •*➡*[g] T2. + ⦃G, L⦄ ⊢ T1 •*[h, g] T → ⦃G, L⦄ ⊢ T •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2. #h #g #L #T1 #T #T2 #HT1 * #T0 #HT0 #HT02 lapply (sstas_trans … HT1 … HT0) -T /2 width=3/ qed-. (* Advanced inversion lemmas ************************************************) -lemma cpds_inv_abst1: ∀h,g,a,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓛ{a}V1. T1 •*➡*[g] U2 → - ∃∃V2,T2. L ⊢ V1 ➡* V2 & ⦃h, L.ⓛV1⦄ ⊢ T1 •*➡*[g] T2 & +lemma cpds_inv_abst1: ∀h,g,a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1. T1 •*➡*[h, g] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃h, L.ⓛV1⦄ ⊢ T1 •*➡*[h, g] T2 & U2 = ⓛ{a}V2. T2. #h #g #a #L #V1 #T1 #U2 * #X #H1 #H2 elim (sstas_inv_bind1 … H1) -H1 #U #HTU1 #H destruct elim (cprs_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HUT2 #H destruct /3 width=5/ qed-. -lemma cpds_inv_abbr_abst: ∀h,g,a1,a2,L,V1,W2,T1,T2. ⦃h, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[g] ⓛ{a2}W2.T2 → - ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 •*➡*[g] T & ⇧[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true. +lemma cpds_inv_abbr_abst: ∀h,g,a1,a2,L,V1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, g] ⓛ{a2}W2.T2 → + ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 •*➡*[h, g] T & ⇧[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true. #h #g #a1 #a2 #L #V1 #W2 #T1 #T2 * #X #H1 #H2 elim (sstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct elim (cprs_inv_abbr1 … H2) -H2 * @@ -55,6 +55,6 @@ qed-. (* Advanced forward lemmas **************************************************) -lemma cpds_fwd_cpxs: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 •*➡*[g] T2 → ⦃h, L⦄ ⊢ T1 ➡*[g] T2. +lemma cpds_fwd_cpxs: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2. #h #g #L #T1 #T2 * /3 width=3 by cpxs_trans, sstas_cpxs, cprs_cpxs/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_lift.ma index 1dd32f7a0..40edbb07f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpds_lift.ma @@ -21,15 +21,15 @@ include "basic_2/computation/cpds.ma". (* Relocation properties ****************************************************) lemma cpds_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → ∀T1,U1. ⇧[d, e] T1 ≡ U1 → - ∀h,g,T2. ⦃h, K⦄ ⊢ T1 •*➡*[g] T2 → ∀U2. ⇧[d, e] T2 ≡ U2 → - ⦃h, L⦄ ⊢ U1 •*➡*[g] U2. + ∀h,g,T2. ⦃h, K⦄ ⊢ T1 •*➡*[h, g] T2 → ∀U2. ⇧[d, e] T2 ≡ U2 → + ⦃G, L⦄ ⊢ U1 •*➡*[h, g] U2. #L #K #d #e #HLK #T1 #U1 #HTU1 #h #g #T2 * #T elim (lift_total T d e) /3 width=11/ qed. lemma cpds_inv_lift1: ∀L,K,d,e. ⇩[d, e] L ≡ K → - ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀h,g,U2. ⦃h, L⦄ ⊢ U1 •*➡*[g] U2 → - ∃∃T2. ⇧[d, e] T2 ≡ U2 & ⦃h, K⦄ ⊢ T1 •*➡*[g] T2. + ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀h,g,U2. ⦃G, L⦄ ⊢ U1 •*➡*[h, g] U2 → + ∃∃T2. ⇧[d, e] T2 ≡ U2 & ⦃h, K⦄ ⊢ T1 •*➡*[h, g] T2. #L #K #d #e #HLK #T1 #U1 #HTU1 #h #g #U2 * #U #HU1 #HU2 elim (sstas_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HT1 #HTU elim (cprs_inv_lift1 … HU2 … HLK … HTU) -U -L /3 width=5/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpre.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpre.ma index a31e4023d..d6f6e3dce 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpre.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpre.ma @@ -19,7 +19,7 @@ include "basic_2/computation/csn.ma". (* CONTEXT-SENSITIVE PARALLEL EVALUATION ON TERMS **************************) definition cpre: lenv → relation term ≝ - λL,T1,T2. L ⊢ T1 ➡* T2 ∧ L ⊢ 𝐍⦃T2⦄. + λL,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 ∧ ⦃G, L⦄ ⊢ 𝐍⦃T2⦄. interpretation "context-sensitive parallel evaluation (term)" 'PEval L T1 T2 = (cpre L T1 T2). @@ -27,7 +27,7 @@ interpretation "context-sensitive parallel evaluation (term)" (* Basic_properties *********************************************************) (* Basic_1: was just: nf2_sn3 *) -lemma csn_cpre: ∀h,g,L,T1. ⦃h, L⦄ ⊢ ⬊*[g] T1 → ∃T2. L ⊢ T1 ➡* 𝐍⦃T2⦄. +lemma csn_cpre: ∀h,g,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡* 𝐍⦃T2⦄. #h #g #L #T1 #H @(csn_ind … H) -T1 #T1 #_ #IHT1 elim (cnr_dec L T1) /3 width=3/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpre_cpre.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpre_cpre.ma index cac970e9a..353350192 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpre_cpre.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpre_cpre.ma @@ -20,7 +20,7 @@ include "basic_2/computation/cpre.ma". (* Main properties *********************************************************) (* Basic_1: was: nf2_pr3_confluence *) -theorem cpre_mono: ∀L,T,T1. L ⊢ T ➡* 𝐍⦃T1⦄ → ∀T2. L ⊢ T ➡* 𝐍⦃T2⦄ → T1 = T2. +theorem cpre_mono: ∀L,T,T1. ⦃G, L⦄ ⊢ T ➡* 𝐍⦃T1⦄ → ∀T2. ⦃G, L⦄ ⊢ T ➡* 𝐍⦃T2⦄ → T1 = T2. #L #T #T1 * #H1T1 #H2T1 #T2 * #H1T2 #H2T2 elim (cprs_conf … H1T1 … H1T2) -T #T #HT1 >(cprs_inv_cnr1 … HT1 H2T1) -T1 #HT2 diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs.ma index 1bcc38d80..e4af5fdfb 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs.ma @@ -26,15 +26,15 @@ interpretation "context-sensitive parallel computation (term)" (* Basic eliminators ********************************************************) lemma cprs_ind: ∀L,T1. ∀R:predicate term. R T1 → - (∀T,T2. L ⊢ T1 ➡* T → L ⊢ T ➡ T2 → R T → R T2) → - ∀T2. L ⊢ T1 ➡* T2 → R T2. + (∀T,T2. ⦃G, L⦄ ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ T ➡ T2 → R T → R T2) → + ∀T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → R T2. #L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // qed-. lemma cprs_ind_dx: ∀L,T2. ∀R:predicate term. R T2 → - (∀T1,T. L ⊢ T1 ➡ T → L ⊢ T ➡* T2 → R T → R T1) → - ∀T1. L ⊢ T1 ➡* T2 → R T1. + (∀T1,T. ⦃G, L⦄ ⊢ T1 ➡ T → ⦃G, L⦄ ⊢ T ➡* T2 → R T → R T1) → + ∀T1. ⦃G, L⦄ ⊢ T1 ➡* T2 → R T1. #L #T2 #R #HT2 #IHT2 #T1 #HT12 @(TC_star_ind_dx … HT2 IHT2 … HT12) // qed-. @@ -42,20 +42,20 @@ qed-. (* Basic properties *********************************************************) (* Basic_1: was: pr3_pr2 *) -lemma cpr_cprs: ∀L,T1,T2. L ⊢ T1 ➡ T2 → L ⊢ T1 ➡* T2. +lemma cpr_cprs: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡* T2. /2 width=1/ qed. (* Basic_1: was: pr3_refl *) -lemma cprs_refl: ∀L,T. L ⊢ T ➡* T. +lemma cprs_refl: ∀L,T. ⦃G, L⦄ ⊢ T ➡* T. /2 width=1/ qed. lemma cprs_strap1: ∀L,T1,T,T2. - L ⊢ T1 ➡* T → L ⊢ T ➡ T2 → L ⊢ T1 ➡* T2. + ⦃G, L⦄ ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡* T2. normalize /2 width=3/ qed. (* Basic_1: was: pr3_step *) lemma cprs_strap2: ∀L,T1,T,T2. - L ⊢ T1 ➡ T → L ⊢ T ➡* T2 → L ⊢ T1 ➡* T2. + ⦃G, L⦄ ⊢ T1 ➡ T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡* T2. normalize /2 width=3/ qed. lemma lsubr_cprs_trans: lsub_trans … cprs lsubr. @@ -63,50 +63,50 @@ lemma lsubr_cprs_trans: lsub_trans … cprs lsubr. qed-. (* Basic_1: was: pr3_pr1 *) -lemma tprs_cprs: ∀L,T1,T2. ⋆ ⊢ T1 ➡* T2 → L ⊢ T1 ➡* T2. +lemma tprs_cprs: ∀L,T1,T2. ⋆ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡* T2. #L #T1 #T2 #H @(lsubr_cprs_trans … H) -H // qed. -lemma cprs_bind_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀I,T1,T2. L. ⓑ{I}V1 ⊢ T1 ➡* T2 → - ∀a. L ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. +lemma cprs_bind_dx: ∀L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀I,T1,T2. L. ⓑ{I}V1 ⊢ T1 ➡* T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 #a @(cprs_ind_dx … HT12) -T1 /3 width=1/ /3 width=3/ qed. (* Basic_1: was only: pr3_thin_dx *) -lemma cprs_flat_dx: ∀I,L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L ⊢ T1 ➡* T2 → - L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. +lemma cprs_flat_dx: ∀I,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → + ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. #I #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind … HT12) -T2 /3 width=1/ #T #T2 #_ #HT2 #IHT1 @(cprs_strap1 … IHT1) -V1 -T1 /2 width=1/ qed. -lemma cprs_flat_sn: ∀I,L,T1,T2. L ⊢ T1 ➡ T2 → ∀V1,V2. L ⊢ V1 ➡* V2 → - L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. +lemma cprs_flat_sn: ∀I,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 → + ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. #I #L #T1 #T2 #HT12 #V1 #V2 #H @(cprs_ind … H) -V2 /3 width=1/ #V #V2 #_ #HV2 #IHV1 @(cprs_strap1 … IHV1) -V1 -T1 /2 width=1/ qed. lemma cprs_zeta: ∀L,V,T1,T,T2. ⇧[0, 1] T2 ≡ T → - L.ⓓV ⊢ T1 ➡* T → L ⊢ +ⓓV.T1 ➡* T2. + L.ⓓV ⊢ T1 ➡* T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡* T2. #L #V #T1 #T #T2 #HT2 #H @(TC_ind_dx … T1 H) -T1 /3 width=3/ qed. -lemma cprs_tau: ∀L,T1,T2. L ⊢ T1 ➡* T2 → ∀V. L ⊢ ⓝV.T1 ➡* T2. +lemma cprs_tau: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ∀V. ⦃G, L⦄ ⊢ ⓝV.T1 ➡* T2. #L #T1 #T2 #H elim H -T2 /2 width=3/ /3 width=1/ qed. lemma cprs_beta_dx: ∀a,L,V1,V2,W1,W2,T1,T2. - L ⊢ V1 ➡ V2 → L ⊢ W1 ➡ W2 → L.ⓛW1 ⊢ T1 ➡* T2 → - L ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2. + ⦃G, L⦄ ⊢ V1 ➡ V2 → ⦃G, L⦄ ⊢ W1 ➡ W2 → L.ⓛW1 ⊢ T1 ➡* T2 → + ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2. #a #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HW12 * -T2 /3 width=1/ /4 width=7 by cprs_strap1, cprs_bind_dx, cprs_flat_dx, cpr_beta/ (**) (* auto too slow without trace *) qed. lemma cprs_theta_dx: ∀a,L,V1,V,V2,W1,W2,T1,T2. - L ⊢ V1 ➡ V → ⇧[0, 1] V ≡ V2 → L.ⓓW1 ⊢ T1 ➡* T2 → - L ⊢ W1 ➡ W2 → L ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2. + ⦃G, L⦄ ⊢ V1 ➡ V → ⇧[0, 1] V ≡ V2 → L.ⓓW1 ⊢ T1 ➡* T2 → + ⦃G, L⦄ ⊢ W1 ➡ W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2. #a #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 * -T2 [ /3 width=3/ ] /4 width=9 by cprs_strap1, cprs_bind_dx, cprs_flat_dx, cpr_theta/ (**) (* auto too slow without trace *) qed. @@ -114,15 +114,15 @@ qed. (* Basic inversion lemmas ***************************************************) (* Basic_1: was: pr3_gen_sort *) -lemma cprs_inv_sort1: ∀L,U2,k. L ⊢ ⋆k ➡* U2 → U2 = ⋆k. +lemma cprs_inv_sort1: ∀L,U2,k. ⦃G, L⦄ ⊢ ⋆k ➡* U2 → U2 = ⋆k. #L #U2 #k #H @(cprs_ind … H) -U2 // #U2 #U #_ #HU2 #IHU2 destruct >(cpr_inv_sort1 … HU2) -HU2 // qed-. (* Basic_1: was: pr3_gen_cast *) -lemma cprs_inv_cast1: ∀L,W1,T1,U2. L ⊢ ⓝW1.T1 ➡* U2 → L ⊢ T1 ➡* U2 ∨ - ∃∃W2,T2. L ⊢ W1 ➡* W2 & L ⊢ T1 ➡* T2 & U2 = ⓝW2.T2. +lemma cprs_inv_cast1: ∀L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓝW1.T1 ➡* U2 → ⦃G, L⦄ ⊢ T1 ➡* U2 ∨ + ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L⦄ ⊢ T1 ➡* T2 & U2 = ⓝW2.T2. #L #W1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/ #U2 #U #_ #HU2 * /3 width=3/ * #W #T #HW1 #HT1 #H destruct @@ -131,7 +131,7 @@ elim (cpr_inv_cast1 … HU2) -HU2 /3 width=3/ * qed-. (* Basic_1: was: nf2_pr3_unfold *) -lemma cprs_inv_cnr1: ∀L,T,U. L ⊢ T ➡* U → L ⊢ 𝐍⦃T⦄ → T = U. +lemma cprs_inv_cnr1: ∀L,T,U. ⦃G, L⦄ ⊢ T ➡* U → ⦃G, L⦄ ⊢ 𝐍⦃T⦄ → T = U. #L #T #U #H @(cprs_ind_dx … H) -T // #T0 #T #H1T0 #_ #IHT #H2T0 lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_cprs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_cprs.ma index 7e8723fb1..1941a33bb 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_cprs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_cprs.ma @@ -29,48 +29,48 @@ theorem cprs_trans: ∀L. Transitive … (cprs L). theorem cprs_conf: ∀L. confluent2 … (cprs L) (cprs L). #L @TC_confluent2 /2 width=3 by cpr_conf/ qed-. (**) (* auto /3 width=3/ does not work because a δ-expansion gets in the way *) -theorem cprs_bind: ∀a,I,L,V1,V2,T1,T2. L. ⓑ{I}V1 ⊢ T1 ➡* T2 → L ⊢ V1 ➡* V2 → - L ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. +theorem cprs_bind: ∀a,I,L,V1,V2,T1,T2. L. ⓑ{I}V1 ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ V1 ➡* V2 → + ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. #a #I #L #V1 #V2 #T1 #T2 #HT12 #H @(cprs_ind … H) -V2 /2 width=1/ #V #V2 #_ #HV2 #IHV1 @(cprs_trans … IHV1) -V1 /2 width=1/ qed. (* Basic_1: was: pr3_flat *) -theorem cprs_flat: ∀I,L,V1,V2,T1,T2. L ⊢ T1 ➡* T2 → L ⊢ V1 ➡* V2 → - L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. +theorem cprs_flat: ∀I,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ V1 ➡* V2 → + ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2. #I #L #V1 #V2 #T1 #T2 #HT12 #H @(cprs_ind … H) -V2 /2 width=1/ #V #V2 #_ #HV2 #IHV1 @(cprs_trans … IHV1) -IHV1 /2 width=1/ qed. theorem cprs_beta_rc: ∀a,L,V1,V2,W1,W2,T1,T2. - L ⊢ V1 ➡ V2 → L.ⓛW1 ⊢ T1 ➡* T2 → L ⊢ W1 ➡* W2 → - L ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2. + ⦃G, L⦄ ⊢ V1 ➡ V2 → L.ⓛW1 ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → + ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2. #a #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HT12 #H @(cprs_ind … H) -W2 /2 width=1/ #W #W2 #_ #HW2 #IHW1 @(cprs_trans … IHW1) -IHW1 /3 width=1/ qed. theorem cprs_beta: ∀a,L,V1,V2,W1,W2,T1,T2. - L.ⓛW1 ⊢ T1 ➡* T2 → L ⊢ W1 ➡* W2 → L ⊢ V1 ➡* V2 → - L ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2. + L.ⓛW1 ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L⦄ ⊢ V1 ➡* V2 → + ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡* ⓓ{a}ⓝW2.V2.T2. #a #L #V1 #V2 #W1 #W2 #T1 #T2 #HT12 #HW12 #H @(cprs_ind … H) -V2 /2 width=1/ #V #V2 #_ #HV2 #IHV1 @(cprs_trans … IHV1) -IHV1 /3 width=1/ qed. theorem cprs_theta_rc: ∀a,L,V1,V,V2,W1,W2,T1,T2. - L ⊢ V1 ➡ V → ⇧[0, 1] V ≡ V2 → L.ⓓW1 ⊢ T1 ➡* T2 → - L ⊢ W1 ➡* W2 → L ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2. + ⦃G, L⦄ ⊢ V1 ➡ V → ⇧[0, 1] V ≡ V2 → L.ⓓW1 ⊢ T1 ➡* T2 → + ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2. #a #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H elim H -W2 /2 width=3/ #W #W2 #_ #HW2 #IHW1 @(cprs_trans … IHW1) /2 width=1/ qed. theorem cprs_theta: ∀a,L,V1,V,V2,W1,W2,T1,T2. - ⇧[0, 1] V ≡ V2 → L ⊢ W1 ➡* W2 → L.ⓓW1 ⊢ T1 ➡* T2 → - L ⊢ V1 ➡* V → L ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2. + ⇧[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → L.ⓓW1 ⊢ T1 ➡* T2 → + ⦃G, L⦄ ⊢ V1 ➡* V → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2. #a #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(TC_ind_dx … V1 H) -V1 /2 width=3/ #V1 #V0 #HV10 #_ #IHV0 @(cprs_trans … IHV0) /2 width=1/ @@ -79,14 +79,14 @@ qed. (* Advanced inversion lemmas ************************************************) (* Basic_1: was pr3_gen_appl *) -lemma cprs_inv_appl1: ∀L,V1,T1,U2. L ⊢ ⓐV1.T1 ➡* U2 → - ∨∨ ∃∃V2,T2. L ⊢ V1 ➡* V2 & L ⊢ T1 ➡* T2 & +lemma cprs_inv_appl1: ∀L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡* U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L⦄ ⊢ T1 ➡* T2 & U2 = ⓐV2. T2 - | ∃∃a,W,T. L ⊢ T1 ➡* ⓛ{a}W.T & - L ⊢ ⓓ{a}ⓝW.V1.T ➡* U2 - | ∃∃a,V0,V2,V,T. L ⊢ V1 ➡* V0 & ⇧[0,1] V0 ≡ V2 & - L ⊢ T1 ➡* ⓓ{a}V.T & - L ⊢ ⓓ{a}V.ⓐV2.T ➡* U2. + | ∃∃a,W,T. ⦃G, L⦄ ⊢ T1 ➡* ⓛ{a}W.T & + ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡* U2 + | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡* V0 & ⇧[0,1] V0 ≡ V2 & + ⦃G, L⦄ ⊢ T1 ➡* ⓓ{a}V.T & + ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡* U2. #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 [ /3 width=5/ ] #U #U2 #_ #HU2 * * [ #V0 #T0 #HV10 #HT10 #H destruct @@ -128,8 +128,8 @@ lemma lpr_cpr_trans: s_r_trans … cpr lpr. ] qed-. -lemma cpr_bind2: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀I,T1,T2. L. ⓑ{I}V2 ⊢ T1 ➡ T2 → - ∀a. L ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. +lemma cpr_bind2: ∀L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀I,T1,T2. L. ⓑ{I}V2 ⊢ T1 ➡ T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 lapply (lpr_cpr_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1/ qed. @@ -166,8 +166,8 @@ elim (cprs_lpr_conf_dx … HT01 … HL01) -HT01 #T #HT1 lapply (lpr_cprs_trans … HT1 … HL01) -HT1 /2 width=3/ qed-. -lemma cprs_bind2_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀I,T1,T2. L. ⓑ{I}V2 ⊢ T1 ➡* T2 → - ∀a. L ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. +lemma cprs_bind2_dx: ∀L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀I,T1,T2. L. ⓑ{I}V2 ⊢ T1 ➡* T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 lapply (lpr_cprs_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma index 7e75767b1..7e5961604 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma @@ -22,7 +22,7 @@ include "basic_2/computation/cprs.ma". (* Note: apparently this was missing in basic_1 *) lemma cprs_delta: ∀L,K,V,V2,i. ⇩[0, i] L ≡ K. ⓓV → K ⊢ V ➡* V2 → - ∀W2. ⇧[0, i + 1] V2 ≡ W2 → L ⊢ #i ➡* W2. + ∀W2. ⇧[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡* W2. #L #K #V #V2 #i #HLK #H elim H -V2 [ /3 width=6/ ] #V1 #V2 #_ #HV12 #IHV1 #W2 #HVW2 lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK @@ -32,7 +32,7 @@ qed. (* Advanced inversion lemmas ************************************************) (* Basic_1: was: pr3_gen_lref *) -lemma cprs_inv_lref1: ∀L,T2,i. L ⊢ #i ➡* T2 → +lemma cprs_inv_lref1: ∀L,T2,i. ⦃G, L⦄ ⊢ #i ➡* T2 → T2 = #i ∨ ∃∃K,V1,T1. ⇩[0, i] L ≡ K. ⓓV1 & K ⊢ V1 ➡* T1 & diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxe.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxe.ma index 0a150d198..fd875138c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxe.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxe.ma @@ -19,14 +19,14 @@ include "basic_2/computation/csn.ma". (* CONTEXT-SENSITIVE EXTENDED PARALLEL EVALUATION ON TERMS ******************) definition cpxe: ∀h. sd h → lenv → relation term ≝ - λh,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 ∧ ⦃h, L⦄ ⊢ 𝐍[g]⦃T2⦄. + λh,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 ∧ ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T2⦄. interpretation "context-sensitive extended parallel evaluation (term)" 'PEval h g L T1 T2 = (cpxe h g L T1 T2). (* Basic_properties *********************************************************) -lemma csn_cpxe: ∀h,g,L,T1. ⦃h, L⦄ ⊢ ⬊*[g] T1 → ∃T2. ⦃h, L⦄ ⊢ T1 ➡*[g] 𝐍⦃T2⦄. +lemma csn_cpxe: ∀h,g,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → ∃T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] 𝐍⦃T2⦄. #h #g #L #T1 #H @(csn_ind … H) -T1 #T1 #_ #IHT1 elim (cnx_dec h g L T1) /3 width=3/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs.ma index 635aff910..ba5fd7c5a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs.ma @@ -28,97 +28,97 @@ interpretation "extended context-sensitive parallel computation (term)" (* Basic eliminators ********************************************************) lemma cpxs_ind: ∀h,g,L,T1. ∀R:predicate term. R T1 → - (∀T,T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T → ⦃h, L⦄ ⊢ T ➡[g] T2 → R T → R T2) → - ∀T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → R T2. + (∀T,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T → ⦃G, L⦄ ⊢ T ➡[h, g] T2 → R T → R T2) → + ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → R T2. #h #g #L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // qed-. lemma cpxs_ind_dx: ∀h,g,L,T2. ∀R:predicate term. R T2 → - (∀T1,T. ⦃h, L⦄ ⊢ T1 ➡[g] T → ⦃h, L⦄ ⊢ T ➡*[g] T2 → R T → R T1) → - ∀T1. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → R T1. + (∀T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T → ⦃G, L⦄ ⊢ T ➡*[h, g] T2 → R T → R T1) → + ∀T1. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → R T1. #h #g #L #T2 #R #HT2 #IHT2 #T1 #HT12 @(TC_star_ind_dx … HT2 IHT2 … HT12) // qed-. (* Basic properties *********************************************************) -lemma cpxs_refl: ∀h,g,L,T. ⦃h, L⦄ ⊢ T ➡*[g] T. +lemma cpxs_refl: ∀h,g,L,T. ⦃G, L⦄ ⊢ T ➡*[h, g] T. /2 width=1/ qed. -lemma cpx_cpxs: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → ⦃h, L⦄ ⊢ T1 ➡*[g] T2. +lemma cpx_cpxs: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2. /2 width=1/ qed. -lemma cpxs_strap1: ∀h,g,L,T1,T. ⦃h, L⦄ ⊢ T1 ➡*[g] T → - ∀T2. ⦃h, L⦄ ⊢ T ➡[g] T2 → ⦃h, L⦄ ⊢ T1 ➡*[g] T2. +lemma cpxs_strap1: ∀h,g,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T → + ∀T2. ⦃G, L⦄ ⊢ T ➡[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2. normalize /2 width=3/ qed. -lemma cpxs_strap2: ∀h,g,L,T1,T. ⦃h, L⦄ ⊢ T1 ➡[g] T → - ∀T2. ⦃h, L⦄ ⊢ T ➡*[g] T2 → ⦃h, L⦄ ⊢ T1 ➡*[g] T2. +lemma cpxs_strap2: ∀h,g,L,T1,T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T → + ∀T2. ⦃G, L⦄ ⊢ T ➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2. normalize /2 width=3/ qed. lemma lsubr_cpxs_trans: ∀h,g. lsub_trans … (cpxs h g) lsubr. /3 width=5 by lsubr_cpx_trans, TC_lsub_trans/ qed-. -lemma sstas_cpxs: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 •* [g] T2 → ⦃h, L⦄ ⊢ T1 ➡*[g] T2. +lemma sstas_cpxs: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 •* [h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2. #h #g #L #T1 #T2 #H @(sstas_ind … H) -T2 // /3 width=4 by cpxs_strap1, ssta_cpx/ qed. -lemma cprs_cpxs: ∀h,g,L,T1,T2. L ⊢ T1 ➡* T2 → ⦃h, L⦄ ⊢ T1 ➡*[g] T2. +lemma cprs_cpxs: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2. #h #g #L #T1 #T2 #H @(cprs_ind … H) -T2 // /3 width=3/ qed. -lemma cpxs_bind_dx: ∀h,g,L,V1,V2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 → - ∀I,T1,T2. ⦃h, L. ⓑ{I}V1⦄ ⊢ T1 ➡*[g] T2 → - ∀a. ⦃h, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[g] ⓑ{a,I}V2.T2. +lemma cpxs_bind_dx: ∀h,g,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → + ∀I,T1,T2. ⦃h, L. ⓑ{I}V1⦄ ⊢ T1 ➡*[h, g] T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2. #h #g #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 #a @(cpxs_ind_dx … HT12) -T1 /3 width=1/ /3 width=3/ qed. -lemma cpxs_flat_dx: ∀h,g,L,V1,V2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 → - ∀T1,T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → - ∀I. ⦃h, L⦄ ⊢ ⓕ{I} V1. T1 ➡*[g] ⓕ{I} V2. T2. +lemma cpxs_flat_dx: ∀h,g,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → + ∀T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → + ∀I. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡*[h, g] ⓕ{I} V2. T2. #h #g #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cpxs_ind … HT12) -T2 /3 width=1/ /3 width=5/ qed. -lemma cpxs_flat_sn: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → - ∀V1,V2. ⦃h, L⦄ ⊢ V1 ➡*[g] V2 → - ∀I. ⦃h, L⦄ ⊢ ⓕ{I} V1. T1 ➡*[g] ⓕ{I} V2. T2. +lemma cpxs_flat_sn: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → + ∀V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 → + ∀I. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ➡*[h, g] ⓕ{I} V2. T2. #h #g #L #T1 #T2 #HT12 #V1 #V2 #H @(cpxs_ind … H) -V2 /3 width=1/ /3 width=5/ qed. lemma cpxs_zeta: ∀h,g,L,V,T1,T,T2. ⇧[0, 1] T2 ≡ T → - ⦃h, L.ⓓV⦄ ⊢ T1 ➡*[g] T → ⦃h, L⦄ ⊢ +ⓓV.T1 ➡*[g] T2. + ⦃h, L.ⓓV⦄ ⊢ T1 ➡*[h, g] T → ⦃G, L⦄ ⊢ +ⓓV.T1 ➡*[h, g] T2. #h #g #L #V #T1 #T #T2 #HT2 #H @(TC_ind_dx … T1 H) -T1 /3 width=3/ qed. -lemma cpxs_tau: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → ∀V. ⦃h, L⦄ ⊢ ⓝV.T1 ➡*[g] T2. +lemma cpxs_tau: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ∀V. ⦃G, L⦄ ⊢ ⓝV.T1 ➡*[h, g] T2. #h #g #L #T1 #T2 #H elim H -T2 /2 width=3/ /3 width=1/ qed. -lemma cpxs_ti: ∀h,g,L,V1,V2. ⦃h, L⦄ ⊢ V1 ➡*[g] V2 → ∀T. ⦃h, L⦄ ⊢ ⓝV1.T ➡*[g] V2. +lemma cpxs_ti: ∀h,g,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 → ∀T. ⦃G, L⦄ ⊢ ⓝV1.T ➡*[h, g] V2. #h #g #L #V1 #V2 #H elim H -V2 /2 width=3/ /3 width=1/ qed. lemma cpxs_beta_dx: ∀h,g,a,L,V1,V2,W1,W2,T1,T2. - ⦃h, L⦄ ⊢ V1 ➡[g] V2 → ⦃h, L.ⓛW1⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ W1 ➡[g] W2 → - ⦃h, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[g] ⓓ{a}ⓝW2.V2.T2. + ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → ⦃h, L.ⓛW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 → + ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, g] ⓓ{a}ⓝW2.V2.T2. #h #g #a #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 * -T2 /3 width=1/ /4 width=7 by cpxs_strap1, cpxs_bind_dx, cpxs_flat_dx, cpx_beta/ (**) (* auto too slow without trace *) qed. lemma cpxs_theta_dx: ∀h,g,a,L,V1,V,V2,W1,W2,T1,T2. - ⦃h, L⦄ ⊢ V1 ➡[g] V → ⇧[0, 1] V ≡ V2 → ⦃h, L.ⓓW1⦄ ⊢ T1 ➡*[g] T2 → - ⦃h, L⦄ ⊢ W1 ➡[g] W2 → ⦃h, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[g] ⓓ{a}W2.ⓐV2.T2. + ⦃G, L⦄ ⊢ V1 ➡[h, g] V → ⇧[0, 1] V ≡ V2 → ⦃h, L.ⓓW1⦄ ⊢ T1 ➡*[h, g] T2 → + ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, g] ⓓ{a}W2.ⓐV2.T2. #h #g #a #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 * -T2 [ /3 width=3/ ] /4 width=9 by cpxs_strap1, cpxs_bind_dx, cpxs_flat_dx, cpx_theta/ (**) (* auto too slow without trace *) qed. (* Basic inversion lemmas ***************************************************) -lemma cpxs_inv_sort1: ∀h,g,L,U2,k. ⦃h, L⦄ ⊢ ⋆k ➡*[g] U2 → +lemma cpxs_inv_sort1: ∀h,g,L,U2,k. ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] U2 → ∃∃n,l. deg h g k (n+l) & U2 = ⋆((next h)^n k). #h #g #L #U2 #k #H @(cpxs_ind … H) -U2 [ elim (deg_total h g k) #l #Hkl @@ -132,10 +132,10 @@ lemma cpxs_inv_sort1: ∀h,g,L,U2,k. ⦃h, L⦄ ⊢ ⋆k ➡*[g] U2 → ] qed-. -lemma cpxs_inv_cast1: ∀h,g,L,W1,T1,U2. ⦃h, L⦄ ⊢ ⓝW1.T1 ➡*[g] U2 → - ∨∨ ∃∃W2,T2. ⦃h, L⦄ ⊢ W1 ➡*[g] W2 & ⦃h, L⦄ ⊢ T1 ➡*[g] T2 & U2 = ⓝW2.T2 - | ⦃h, L⦄ ⊢ T1 ➡*[g] U2 - | ⦃h, L⦄ ⊢ W1 ➡*[g] U2. +lemma cpxs_inv_cast1: ∀h,g,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓝW1.T1 ➡*[h, g] U2 → + ∨∨ ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 & ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 & U2 = ⓝW2.T2 + | ⦃G, L⦄ ⊢ T1 ➡*[h, g] U2 + | ⦃G, L⦄ ⊢ W1 ➡*[h, g] U2. #h #g #L #W1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5/ #U2 #U #_ #HU2 * /3 width=3/ * #W #T #HW1 #HT1 #H destruct @@ -145,7 +145,7 @@ lapply (cpxs_strap1 … HW1 … HW2) -W lapply (cpxs_strap1 … HT1 … HT2) -T /3 width=5/ qed-. -lemma cpxs_inv_cnx1: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T ➡*[g] U → ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → T = U. +lemma cpxs_inv_cnx1: ∀h,g,L,T,U. ⦃G, L⦄ ⊢ T ➡*[h, g] U → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → T = U. #h #g #L #T #U #H @(cpxs_ind_dx … H) -T // #T0 #T #H1T0 #_ #IHT #H2T0 lapply (H2T0 … H1T0) -H1T0 #H destruct /2 width=1/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_aaa.ma index ed2329c54..5f377731a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_aaa.ma @@ -19,11 +19,11 @@ include "basic_2/computation/cpxs.ma". (* Properties about atomic arity assignment on terms ************************) -lemma aaa_cpxs_conf: ∀h,g,L,T1,A. L ⊢ T1 ⁝ A → - ∀T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → L ⊢ T2 ⁝ A. +lemma aaa_cpxs_conf: ∀h,g,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → + ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A. #h #g #L #T1 #A #HT1 #T2 #HT12 @(TC_Conf3 … HT1 ? HT12) -A -T1 -T2 /2 width=5 by aaa_cpx_conf/ qed-. -lemma aaa_cprs_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → ∀T2. L ⊢ T1 ➡* T2 → L ⊢ T2 ⁝ A. +lemma aaa_cprs_conf: ∀L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T2 ⁝ A. /3 width=5 by aaa_cpxs_conf, cprs_cpxs/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpxs.ma index 2399ca665..895832329 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpxs.ma @@ -22,51 +22,51 @@ include "basic_2/computation/cpxs_lift.ma". theorem cpxs_trans: ∀h,g,L. Transitive … (cpxs h g L). #h #g #L #T1 #T #HT1 #T2 @trans_TC @HT1 qed-. (**) (* auto /3 width=3/ does not work because a δ-expansion gets in the way *) -theorem cpxs_bind: ∀h,g,a,I,L,V1,V2,T1,T2. ⦃h, L.ⓑ{I}V1⦄ ⊢ T1 ➡*[g] T2 → - ⦃h, L⦄ ⊢ V1 ➡*[g] V2 → - ⦃h, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[g] ⓑ{a,I}V2.T2. +theorem cpxs_bind: ∀h,g,a,I,L,V1,V2,T1,T2. ⦃h, L.ⓑ{I}V1⦄ ⊢ T1 ➡*[h, g] T2 → + ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 → + ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2. #h #g #a #I #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2 /2 width=1/ #V #V2 #_ #HV2 #IHV1 @(cpxs_trans … IHV1) -V1 /2 width=1/ qed. -theorem cpxs_flat: ∀h,g,I,L,V1,V2,T1,T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → - ⦃h, L⦄ ⊢ V1 ➡*[g] V2 → - ⦃h, L⦄ ⊢ ⓕ{I} V1.T1 ➡*[g] ⓕ{I} V2.T2. +theorem cpxs_flat: ∀h,g,I,L,V1,V2,T1,T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → + ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 → + ⦃G, L⦄ ⊢ ⓕ{I} V1.T1 ➡*[h, g] ⓕ{I} V2.T2. #h #g #I #L #V1 #V2 #T1 #T2 #HT12 #H @(cpxs_ind … H) -V2 /2 width=1/ #V #V2 #_ #HV2 #IHV1 @(cpxs_trans … IHV1) -IHV1 /2 width=1/ qed. theorem cpxs_beta_rc: ∀h,g,a,L,V1,V2,W1,W2,T1,T2. - ⦃h, L⦄ ⊢ V1 ➡[g] V2 → ⦃h, L.ⓛW1⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ W1 ➡*[g] W2 → - ⦃h, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[g] ⓓ{a}ⓝW2.V2.T2. + ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → ⦃h, L.ⓛW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 → + ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, g] ⓓ{a}ⓝW2.V2.T2. #h #g #a #L #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HT12 #H @(cpxs_ind … H) -W2 /2 width=1/ #W #W2 #_ #HW2 #IHW1 @(cpxs_trans … IHW1) -IHW1 /3 width=1/ qed. theorem cpxs_beta: ∀h,g,a,L,V1,V2,W1,W2,T1,T2. - ⦃h, L.ⓛW1⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ W1 ➡*[g] W2 → ⦃h, L⦄ ⊢ V1 ➡*[g] V2 → - ⦃h, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[g] ⓓ{a}ⓝW2.V2.T2. + ⦃h, L.ⓛW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 → ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 → + ⦃G, L⦄ ⊢ ⓐV1.ⓛ{a}W1.T1 ➡*[h, g] ⓓ{a}ⓝW2.V2.T2. #h #g #a #L #V1 #V2 #W1 #W2 #T1 #T2 #HT12 #HW12 #H @(cpxs_ind … H) -V2 /2 width=1/ #V #V2 #_ #HV2 #IHV1 @(cpxs_trans … IHV1) -IHV1 /3 width=1/ qed. theorem cpxs_theta_rc: ∀h,g,a,L,V1,V,V2,W1,W2,T1,T2. - ⦃h, L⦄ ⊢ V1 ➡[g] V → ⇧[0, 1] V ≡ V2 → - ⦃h, L.ⓓW1⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ W1 ➡*[g] W2 → - ⦃h, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[g] ⓓ{a}W2.ⓐV2.T2. + ⦃G, L⦄ ⊢ V1 ➡[h, g] V → ⇧[0, 1] V ≡ V2 → + ⦃h, L.ⓓW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 → + ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, g] ⓓ{a}W2.ⓐV2.T2. #h #g #a #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H elim H -W2 /2 width=3/ #W #W2 #_ #HW2 #IHW1 @(cpxs_trans … IHW1) -IHW1 /2 width=1/ qed. theorem cpxs_theta: ∀h,g,a,L,V1,V,V2,W1,W2,T1,T2. - ⇧[0, 1] V ≡ V2 → ⦃h, L⦄ ⊢ W1 ➡*[g] W2 → - ⦃h, L.ⓓW1⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ V1 ➡*[g] V → - ⦃h, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[g] ⓓ{a}W2.ⓐV2.T2. + ⇧[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡*[h, g] W2 → + ⦃h, L.ⓓW1⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ V1 ➡*[h, g] V → + ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡*[h, g] ⓓ{a}W2.ⓐV2.T2. #h #g #a #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(TC_ind_dx … V1 H) -V1 /2 width=3/ #V1 #V0 #HV10 #_ #IHV0 @(cpxs_trans … IHV0) -IHV0 /2 width=1/ @@ -74,12 +74,12 @@ qed. (* Advanced inversion lemmas ************************************************) -lemma cpxs_inv_appl1: ∀h,g,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓐV1.T1 ➡*[g] U2 → - ∨∨ ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡*[g] V2 & ⦃h, L⦄ ⊢ T1 ➡*[g] T2 & +lemma cpxs_inv_appl1: ∀h,g,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡*[h, g] U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 & ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 & U2 = ⓐV2. T2 - | ∃∃a,W,T. ⦃h, L⦄ ⊢ T1 ➡*[g] ⓛ{a}W.T & ⦃h, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡*[g] U2 - | ∃∃a,V0,V2,V,T. ⦃h, L⦄ ⊢ V1 ➡*[g] V0 & ⇧[0,1] V0 ≡ V2 & - ⦃h, L⦄ ⊢ T1 ➡*[g] ⓓ{a}V.T & ⦃h, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[g] U2. + | ∃∃a,W,T. ⦃G, L⦄ ⊢ T1 ➡*[h, g] ⓛ{a}W.T & ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡*[h, g] U2 + | ∃∃a,V0,V2,V,T. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V0 & ⇧[0,1] V0 ≡ V2 & + ⦃G, L⦄ ⊢ T1 ➡*[h, g] ⓓ{a}V.T & ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, g] U2. #h #g #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 [ /3 width=5/ ] #U #U2 #_ #HU2 * * [ #V0 #T0 #HV10 #HT10 #H destruct @@ -120,9 +120,9 @@ lemma lpx_cpx_trans: ∀h,g. s_r_trans … (cpx h g) (lpx h g). ] qed-. -lemma cpx_bind2: ∀h,g,L,V1,V2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 → - ∀I,T1,T2. ⦃h, L.ⓑ{I}V2⦄ ⊢ T1 ➡[g] T2 → - ∀a. ⦃h, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[g] ⓑ{a,I}V2.T2. +lemma cpx_bind2: ∀h,g,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → + ∀I,T1,T2. ⦃h, L.ⓑ{I}V2⦄ ⊢ T1 ➡[h, g] T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2. #h #g #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 lapply (lpx_cpx_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1/ qed. @@ -132,9 +132,9 @@ qed. lemma lpx_cpxs_trans: ∀h,g. s_rs_trans … (cpx h g) (lpx h g). /3 width=5 by s_r_trans_TC1, lpx_cpx_trans/ qed-. -lemma cpxs_bind2_dx: ∀h,g,L,V1,V2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 → - ∀I,T1,T2. ⦃h, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[g] T2 → - ∀a. ⦃h, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[g] ⓑ{a,I}V2.T2. +lemma cpxs_bind2_dx: ∀h,g,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → + ∀I,T1,T2. ⦃h, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, g] T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2. #h #g #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 lapply (lpx_cpxs_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_lift.ma index 2e86e82f4..7bba480e4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_lift.ma @@ -21,8 +21,8 @@ include "basic_2/computation/cpxs.ma". (* Advanced properties ******************************************************) lemma cpxs_delta: ∀h,g,I,L,K,V,V2,i. - ⇩[0, i] L ≡ K. ⓑ{I}V → ⦃h, K⦄ ⊢ V ➡*[g] V2 → - ∀W2. ⇧[0, i + 1] V2 ≡ W2 → ⦃h, L⦄ ⊢ #i ➡*[g] W2. + ⇩[0, i] L ≡ K. ⓑ{I}V → ⦃h, K⦄ ⊢ V ➡*[h, g] V2 → + ∀W2. ⇧[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡*[h, g] W2. #h #g #I #L #K #V #V2 #i #HLK #H elim H -V2 [ /3 width=9/ ] #V1 #V2 #_ #HV12 #IHV1 #W2 #HVW2 lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK @@ -31,9 +31,9 @@ qed. (* Advanced inversion lemmas ************************************************) -lemma cpxs_inv_lref1: ∀h,g,L,T2,i. ⦃h, L⦄ ⊢ #i ➡*[g] T2 → +lemma cpxs_inv_lref1: ∀h,g,L,T2,i. ⦃G, L⦄ ⊢ #i ➡*[h, g] T2 → T2 = #i ∨ - ∃∃I,K,V1,T1. ⇩[0, i] L ≡ K.ⓑ{I}V1 & ⦃h, K⦄ ⊢ V1 ➡*[g] T1 & + ∃∃I,K,V1,T1. ⇩[0, i] L ≡ K.ⓑ{I}V1 & ⦃h, K⦄ ⊢ V1 ➡*[h, g] T1 & ⇧[0, i + 1] T1 ≡ T2. #h #g #L #T2 #i #H @(cpxs_ind … H) -T2 /2 width=1/ #T #T2 #_ #HT2 * @@ -57,9 +57,9 @@ qed-. (* Properties on supclosure *************************************************) -lemma fsupq_cpxs_trans: ∀h,g,L1,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 ➡*[g] U2 → +lemma fsupq_cpxs_trans: ∀h,g,L1,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 ➡*[h, g] U2 → ∀T1. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡*[g] U1 & ⦃L1, U1⦄ ⊃* ⦃L2, U2⦄. + ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃L1, U1⦄ ⊃* ⦃L2, U2⦄. #h #g #L1 #L2 #T2 #U2 #H @(cpxs_ind_dx … H) -T2 [ /3 width=3/ ] #T #T2 #HT2 #_ #IHTU2 #T1 #HT1 elim (fsupq_cpx_trans … HT1 … HT2) -T #T #HT1 #HT2 @@ -67,8 +67,8 @@ elim (IHTU2 … HT2) -T2 /3 width=3/ qed-. lemma fsups_cpxs_trans: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃* ⦃L2, T2⦄ → - ∀U2. ⦃h, L2⦄ ⊢ T2 ➡*[g] U2 → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡*[g] U1 & ⦃L1, U1⦄ ⊃* ⦃L2, U2⦄. + ∀U2. ⦃h, L2⦄ ⊢ T2 ➡*[h, g] U2 → + ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡*[h, g] U1 & ⦃L1, U1⦄ ⊃* ⦃L2, U2⦄. #h #g #L1 #L2 #T1 #T2 #H @(fsups_ind … H) -L2 -T2 [ /2 width=3/ ] #L #L2 #T #T2 #_ #HT2 #IHT1 #U2 #HTU2 elim (fsupq_cpxs_trans … HTU2 … HT2) -T2 #T2 #HT2 #HTU2 @@ -77,6 +77,6 @@ lapply (fsups_trans … HT2 … HTU2) -L -T2 /2 width=3/ qed-. lemma fsup_ssta_trans: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → - ∀U2,l. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l+1, U2⦄ → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. + ∀U2,l. ⦃h, L2⦄ ⊢ T2 •[h, g] ⦃l+1, U2⦄ → + ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. /3 width=4 by fsup_cpx_trans, ssta_cpx/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_tstc.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_tstc.ma index 87ae419aa..03566d556 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_tstc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_tstc.ma @@ -19,13 +19,13 @@ include "basic_2/computation/lpxs_cpxs.ma". (* Forward lemmas involving same top term constructor ***********************) -lemma cpxs_fwd_cnx: ∀h,g,L,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → ∀U. ⦃h, L⦄ ⊢ T ➡*[g] U → T ≃ U. +lemma cpxs_fwd_cnx: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → ∀U. ⦃G, L⦄ ⊢ T ➡*[h, g] U → T ≃ U. #h #g #L #T #HT #U #H >(cpxs_inv_cnx1 … H HT) -L -T // qed-. -lemma cpxs_fwd_sort: ∀h,g,L,U,k. ⦃h, L⦄ ⊢ ⋆k ➡*[g] U → - ⋆k ≃ U ∨ ⦃h, L⦄ ⊢ ⋆(next h k) ➡*[g] U. +lemma cpxs_fwd_sort: ∀h,g,L,U,k. ⦃G, L⦄ ⊢ ⋆k ➡*[h, g] U → + ⋆k ≃ U ∨ ⦃G, L⦄ ⊢ ⋆(next h k) ➡*[h, g] U. #h #g #L #U #k #H elim (cpxs_inv_sort1 … H) -H #n #l generalize in match k; -k @(nat_ind_plus … n) -n [ #k #_ #H -l destruct /2 width=1/ @@ -41,8 +41,8 @@ elim (cpxs_inv_sort1 … H) -H #n #l generalize in match k; -k @(nat_ind_plus qed-. (* Basic_1: was just: pr3_iso_beta *) -lemma cpxs_fwd_beta: ∀h,g,a,L,V,W,T,U. ⦃h, L⦄ ⊢ ⓐV.ⓛ{a}W.T ➡*[g] U → - ⓐV.ⓛ{a}W.T ≃ U ∨ ⦃h, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[g] U. +lemma cpxs_fwd_beta: ∀h,g,a,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{a}W.T ➡*[h, g] U → + ⓐV.ⓛ{a}W.T ≃ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[h, g] U. #h #g #a #L #V #W #T #U #H elim (cpxs_inv_appl1 … H) -H * [ #V0 #T0 #_ #_ #H destruct /2 width=1/ @@ -58,8 +58,8 @@ qed-. (* Note: probably this is an inversion lemma *) lemma cpxs_fwd_delta: ∀h,g,I,L,K,V1,i. ⇩[0, i] L ≡ K.ⓑ{I}V1 → ∀V2. ⇧[0, i + 1] V1 ≡ V2 → - ∀U. ⦃h, L⦄ ⊢ #i ➡*[g] U → - #i ≃ U ∨ ⦃h, L⦄ ⊢ V2 ➡*[g] U. + ∀U. ⦃G, L⦄ ⊢ #i ➡*[h, g] U → + #i ≃ U ∨ ⦃G, L⦄ ⊢ V2 ➡*[h, g] U. #h #g #I #L #K #V1 #i #HLK #V2 #HV12 #U #H elim (cpxs_inv_lref1 … H) -H /2 width=1/ * #I0 #K0 #V0 #U0 #HLK0 #HVU0 #HU0 @@ -67,9 +67,9 @@ lapply (ldrop_mono … HLK0 … HLK) -HLK0 #H destruct lapply (ldrop_fwd_ldrop2 … HLK) -HLK /3 width=9/ qed-. -lemma cpxs_fwd_theta: ∀h,g,a,L,V1,V,T,U. ⦃h, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ➡*[g] U → +lemma cpxs_fwd_theta: ∀h,g,a,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ➡*[h, g] U → ∀V2. ⇧[0, 1] V1 ≡ V2 → ⓐV1.ⓓ{a}V.T ≃ U ∨ - ⦃h, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[g] U. + ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, g] U. #h #g #a #L #V1 #V #T #U #H #V2 #HV12 elim (cpxs_inv_appl1 … H) -H * [ -HV12 #V0 #T0 #_ #_ #H destruct /2 width=1/ @@ -99,8 +99,8 @@ elim (cpxs_inv_appl1 … H) -H * ] qed-. -lemma cpxs_fwd_cast: ∀h,g,L,W,T,U. ⦃h, L⦄ ⊢ ⓝW.T ➡*[g] U → - ∨∨ ⓝW. T ≃ U | ⦃h, L⦄ ⊢ T ➡*[g] U | ⦃h, L⦄ ⊢ W ➡*[g] U. +lemma cpxs_fwd_cast: ∀h,g,L,W,T,U. ⦃G, L⦄ ⊢ ⓝW.T ➡*[h, g] U → + ∨∨ ⓝW. T ≃ U | ⦃G, L⦄ ⊢ T ➡*[h, g] U | ⦃G, L⦄ ⊢ W ➡*[h, g] U. #h #g #L #W #T #U #H elim (cpxs_inv_cast1 … H) -H /2 width=1/ * #W0 #T0 #_ #_ #H destruct /2 width=1/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_tstc_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_tstc_vector.ma index 61f43dda6..a166abc43 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_tstc_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_tstc_vector.ma @@ -21,8 +21,8 @@ include "basic_2/computation/cpxs_tstc.ma". (* Vector form of forward lemmas involving same top term constructor ********) (* Basic_1: was just: nf2_iso_appls_lref *) -lemma cpxs_fwd_cnx_vector: ∀h,g,L,T. 𝐒⦃T⦄ → ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → - ∀Vs,U. ⦃h, L⦄ ⊢ ⒶVs.T ➡*[g] U → ⒶVs.T ≃ U. +lemma cpxs_fwd_cnx_vector: ∀h,g,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → + ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, g] U → ⒶVs.T ≃ U. #h #g #L #T #H1T #H2T #Vs elim Vs -Vs [ @(cpxs_fwd_cnx … H2T) ] (**) (* /2 width=3 by cpxs_fwd_cnx/ does not work *) #V #Vs #IHVs #U #H elim (cpxs_inv_appl1 … H) -H * @@ -36,8 +36,8 @@ elim (cpxs_inv_appl1 … H) -H * ] qed-. -lemma cpxs_fwd_sort_vector: ∀h,g,L,k,Vs,U. ⦃h, L⦄ ⊢ ⒶVs.⋆k ➡*[g] U → - ⒶVs.⋆k ≃ U ∨ ⦃h, L⦄ ⊢ ⒶVs.⋆(next h k) ➡*[g] U. +lemma cpxs_fwd_sort_vector: ∀h,g,L,k,Vs,U. ⦃G, L⦄ ⊢ ⒶVs.⋆k ➡*[h, g] U → + ⒶVs.⋆k ≃ U ∨ ⦃G, L⦄ ⊢ ⒶVs.⋆(next h k) ➡*[h, g] U. #h #g #L #k #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_sort/ #V #Vs #IHVs #U #H elim (cpxs_inv_appl1 … H) -H * @@ -61,8 +61,8 @@ qed-. (* Basic_1: was just: pr3_iso_appls_beta *) -lemma cpxs_fwd_beta_vector: ∀h,g,a,L,Vs,V,W,T,U. ⦃h, L⦄ ⊢ ⒶVs.ⓐV.ⓛ{a}W.T ➡*[g] U → - ⒶVs. ⓐV. ⓛ{a}W. T ≃ U ∨ ⦃h, L⦄ ⊢ ⒶVs.ⓓ{a}ⓝW.V.T ➡*[g] U. +lemma cpxs_fwd_beta_vector: ∀h,g,a,L,Vs,V,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓐV.ⓛ{a}W.T ➡*[h, g] U → + ⒶVs. ⓐV. ⓛ{a}W. T ≃ U ∨ ⦃G, L⦄ ⊢ ⒶVs.ⓓ{a}ⓝW.V.T ➡*[h, g] U. #h #g #a #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_beta/ #V0 #Vs #IHVs #V #W #T #U #H elim (cpxs_inv_appl1 … H) -H * @@ -86,8 +86,8 @@ qed-. lemma cpxs_fwd_delta_vector: ∀h,g,I,L,K,V1,i. ⇩[0, i] L ≡ K.ⓑ{I}V1 → ∀V2. ⇧[0, i + 1] V1 ≡ V2 → - ∀Vs,U. ⦃h, L⦄ ⊢ ⒶVs.#i ➡*[g] U → - ⒶVs.#i ≃ U ∨ ⦃h, L⦄ ⊢ ⒶVs.V2 ➡*[g] U. + ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.#i ➡*[h, g] U → + ⒶVs.#i ≃ U ∨ ⦃G, L⦄ ⊢ ⒶVs.V2 ➡*[h, g] U. #h #g #I #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs /2 width=5 by cpxs_fwd_delta/ #V #Vs #IHVs #U #H -K -V1 elim (cpxs_inv_appl1 … H) -H * @@ -111,8 +111,8 @@ qed-. (* Basic_1: was just: pr3_iso_appls_abbr *) lemma cpxs_fwd_theta_vector: ∀h,g,L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → - ∀a,V,T,U. ⦃h, L⦄ ⊢ ⒶV1s.ⓓ{a}V.T ➡*[g] U → - ⒶV1s. ⓓ{a}V. T ≃ U ∨ ⦃h, L⦄ ⊢ ⓓ{a}V.ⒶV2s.T ➡*[g] U. + ∀a,V,T,U. ⦃G, L⦄ ⊢ ⒶV1s.ⓓ{a}V.T ➡*[h, g] U → + ⒶV1s. ⓓ{a}V. T ≃ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}V.ⒶV2s.T ➡*[h, g] U. #h #g #L #V1s #V2s * -V1s -V2s /3 width=1/ #V1s #V2s #V1a #V2a #HV12a #HV12s #a generalize in match HV12a; -HV12a @@ -159,10 +159,10 @@ elim (cpxs_inv_appl1 … H) -H * qed-. (* Basic_1: was just: pr3_iso_appls_cast *) -lemma cpxs_fwd_cast_vector: ∀h,g,L,Vs,W,T,U. ⦃h, L⦄ ⊢ ⒶVs.ⓝW.T ➡*[g] U → +lemma cpxs_fwd_cast_vector: ∀h,g,L,Vs,W,T,U. ⦃G, L⦄ ⊢ ⒶVs.ⓝW.T ➡*[h, g] U → ∨∨ ⒶVs. ⓝW. T ≃ U - | ⦃h, L⦄ ⊢ ⒶVs.T ➡*[g] U - | ⦃h, L⦄ ⊢ ⒶVs.W ➡*[g] U. + | ⦃G, L⦄ ⊢ ⒶVs.T ➡*[h, g] U + | ⦃G, L⦄ ⊢ ⒶVs.W ➡*[h, g] U. #h #g #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_cast/ #V #Vs #IHVs #W #T #U #H elim (cpxs_inv_appl1 … H) -H * diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/csn.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/csn.ma index 9aead33d3..e9eca9679 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/csn.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/csn.ma @@ -27,11 +27,11 @@ interpretation (* Basic eliminators ********************************************************) lemma csn_ind: ∀h,g,L. ∀R:predicate term. - (∀T1. ⦃h, L⦄ ⊢ ⬊*[g] T1 → - (∀T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → (T1 = T2 → ⊥) → R T2) → + (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → + (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1 ) → - ∀T. ⦃h, L⦄ ⊢ ⬊*[g] T → R T. + ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R T. #h #g #L #R #H0 #T1 #H elim H -T1 #T1 #HT1 #IHT1 @H0 -H0 /3 width=1/ -IHT1 /4 width=1/ qed-. @@ -40,12 +40,12 @@ qed-. (* Basic_1: was just: sn3_pr2_intro *) lemma csn_intro: ∀h,g,L,T1. - (∀T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → (T1 = T2 → ⊥) → ⦃h, L⦄ ⊢ ⬊*[g] T2) → - ⦃h, L⦄ ⊢ ⬊*[g] T1. + (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] T2) → + ⦃G, L⦄ ⊢ ⬊*[h, g] T1. /4 width=1/ qed. -lemma csn_cpx_trans: ∀h,g,L,T1. ⦃h, L⦄ ⊢ ⬊*[g] T1 → - ∀T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → ⦃h, L⦄ ⊢ ⬊*[g] T2. +lemma csn_cpx_trans: ∀h,g,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → + ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ⦃G, L⦄ ⊢ ⬊*[h, g] T2. #h #g #L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12 @csn_intro #T #HLT2 #HT2 elim (term_eq_dec T1 T2) #HT12 @@ -54,10 +54,10 @@ elim (term_eq_dec T1 T2) #HT12 qed-. (* Basic_1: was just: sn3_nf2 *) -lemma cnx_csn: ∀h,g,L,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → ⦃h, L⦄ ⊢ ⬊*[g] T. +lemma cnx_csn: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ ⬊*[h, g] T. /2 width=1/ qed. -lemma cnx_sort: ∀h,g,L,k. ⦃h, L⦄ ⊢ ⬊*[g] ⋆k. +lemma cnx_sort: ∀h,g,L,k. ⦃G, L⦄ ⊢ ⬊*[h, g] ⋆k. #h #g #L #k elim (deg_total h g k) #l generalize in match k; -k @(nat_ind_plus … l) -l /3 width=1/ #l #IHl #k #Hkl lapply (deg_next_SO … Hkl) -Hkl @@ -68,8 +68,8 @@ lemma cnx_sort: ∀h,g,L,k. ⦃h, L⦄ ⊢ ⬊*[g] ⋆k. qed. (* Basic_1: was just: sn3_cast *) -lemma csn_cast: ∀h,g,L,W. ⦃h, L⦄ ⊢ ⬊*[g] W → - ∀T. ⦃h, L⦄ ⊢ ⬊*[g] T → ⦃h, L⦄ ⊢ ⬊*[g] ⓝW.T. +lemma csn_cast: ∀h,g,L,W. ⦃G, L⦄ ⊢ ⬊*[h, g] W → + ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓝW.T. #h #g #L #W #HW @(csn_ind … HW) -W #W #HW #IHW #T #HT @(csn_ind … HT) -T #T #HT #IHT @csn_intro #X #H1 #H2 elim (cpx_inv_cast1 … H1) -H1 @@ -84,37 +84,37 @@ qed. (* Basic forward lemmas *****************************************************) -fact csn_fwd_pair_sn_aux: ∀h,g,L,U. ⦃h, L⦄ ⊢ ⬊*[g] U → - ∀I,V,T. U = ②{I}V.T → ⦃h, L⦄ ⊢ ⬊*[g] V. +fact csn_fwd_pair_sn_aux: ∀h,g,L,U. ⦃G, L⦄ ⊢ ⬊*[h, g] U → + ∀I,V,T. U = ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] V. #h #g #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct @csn_intro #V2 #HLV2 #HV2 @(IH (②{I}V2.T)) -IH // /2 width=1/ -HLV2 #H destruct /2 width=1/ qed-. (* Basic_1: was just: sn3_gen_head *) -lemma csn_fwd_pair_sn: ∀h,g,I,L,V,T. ⦃h, L⦄ ⊢ ⬊*[g] ②{I}V.T → ⦃h, L⦄ ⊢ ⬊*[g] V. +lemma csn_fwd_pair_sn: ∀h,g,I,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ②{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] V. /2 width=5 by csn_fwd_pair_sn_aux/ qed-. -fact csn_fwd_bind_dx_aux: ∀h,g,L,U. ⦃h, L⦄ ⊢ ⬊*[g] U → - ∀a,I,V,T. U = ⓑ{a,I}V.T → ⦃h, L.ⓑ{I}V⦄ ⊢ ⬊*[g] T. +fact csn_fwd_bind_dx_aux: ∀h,g,L,U. ⦃G, L⦄ ⊢ ⬊*[h, g] U → + ∀a,I,V,T. U = ⓑ{a,I}V.T → ⦃h, L.ⓑ{I}V⦄ ⊢ ⬊*[h, g] T. #h #g #L #U #H elim H -H #U0 #_ #IH #a #I #V #T #H destruct @csn_intro #T2 #HLT2 #HT2 @(IH (ⓑ{a,I} V. T2)) -IH // /2 width=1/ -HLT2 #H destruct /2 width=1/ qed-. (* Basic_1: was just: sn3_gen_bind *) -lemma csn_fwd_bind_dx: ∀h,g,a,I,L,V,T. ⦃h, L⦄ ⊢ ⬊*[g] ⓑ{a,I}V.T → ⦃h, L.ⓑ{I}V⦄ ⊢ ⬊*[g] T. +lemma csn_fwd_bind_dx: ∀h,g,a,I,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓑ{a,I}V.T → ⦃h, L.ⓑ{I}V⦄ ⊢ ⬊*[h, g] T. /2 width=4 by csn_fwd_bind_dx_aux/ qed-. -fact csn_fwd_flat_dx_aux: ∀h,g,L,U. ⦃h, L⦄ ⊢ ⬊*[g] U → - ∀I,V,T. U = ⓕ{I}V.T → ⦃h, L⦄ ⊢ ⬊*[g] T. +fact csn_fwd_flat_dx_aux: ∀h,g,L,U. ⦃G, L⦄ ⊢ ⬊*[h, g] U → + ∀I,V,T. U = ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] T. #h #g #L #U #H elim H -H #U0 #_ #IH #I #V #T #H destruct @csn_intro #T2 #HLT2 #HT2 @(IH (ⓕ{I}V.T2)) -IH // /2 width=1/ -HLT2 #H destruct /2 width=1/ qed-. (* Basic_1: was just: sn3_gen_flat *) -lemma csn_fwd_flat_dx: ∀h,g,I,L,V,T. ⦃h, L⦄ ⊢ ⬊*[g] ⓕ{I}V.T → ⦃h, L⦄ ⊢ ⬊*[g] T. +lemma csn_fwd_flat_dx: ∀h,g,I,L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓕ{I}V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] T. /2 width=5 by csn_fwd_flat_dx_aux/ qed-. (* Basic_1: removed theorems 14: diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_aaa.ma index cd8d5a64b..f5229400f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_aaa.ma @@ -19,7 +19,7 @@ include "basic_2/computation/csn_tstc_vector.ma". (* Main properties concerning atomic arity assignment ***********************) -theorem csn_aaa: ∀h,g,L,T,A. L ⊢ T ⁝ A → ⦃h, L⦄ ⊢ ⬊*[g] T. +theorem csn_aaa: ∀h,g,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ⦃G, L⦄ ⊢ ⬊*[h, g] T. #h #g #L #T #A #H @(acp_aaa … (csn_acp h g) (csn_acr h g) … H) qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_alt.ma index 954cb9c01..07716de62 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_alt.ma @@ -29,10 +29,10 @@ interpretation (* Basic eliminators ********************************************************) lemma csna_ind: ∀h,g,L. ∀R:predicate term. - (∀T1. ⦃h, L⦄ ⊢ ⬊⬊*[g] T1 → - (∀T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → (T1 = T2 → ⊥) → R T2) → R T1 + (∀T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1 → + (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1 ) → - ∀T. ⦃h, L⦄ ⊢ ⬊⬊*[g] T → R T. + ∀T. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T → R T. #h #g #L #R #H0 #T1 #H elim H -T1 #T1 #HT1 #IHT1 @H0 -H0 /3 width=1/ -IHT1 /4 width=1/ qed-. @@ -41,18 +41,18 @@ qed-. (* Basic_1: was just: sn3_intro *) lemma csna_intro: ∀h,g,L,T1. - (∀T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → (T1 = T2 → ⊥) → ⦃h, L⦄ ⊢ ⬊⬊*[g] T2) → - ⦃h, L⦄ ⊢ ⬊⬊*[g] T1. + (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2) → + ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1. /4 width=1/ qed. fact csna_intro_aux: ∀h,g,L,T1. ( - ∀T,T2. ⦃h, L⦄ ⊢ T ➡*[g] T2 → T1 = T → (T1 = T2 → ⊥) → ⦃h, L⦄ ⊢ ⬊⬊*[g] T2 - ) → ⦃h, L⦄ ⊢ ⬊⬊*[g] T1. + ∀T,T2. ⦃G, L⦄ ⊢ T ➡*[h, g] T2 → T1 = T → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2 + ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1. /4 width=3/ qed-. (* Basic_1: was just: sn3_pr3_trans (old version) *) -lemma csna_cpxs_trans: ∀h,g,L,T1. ⦃h, L⦄ ⊢ ⬊⬊*[g] T1 → - ∀T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ ⬊⬊*[g] T2. +lemma csna_cpxs_trans: ∀h,g,L,T1. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1 → + ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2. #h #g #L #T1 #H elim H -T1 #T1 #HT1 #IHT1 #T2 #HLT12 @csna_intro #T #HLT2 #HT2 elim (term_eq_dec T1 T2) #HT12 @@ -62,8 +62,8 @@ qed. (* Basic_1: was just: sn3_pr2_intro (old version) *) lemma csna_intro_cpx: ∀h,g,L,T1. ( - ∀T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → (T1 = T2 → ⊥) → ⦃h, L⦄ ⊢ ⬊⬊*[g] T2 - ) → ⦃h, L⦄ ⊢ ⬊⬊*[g] T1. + ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T2 + ) → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T1. #h #g #L #T1 #H @csna_intro_aux #T #T2 #H @(cpxs_ind_dx … H) -T [ -H #H destruct #H @@ -78,27 +78,27 @@ qed. (* Main properties **********************************************************) -theorem csn_csna: ∀h,g,L,T. ⦃h, L⦄ ⊢ ⬊*[g] T → ⦃h, L⦄ ⊢ ⬊⬊*[g] T. +theorem csn_csna: ∀h,g,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T. #h #g #L #T #H @(csn_ind … H) -T /4 width=1/ qed. -theorem csna_csn: ∀h,g,L,T. ⦃h, L⦄ ⊢ ⬊⬊*[g] T → ⦃h, L⦄ ⊢ ⬊*[g] T. +theorem csna_csn: ∀h,g,L,T. ⦃G, L⦄ ⊢ ⬊⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] T. #h #g #L #T #H @(csna_ind … H) -T /4 width=1/ qed. (* Basic_1: was just: sn3_pr3_trans *) -lemma csn_cpxs_trans: ∀h,g,L,T1. ⦃h, L⦄ ⊢ ⬊*[g] T1 → - ∀T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → ⦃h, L⦄ ⊢ ⬊*[g] T2. +lemma csn_cpxs_trans: ∀h,g,L,T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → + ∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → ⦃G, L⦄ ⊢ ⬊*[h, g] T2. #h #g #L #T1 #HT1 #T2 #H @(cpxs_ind … H) -T2 // /2 width=3 by csn_cpx_trans/ qed-. (* Main eliminators *********************************************************) lemma csn_ind_alt: ∀h,g,L. ∀R:predicate term. - (∀T1. ⦃h, L⦄ ⊢ ⬊*[g] T1 → - (∀T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → (T1 = T2 → ⊥) → R T2) → R T1 + (∀T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → + (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 = T2 → ⊥) → R T2) → R T1 ) → - ∀T. ⦃h, L⦄ ⊢ ⬊*[g] T → R T. + ∀T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → R T. #h #g #L #R #H0 #T1 #H @(csna_ind … (csn_csna … H)) -T1 #T1 #HT1 #IHT1 @H0 -H0 /2 width=1/ -HT1 /3 width=1/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_lift.ma index e57f37d9d..7a4513fb9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_lift.ma @@ -21,8 +21,8 @@ include "basic_2/computation/csn.ma". (* Relocation properties ****************************************************) (* Basic_1: was just: sn3_lift *) -lemma csn_lift: ∀h,g,L2,L1,T1,d,e. ⦃h, L1⦄ ⊢ ⬊*[g] T1 → - ∀T2. ⇩[d, e] L2 ≡ L1 → ⇧[d, e] T1 ≡ T2 → ⦃h, L2⦄ ⊢ ⬊*[g] T2. +lemma csn_lift: ∀h,g,L2,L1,T1,d,e. ⦃h, L1⦄ ⊢ ⬊*[h, g] T1 → + ∀T2. ⇩[d, e] L2 ≡ L1 → ⇧[d, e] T1 ≡ T2 → ⦃h, L2⦄ ⊢ ⬊*[h, g] T2. #h #g #L2 #L1 #T1 #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12 @csn_intro #T #HLT2 #HT2 elim (cpx_inv_lift1 … HLT2 … HL21 … HT12) -HLT2 #T0 #HT0 #HLT10 @@ -31,8 +31,8 @@ elim (cpx_inv_lift1 … HLT2 … HL21 … HT12) -HLT2 #T0 #HT0 #HLT10 qed. (* Basic_1: was just: sn3_gen_lift *) -lemma csn_inv_lift: ∀h,g,L2,L1,T1,d,e. ⦃h, L1⦄ ⊢ ⬊*[g] T1 → - ∀T2. ⇩[d, e] L1 ≡ L2 → ⇧[d, e] T2 ≡ T1 → ⦃h, L2⦄ ⊢ ⬊*[g] T2. +lemma csn_inv_lift: ∀h,g,L2,L1,T1,d,e. ⦃h, L1⦄ ⊢ ⬊*[h, g] T1 → + ∀T2. ⇩[d, e] L1 ≡ L2 → ⇧[d, e] T2 ≡ T1 → ⦃h, L2⦄ ⊢ ⬊*[h, g] T2. #h #g #L2 #L1 #T1 #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21 @csn_intro #T #HLT2 #HT2 elim (lift_total T d e) #T0 #HT0 @@ -44,7 +44,7 @@ qed. (* Advanced properties ******************************************************) (* Basic_1: was just: sn3_abbr *) -lemma csn_lref_bind: ∀h,g,I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃h, K⦄ ⊢ ⬊*[g] V → ⦃h, L⦄ ⊢ ⬊*[g] #i. +lemma csn_lref_bind: ∀h,g,I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃h, K⦄ ⊢ ⬊*[h, g] V → ⦃G, L⦄ ⊢ ⬊*[h, g] #i. #h #g #I #L #K #V #i #HLK #HV @csn_intro #X #H #Hi elim (cpx_inv_lref1 … H) -H @@ -57,9 +57,9 @@ elim (cpx_inv_lref1 … H) -H ] qed. -lemma csn_appl_simple: ∀h,g,L,V. ⦃h, L⦄ ⊢ ⬊*[g] V → ∀T1. - (∀T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → (T1 = T2 → ⊥) → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.T2) → - 𝐒⦃T1⦄ → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.T1. +lemma csn_appl_simple: ∀h,g,L,V. ⦃G, L⦄ ⊢ ⬊*[h, g] V → ∀T1. + (∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → (T1 = T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T2) → + 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T1. #h #g #L #V #H @(csn_ind … H) -V #V #_ #IHV #T1 #IHT1 #HT1 @csn_intro #X #H1 #H2 elim (cpx_inv_appl1_simple … H1) // -H1 @@ -80,7 +80,7 @@ qed. (* Basic_1: was: sn3_gen_def *) lemma csn_inv_lref_bind: ∀h,g,I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → - ⦃h, L⦄ ⊢ ⬊*[g] #i → ⦃h, K⦄ ⊢ ⬊*[g] V. + ⦃G, L⦄ ⊢ ⬊*[h, g] #i → ⦃h, K⦄ ⊢ ⬊*[h, g] V. #h #g #I #L #K #V #i #HLK #Hi elim (lift_total V 0 (i+1)) #V0 #HV0 lapply (ldrop_fwd_ldrop2 … HLK) #H0LK diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_lpx.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_lpx.ma index 3f97c882a..ca40d78e9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_lpx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_lpx.ma @@ -21,15 +21,15 @@ include "basic_2/computation/csn_lift.ma". (* Advanced properties ******************************************************) -lemma csn_lpx_conf: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡[g] L2 → - ∀T. ⦃h, L1⦄ ⊢ ⬊*[g] T → ⦃h, L2⦄ ⊢ ⬊*[g] T. +lemma csn_lpx_conf: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡[h, g] L2 → + ∀T. ⦃h, L1⦄ ⊢ ⬊*[h, g] T → ⦃h, L2⦄ ⊢ ⬊*[h, g] T. #h #g #L1 #L2 #HL12 #T #H @(csn_ind_alt … H) -T #T #_ #IHT @csn_intro #T0 #HLT0 #HT0 @IHT /2 width=2/ -IHT -HT0 /2 width=3 by lpx_cpx_trans/ qed. -lemma csn_abst: ∀h,g,a,L,W. ⦃h, L⦄ ⊢ ⬊*[g] W → - ∀T. ⦃h, L.ⓛW⦄ ⊢ ⬊*[g] T → ⦃h, L⦄ ⊢ ⬊*[g] ⓛ{a}W.T. +lemma csn_abst: ∀h,g,a,L,W. ⦃G, L⦄ ⊢ ⬊*[h, g] W → + ∀T. ⦃h, L.ⓛW⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓛ{a}W.T. #h #g #a #L #W #HW @(csn_ind … HW) -W #W #_ #IHW #T #HT @(csn_ind … HT) -T #T #HT #IHT @csn_intro #X #H1 #H2 elim (cpx_inv_abst1 … H1) -H1 @@ -41,8 +41,8 @@ elim (eq_false_inv_tpair_sn … H2) -H2 ] qed. -lemma csn_abbr: ∀h,g,a,L,V. ⦃h, L⦄ ⊢ ⬊*[g] V → - ∀T. ⦃h, L.ⓓV⦄ ⊢ ⬊*[g] T → ⦃h, L⦄ ⊢ ⬊*[g] ⓓ{a}V. T. +lemma csn_abbr: ∀h,g,a,L,V. ⦃G, L⦄ ⊢ ⬊*[h, g] V → + ∀T. ⦃h, L.ⓓV⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓓ{a}V. T. #h #g #a #L #V #HV elim HV -V #V #_ #IHV #T #HT @(csn_ind_alt … HT) -T #T #HT #IHT @csn_intro #X #H1 #H2 elim (cpx_inv_abbr1 … H1) -H1 * @@ -58,8 +58,8 @@ elim (cpx_inv_abbr1 … H1) -H1 * ] qed. -fact csn_appl_beta_aux: ∀h,g,a,L,U1. ⦃h, L⦄ ⊢ ⬊*[g] U1 → - ∀V,W,T1. U1 = ⓓ{a}ⓝW.V.T1 → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.ⓛ{a}W.T1. +fact csn_appl_beta_aux: ∀h,g,a,L,U1. ⦃G, L⦄ ⊢ ⬊*[h, g] U1 → + ∀V,W,T1. U1 = ⓓ{a}ⓝW.V.T1 → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.ⓛ{a}W.T1. #h #g #a #L #X #H @(csn_ind … H) -X #X #HT1 #IHT1 #V #W #T1 #H1 destruct @csn_intro #X #H1 #H2 @@ -76,11 +76,11 @@ elim (cpx_inv_appl1 … H1) -H1 * qed-. (* Basic_1: was just: sn3_beta *) -lemma csn_appl_beta: ∀h,g,a,L,V,W,T. ⦃h, L⦄ ⊢ ⬊*[g] ⓓ{a}ⓝW.V.T → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.ⓛ{a}W.T. +lemma csn_appl_beta: ∀h,g,a,L,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓓ{a}ⓝW.V.T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.ⓛ{a}W.T. /2 width=3 by csn_appl_beta_aux/ qed. -fact csn_appl_theta_aux: ∀h,g,a,L,U. ⦃h, L⦄ ⊢ ⬊*[g] U → ∀V1,V2. ⇧[0, 1] V1 ≡ V2 → - ∀V,T. U = ⓓ{a}V.ⓐV2.T → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV1.ⓓ{a}V.T. +fact csn_appl_theta_aux: ∀h,g,a,L,U. ⦃G, L⦄ ⊢ ⬊*[h, g] U → ∀V1,V2. ⇧[0, 1] V1 ≡ V2 → + ∀V,T. U = ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV1.ⓓ{a}V.T. #h #g #a #L #X #H @(csn_ind_alt … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct lapply (csn_fwd_pair_sn … HVT) #HV lapply (csn_fwd_bind_dx … HVT) -HVT #HVT @@ -117,13 +117,13 @@ elim (cpx_inv_appl1 … HL) -HL * qed-. lemma csn_appl_theta: ∀h,g,a,V1,V2. ⇧[0, 1] V1 ≡ V2 → - ∀L,V,T. ⦃h, L⦄ ⊢ ⬊*[g] ⓓ{a}V.ⓐV2.T → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV1.ⓓ{a}V.T. + ∀L,V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓓ{a}V.ⓐV2.T → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV1.ⓓ{a}V.T. /2 width=5 by csn_appl_theta_aux/ qed. (* Basic_1: was just: sn3_appl_appl *) -lemma csn_appl_simple_tstc: ∀h,g,L,V. ⦃h, L⦄ ⊢ ⬊*[g] V → ∀T1. ⦃h, L⦄ ⊢ ⬊*[g] T1 → - (∀T2. ⦃h, L⦄ ⊢ T1 ➡*[g] T2 → (T1 ≃ T2 → ⊥) → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.T2) → - 𝐒⦃T1⦄ → ⦃h, L⦄ ⊢ ⬊*[g] ⓐV.T1. +lemma csn_appl_simple_tstc: ∀h,g,L,V. ⦃G, L⦄ ⊢ ⬊*[h, g] V → ∀T1. ⦃G, L⦄ ⊢ ⬊*[h, g] T1 → + (∀T2. ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2 → (T1 ≃ T2 → ⊥) → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T2) → + 𝐒⦃T1⦄ → ⦃G, L⦄ ⊢ ⬊*[h, g] ⓐV.T1. #h #g #L #V #H @(csn_ind … H) -V #V #_ #IHV #T1 #H @(csn_ind … H) -T1 #T1 #H1T1 #IHT1 #H2T1 #H3T1 @csn_intro #X #HL #H elim (cpx_inv_appl1_simple … HL) -HL // diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_tstc_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_tstc_vector.ma index 123704979..4b19e3db2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_tstc_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_tstc_vector.ma @@ -22,8 +22,8 @@ include "basic_2/computation/csn_vector.ma". (* Advanced properties ******************************************************) (* Basic_1: was just: sn3_appls_lref *) -lemma csn_applv_cnx: ∀h,g,L,T. 𝐒⦃T⦄ → ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → - ∀Vs. ⦃h, L⦄ ⊢ ⬊*[g] Vs → ⦃h, L⦄ ⊢ ⬊*[g] ⒶVs.T. +lemma csn_applv_cnx: ∀h,g,L,T. 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → + ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Vs → ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.T. #h #g #L #T #H1T #H2T #Vs elim Vs -Vs [ #_ @(cnx_csn … H2T) ] (**) (* /2 width=1/ does not work *) #V #Vs #IHV #H elim (csnv_inv_cons … H) -H #HV #HVs @@ -33,7 +33,7 @@ lapply (cpxs_fwd_cnx_vector … H) -H // -H1T -H2T #H elim (H0) -H0 // qed. -lemma csn_applv_sort: ∀h,g,L,k,Vs. ⦃h, L⦄ ⊢ ⬊*[g] Vs → ⦃h, L⦄ ⊢ ⬊*[g] ⒶVs.⋆k. +lemma csn_applv_sort: ∀h,g,L,k,Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Vs → ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.⋆k. #h #g #L #k elim (deg_total h g k) #l generalize in match k; -k @(nat_ind_plus … l) -l [ /3 width=1/ ] #l #IHl #k #Hkl lapply (deg_next_SO … Hkl) -Hkl @@ -49,8 +49,8 @@ elim (cpxs_fwd_sort_vector … H) -H #H qed. (* Basic_1: was just: sn3_appls_beta *) -lemma csn_applv_beta: ∀h,g,a,L,Vs,V,W,T. ⦃h, L⦄ ⊢ ⬊*[g] ⒶVs.ⓓ{a}ⓝW.V.T → - ⦃h, L⦄ ⊢ ⬊*[g] ⒶVs. ⓐV.ⓛ{a}W.T. +lemma csn_applv_beta: ∀h,g,a,L,Vs,V,W,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.ⓓ{a}ⓝW.V.T → + ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs. ⓐV.ⓛ{a}W.T. #h #g #a #L #Vs elim Vs -Vs /2 width=1/ #V0 #Vs #IHV #V #W #T #H1T lapply (csn_fwd_pair_sn … H1T) #HV0 @@ -65,7 +65,7 @@ qed. lemma csn_applv_delta: ∀h,g,I,L,K,V1,i. ⇩[0, i] L ≡ K.ⓑ{I}V1 → ∀V2. ⇧[0, i + 1] V1 ≡ V2 → - ∀Vs. ⦃h, L⦄ ⊢ ⬊*[g] (ⒶVs.V2) → ⦃h, L⦄ ⊢ ⬊*[g] (ⒶVs.#i). + ∀Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] (ⒶVs.V2) → ⦃G, L⦄ ⊢ ⬊*[h, g] (ⒶVs.#i). #h #g #I #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs [ #H lapply (ldrop_fwd_ldrop2 … HLK) #HLK0 @@ -84,8 +84,8 @@ qed. (* Basic_1: was just: sn3_appls_abbr *) lemma csn_applv_theta: ∀h,g,a,L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → - ∀V,T. ⦃h, L⦄ ⊢ ⬊*[g] ⓓ{a}V.ⒶV2s.T → - ⦃h, L⦄ ⊢ ⬊*[g] ⒶV1s.ⓓ{a}V.T. + ∀V,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⓓ{a}V.ⒶV2s.T → + ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶV1s.ⓓ{a}V.T. #h #g #a #L #V1s #V2s * -V1s -V2s /2 width=1/ #V1s #V2s #V1 #V2 #HV12 #H generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1 @@ -102,8 +102,8 @@ elim (cpxs_fwd_theta_vector … (V2@V2s) … H1) -H1 /2 width=1/ -HV12s -HV12 qed. (* Basic_1: was just: sn3_appls_cast *) -lemma csn_applv_cast: ∀h,g,L,Vs,W,T. ⦃h, L⦄ ⊢ ⬊*[g] ⒶVs.W → ⦃h, L⦄ ⊢ ⬊*[g] ⒶVs.T → - ⦃h, L⦄ ⊢ ⬊*[g] ⒶVs.ⓝW.T. +lemma csn_applv_cast: ∀h,g,L,Vs,W,T. ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.W → ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.T → + ⦃G, L⦄ ⊢ ⬊*[h, g] ⒶVs.ⓝW.T. #h #g #L #Vs elim Vs -Vs /2 width=1/ #V #Vs #IHV #W #T #H1W #H1T lapply (csn_fwd_pair_sn … H1W) #HV @@ -118,7 +118,7 @@ elim (cpxs_fwd_cast_vector … H) -H #H ] qed. -theorem csn_acr: ∀h,g. acr (cpx h g) (eq …) (csn h g) (λL,T. ⦃h, L⦄ ⊢ ⬊*[g] T). +theorem csn_acr: ∀h,g. acr (cpx h g) (eq …) (csn h g) (λL,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T). #h #g @mk_acr // [ /3 width=1/ |2,3,6: /2 width=1/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_vector.ma index c655be403..f29cedc04 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_vector.ma @@ -26,14 +26,14 @@ interpretation (* Basic inversion lemmas ***************************************************) -lemma csnv_inv_cons: ∀h,g,L,T,Ts. ⦃h, L⦄ ⊢ ⬊*[g] T @ Ts → - ⦃h, L⦄ ⊢ ⬊*[g] T ∧ ⦃h, L⦄ ⊢ ⬊*[g] Ts. +lemma csnv_inv_cons: ∀h,g,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, g] T @ Ts → + ⦃G, L⦄ ⊢ ⬊*[h, g] T ∧ ⦃G, L⦄ ⊢ ⬊*[h, g] Ts. normalize // qed-. (* Basic forward lemmas *****************************************************) -lemma csn_fwd_applv: ∀h,g,L,T,Vs. ⦃h, L⦄ ⊢ ⬊*[g] Ⓐ Vs.T → - ⦃h, L⦄ ⊢ ⬊*[g] Vs ∧ ⦃h, L⦄ ⊢ ⬊*[g] T. +lemma csn_fwd_applv: ∀h,g,L,T,Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Ⓐ Vs.T → + ⦃G, L⦄ ⊢ ⬊*[h, g] Vs ∧ ⦃G, L⦄ ⊢ ⬊*[h, g] T. #h #g #L #T #Vs elim Vs -Vs /2 width=1/ #V #Vs #IHVs #HVs lapply (csn_fwd_pair_sn … HVs) #HV diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lprs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lprs.ma index 2916145ee..2e6a25167 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lprs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lprs.ma @@ -26,14 +26,14 @@ interpretation "parallel computation (local environment, sn variant)" (* Basic eliminators ********************************************************) lemma lprs_ind: ∀L1. ∀R:predicate lenv. R L1 → - (∀L,L2. L1 ⊢ ➡* L → L ⊢ ➡ L2 → R L → R L2) → + (∀L,L2. L1 ⊢ ➡* L → ⦃G, L⦄ ⊢ ➡ L2 → R L → R L2) → ∀L2. L1 ⊢ ➡* L2 → R L2. #L1 #R #HL1 #IHL1 #L2 #HL12 @(TC_star_ind … HL1 IHL1 … HL12) // qed-. lemma lprs_ind_dx: ∀L2. ∀R:predicate lenv. R L2 → - (∀L1,L. L1 ⊢ ➡ L → L ⊢ ➡* L2 → R L → R L1) → + (∀L1,L. L1 ⊢ ➡ L → ⦃G, L⦄ ⊢ ➡* L2 → R L → R L1) → ∀L1. L1 ⊢ ➡* L2 → R L1. #L2 #R #HL2 #IHL2 #L1 #HL12 @(TC_star_ind_dx … HL2 IHL2 … HL12) // @@ -44,13 +44,13 @@ qed-. lemma lpr_lprs: ∀L1,L2. L1 ⊢ ➡ L2 → L1 ⊢ ➡* L2. /2 width=1/ qed. -lemma lprs_refl: ∀L. L ⊢ ➡* L. +lemma lprs_refl: ∀L. ⦃G, L⦄ ⊢ ➡* L. /2 width=1/ qed. -lemma lprs_strap1: ∀L1,L,L2. L1 ⊢ ➡* L → L ⊢ ➡ L2 → L1 ⊢ ➡* L2. +lemma lprs_strap1: ∀L1,L,L2. L1 ⊢ ➡* L → ⦃G, L⦄ ⊢ ➡ L2 → L1 ⊢ ➡* L2. /2 width=3/ qed. -lemma lprs_strap2: ∀L1,L,L2. L1 ⊢ ➡ L → L ⊢ ➡* L2 → L1 ⊢ ➡* L2. +lemma lprs_strap2: ∀L1,L,L2. L1 ⊢ ➡ L → ⦃G, L⦄ ⊢ ➡* L2 → L1 ⊢ ➡* L2. /2 width=3/ qed. lemma lprs_pair_refl: ∀L1,L2. L1 ⊢ ➡* L2 → ∀I,V. L1. ⓑ{I} V ⊢ ➡* L2. ⓑ{I} V. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lprs_cprs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lprs_cprs.ma index cc52394c4..c3ba54ef2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lprs_cprs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lprs_cprs.ma @@ -71,8 +71,8 @@ lemma lprs_cpr_conf_sn: ∀L0,T0,T1. L0 ⊢ T0 ➡ T1 → ∀L1. L0 ⊢ ➡* L1 ∃∃T. L0 ⊢ T1 ➡* T & L1 ⊢ T0 ➡* T. /3 width=3 by lprs_cprs_conf_sn, cpr_cprs/ qed-. -lemma cprs_bind2: ∀L,V1,V2. L ⊢ V1 ➡* V2 → ∀I,T1,T2. L. ⓑ{I}V2 ⊢ T1 ➡* T2 → - ∀a. L ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. +lemma cprs_bind2: ∀L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 → ∀I,T1,T2. L. ⓑ{I}V2 ⊢ T1 ➡* T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2. #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 lapply (lprs_cprs_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1/ qed. @@ -80,8 +80,8 @@ qed. (* Inversion lemmas on context-sensitive parallel computation for terms *****) (* Basic_1: was: pr3_gen_abst *) -lemma cprs_inv_abst1: ∀a,L,W1,T1,U2. L ⊢ ⓛ{a}W1.T1 ➡* U2 → - ∃∃W2,T2. L ⊢ W1 ➡* W2 & L.ⓛW1 ⊢ T1 ➡* T2 & +lemma cprs_inv_abst1: ∀a,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* U2 → + ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & L.ⓛW1 ⊢ T1 ➡* T2 & U2 = ⓛ{a}W2.T2. #a #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5/ #U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct @@ -91,15 +91,15 @@ lapply (cprs_strap1 … HV10 … HV02) -V0 lapply (cprs_trans … HT10 … HT02) -T0 /2 width=5/ qed-. -lemma cprs_inv_abst: ∀a,L,W1,W2,T1,T2. L ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 → - L ⊢ W1 ➡* W2 ∧ L.ⓛW1 ⊢ T1 ➡* T2. +lemma cprs_inv_abst: ∀a,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 → + ⦃G, L⦄ ⊢ W1 ➡* W2 ∧ L.ⓛW1 ⊢ T1 ➡* T2. #a #L #W1 #W2 #T1 #T2 #H elim (cprs_inv_abst1 … H) -H #W #T #HW1 #HT1 #H destruct /2 width=1/ qed-. (* Basic_1: was pr3_gen_abbr *) -lemma cprs_inv_abbr1: ∀a,L,V1,T1,U2. L ⊢ ⓓ{a}V1.T1 ➡* U2 → ( - ∃∃V2,T2. L ⊢ V1 ➡* V2 & L. ⓓV1 ⊢ T1 ➡* T2 & +lemma cprs_inv_abbr1: ∀a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡* U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & L. ⓓV1 ⊢ T1 ➡* T2 & U2 = ⓓ{a}V2.T2 ) ∨ ∃∃T2. L. ⓓV1 ⊢ T1 ➡* T2 & ⇧[0, 1] U2 ≡ T2 & a = true. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs.ma index 4cf6a28b4..c2d0eb2cf 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs.ma @@ -26,48 +26,48 @@ interpretation "extended parallel computation (local environment, sn variant)" (* Basic eliminators ********************************************************) lemma lpxs_ind: ∀h,g,L1. ∀R:predicate lenv. R L1 → - (∀L,L2. ⦃h, L1⦄ ⊢ ➡*[g] L → ⦃h, L⦄ ⊢ ➡[g] L2 → R L → R L2) → - ∀L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → R L2. + (∀L,L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L → ⦃G, L⦄ ⊢ ➡[h, g] L2 → R L → R L2) → + ∀L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → R L2. #h #g #L1 #R #HL1 #IHL1 #L2 #HL12 @(TC_star_ind … HL1 IHL1 … HL12) // qed-. lemma lpxs_ind_dx: ∀h,g,L2. ∀R:predicate lenv. R L2 → - (∀L1,L. ⦃h, L1⦄ ⊢ ➡[g] L → ⦃h, L⦄ ⊢ ➡*[g] L2 → R L → R L1) → - ∀L1. ⦃h, L1⦄ ⊢ ➡*[g] L2 → R L1. + (∀L1,L. ⦃h, L1⦄ ⊢ ➡[h, g] L → ⦃G, L⦄ ⊢ ➡*[h, g] L2 → R L → R L1) → + ∀L1. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → R L1. #h #g #L2 #R #HL2 #IHL2 #L1 #HL12 @(TC_star_ind_dx … HL2 IHL2 … HL12) // qed-. (* Basic properties *********************************************************) -lemma lprs_lpxs: ∀h,g,L1,L2. L1 ⊢ ➡* L2 → ⦃h, L1⦄ ⊢ ➡*[g] L2. +lemma lprs_lpxs: ∀h,g,L1,L2. L1 ⊢ ➡* L2 → ⦃h, L1⦄ ⊢ ➡*[h, g] L2. /3 width=3/ qed. -lemma lpx_lpxs: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡[g] L2 → ⦃h, L1⦄ ⊢ ➡*[g] L2. +lemma lpx_lpxs: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡[h, g] L2 → ⦃h, L1⦄ ⊢ ➡*[h, g] L2. /2 width=1/ qed. -lemma lpxs_refl: ∀h,g,L. ⦃h, L⦄ ⊢ ➡*[g] L. +lemma lpxs_refl: ∀h,g,L. ⦃G, L⦄ ⊢ ➡*[h, g] L. /2 width=1/ qed. -lemma lpxs_strap1: ∀h,g,L1,L,L2. ⦃h, L1⦄ ⊢ ➡*[g] L → ⦃h, L⦄ ⊢ ➡[g] L2 → ⦃h, L1⦄ ⊢ ➡*[g] L2. +lemma lpxs_strap1: ∀h,g,L1,L,L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L → ⦃G, L⦄ ⊢ ➡[h, g] L2 → ⦃h, L1⦄ ⊢ ➡*[h, g] L2. /2 width=3/ qed. -lemma lpxs_strap2: ∀h,g,L1,L,L2. ⦃h, L1⦄ ⊢ ➡[g] L → ⦃h, L⦄ ⊢ ➡*[g] L2 → ⦃h, L1⦄ ⊢ ➡*[g] L2. +lemma lpxs_strap2: ∀h,g,L1,L,L2. ⦃h, L1⦄ ⊢ ➡[h, g] L → ⦃G, L⦄ ⊢ ➡*[h, g] L2 → ⦃h, L1⦄ ⊢ ➡*[h, g] L2. /2 width=3/ qed. -lemma lpxs_pair_refl: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → ∀I,V. ⦃h, L1. ⓑ{I}V⦄ ⊢ ➡*[g] L2. ⓑ{I}V. +lemma lpxs_pair_refl: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → ∀I,V. ⦃h, L1. ⓑ{I}V⦄ ⊢ ➡*[h, g] L2. ⓑ{I}V. /2 width=1 by TC_lpx_sn_pair_refl/ qed. (* Basic inversion lemmas ***************************************************) -lemma lpxs_inv_atom1: ∀h,g,L2. ⦃h, ⋆⦄ ⊢ ➡*[g] L2 → L2 = ⋆. +lemma lpxs_inv_atom1: ∀h,g,L2. ⦃h, ⋆⦄ ⊢ ➡*[h, g] L2 → L2 = ⋆. /2 width=2 by TC_lpx_sn_inv_atom1/ qed-. -lemma lpxs_inv_atom2: ∀h,g,L1. ⦃h, L1⦄ ⊢ ➡*[g] ⋆ → L1 = ⋆. +lemma lpxs_inv_atom2: ∀h,g,L1. ⦃h, L1⦄ ⊢ ➡*[h, g] ⋆ → L1 = ⋆. /2 width=2 by TC_lpx_sn_inv_atom2/ qed-. (* Basic forward lemmas *****************************************************) -lemma lpxs_fwd_length: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → |L1| = |L2|. +lemma lpxs_fwd_length: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → |L1| = |L2|. /2 width=2 by TC_lpx_sn_fwd_length/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_aaa.ma index b05385da5..e866102d4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_aaa.ma @@ -20,9 +20,9 @@ include "basic_2/computation/lpxs.ma". (* Properties about atomic arity assignment on terms ************************) lemma aaa_lpxs_conf: ∀h,g,L1,T,A. - L1 ⊢ T ⁝ A → ∀L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → L2 ⊢ T ⁝ A. + L1 ⊢ T ⁝ A → ∀L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → L2 ⊢ T ⁝ A. #h #g #L1 #T #A #HT #L2 #HL12 -@(TC_Conf3 … (λL,A. L ⊢ T ⁝ A) … HT ? HL12) /2 width=5 by aaa_lpx_conf/ +@(TC_Conf3 … (λL,A. ⦃G, L⦄ ⊢ T ⁝ A) … HT ? HL12) /2 width=5 by aaa_lpx_conf/ qed-. lemma aaa_lprs_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. L1 ⊢ ➡* L2 → L2 ⊢ T ⁝ A. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_alt.ma index 7a77b4b30..95a171de8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_alt.ma @@ -26,12 +26,12 @@ interpretation "extended parallel computation (local environment, sn variant) al (* Main properties on the alternative definition ****************************) -theorem lpxsa_lpxs: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡➡*[g] L2 → ⦃h, L1⦄ ⊢ ➡*[g] L2. +theorem lpxsa_lpxs: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡➡*[h, g] L2 → ⦃h, L1⦄ ⊢ ➡*[h, g] L2. /2 width=1 by lpx_sn_LTC_TC_lpx_sn/ qed-. (* Main inversion lemmas on the alternative definition **********************) -theorem lpxs_inv_lpxsa: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → ⦃h, L1⦄ ⊢ ➡➡*[g] L2. +theorem lpxs_inv_lpxsa: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → ⦃h, L1⦄ ⊢ ➡➡*[h, g] L2. /3 width=3 by TC_lpx_sn_inv_lpx_sn_LTC, lpx_cpxs_trans/ qed-. (* Alternative eliminators **************************************************) @@ -39,8 +39,8 @@ theorem lpxs_inv_lpxsa: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → ⦃h, L1⦄ lemma lpxs_ind_alt: ∀h,g. ∀R:relation lenv. R (⋆) (⋆) → ( ∀I,K1,K2,V1,V2. - ⦃h, K1⦄ ⊢ ➡*[g] K2 → ⦃h, K1⦄ ⊢ V1 ➡*[g] V2 → + ⦃h, K1⦄ ⊢ ➡*[h, g] K2 → ⦃h, K1⦄ ⊢ V1 ➡*[h, g] V2 → R K1 K2 → R (K1.ⓑ{I}V1) (K2.ⓑ{I}V2) ) → - ∀L1,L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → R L1 L2. + ∀L1,L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → R L1 L2. /3 width=4 by TC_lpx_sn_ind, lpx_cpxs_trans/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_cpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_cpxs.ma index 960d6778b..381e6e939 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_cpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_cpxs.ma @@ -19,18 +19,18 @@ include "basic_2/computation/lpxs.ma". (* Advanced properties ******************************************************) -lemma lpxs_pair: ∀h,g,I,L1,L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → ∀V1,V2. ⦃h, L1⦄ ⊢ V1 ➡*[g] V2 → - ⦃h, L1.ⓑ{I}V1⦄ ⊢ ➡*[g] L2.ⓑ{I}V2. +lemma lpxs_pair: ∀h,g,I,L1,L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → ∀V1,V2. ⦃h, L1⦄ ⊢ V1 ➡*[h, g] V2 → + ⦃h, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, g] L2.ⓑ{I}V2. /2 width=1 by TC_lpx_sn_pair/ qed. (* Advanced inversion lemmas ************************************************) -lemma lpxs_inv_pair1: ∀h,g,I,K1,L2,V1. ⦃h, K1.ⓑ{I}V1⦄ ⊢ ➡*[g] L2 → - ∃∃K2,V2. ⦃h, K1⦄ ⊢ ➡*[g] K2 & ⦃h, K1⦄ ⊢ V1 ➡*[g] V2 & L2 = K2.ⓑ{I}V2. +lemma lpxs_inv_pair1: ∀h,g,I,K1,L2,V1. ⦃h, K1.ⓑ{I}V1⦄ ⊢ ➡*[h, g] L2 → + ∃∃K2,V2. ⦃h, K1⦄ ⊢ ➡*[h, g] K2 & ⦃h, K1⦄ ⊢ V1 ➡*[h, g] V2 & L2 = K2.ⓑ{I}V2. /3 width=3 by TC_lpx_sn_inv_pair1, lpx_cpxs_trans/ qed-. -lemma lpxs_inv_pair2: ∀h,g,I,L1,K2,V2. ⦃h, L1⦄ ⊢ ➡*[g] K2.ⓑ{I}V2 → - ∃∃K1,V1. ⦃h, K1⦄ ⊢ ➡*[g] K2 & ⦃h, K1⦄ ⊢ V1 ➡*[g] V2 & L1 = K1.ⓑ{I}V1. +lemma lpxs_inv_pair2: ∀h,g,I,L1,K2,V2. ⦃h, L1⦄ ⊢ ➡*[h, g] K2.ⓑ{I}V2 → + ∃∃K1,V1. ⦃h, K1⦄ ⊢ ➡*[h, g] K2 & ⦃h, K1⦄ ⊢ V1 ➡*[h, g] V2 & L1 = K1.ⓑ{I}V1. /3 width=3 by TC_lpx_sn_inv_pair2, lpx_cpxs_trans/ qed-. (* Properties on context-sensitive extended parallel computation for terms **) @@ -41,17 +41,17 @@ lemma lpxs_cpx_trans: ∀h,g. s_r_trans … (cpx h g) (lpxs h g). lemma lpxs_cpxs_trans: ∀h,g. s_rs_trans … (cpx h g) (lpxs h g). /3 width=5 by s_r_trans_TC1, lpxs_cpx_trans/ qed-. -lemma cpxs_bind2: ∀h,g,L,V1,V2. ⦃h, L⦄ ⊢ V1 ➡*[g] V2 → - ∀I,T1,T2. ⦃h, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[g] T2 → - ∀a. ⦃h, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[g] ⓑ{a,I}V2.T2. +lemma cpxs_bind2: ∀h,g,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 → + ∀I,T1,T2. ⦃h, L.ⓑ{I}V2⦄ ⊢ T1 ➡*[h, g] T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡*[h, g] ⓑ{a,I}V2.T2. #h #g #L #V1 #V2 #HV12 #I #T1 #T2 #HT12 lapply (lpxs_cpxs_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1/ qed. (* Inversion lemmas on context-sensitive ext parallel computation for terms *) -lemma cpxs_inv_abst1: ∀h,g,a,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓛ{a}V1.T1 ➡*[g] U2 → - ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡*[g] V2 & ⦃h, L.ⓛV1⦄ ⊢ T1 ➡*[g] T2 & +lemma cpxs_inv_abst1: ∀h,g,a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡*[h, g] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 & ⦃h, L.ⓛV1⦄ ⊢ T1 ➡*[h, g] T2 & U2 = ⓛ{a}V2.T2. #h #g #a #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /2 width=5/ #U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct @@ -61,11 +61,11 @@ lapply (cpxs_strap1 … HV10 … HV02) -V0 lapply (cpxs_trans … HT10 … HT02) -T0 /2 width=5/ qed-. -lemma cpxs_inv_abbr1: ∀h,g,a,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓓ{a}V1.T1 ➡*[g] U2 → ( - ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡*[g] V2 & ⦃h, L.ⓓV1⦄ ⊢ T1 ➡*[g] T2 & +lemma cpxs_inv_abbr1: ∀h,g,a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡*[h, g] U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡*[h, g] V2 & ⦃h, L.ⓓV1⦄ ⊢ T1 ➡*[h, g] T2 & U2 = ⓓ{a}V2.T2 ) ∨ - ∃∃T2. ⦃h, L.ⓓV1⦄ ⊢ T1 ➡*[g] T2 & ⇧[0, 1] U2 ≡ T2 & a = true. + ∃∃T2. ⦃h, L.ⓓV1⦄ ⊢ T1 ➡*[h, g] T2 & ⇧[0, 1] U2 ≡ T2 & a = true. #h #g #a #L #V1 #T1 #U2 #H @(cpxs_ind … H) -U2 /3 width=5/ #U0 #U2 #_ #HU02 * * [ #V0 #T0 #HV10 #HT10 #H destruct @@ -86,6 +86,6 @@ qed-. (* More advanced properties *************************************************) -lemma lpxs_pair2: ∀h,g,I,L1,L2. ⦃h, L1⦄ ⊢ ➡*[g] L2 → - ∀V1,V2. ⦃h, L2⦄ ⊢ V1 ➡*[g] V2 → ⦃h, L1.ⓑ{I}V1⦄ ⊢ ➡*[g] L2.ⓑ{I}V2. +lemma lpxs_pair2: ∀h,g,I,L1,L2. ⦃h, L1⦄ ⊢ ➡*[h, g] L2 → + ∀V1,V2. ⦃h, L2⦄ ⊢ V1 ➡*[h, g] V2 → ⦃h, L1.ⓑ{I}V1⦄ ⊢ ➡*[h, g] L2.ⓑ{I}V2. /3 width=3 by lpxs_pair, lpxs_cpxs_trans/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/conversion/cpc.ma b/matita/matita/contribs/lambdadelta/basic_2/conversion/cpc.ma index 9b78b55c9..5be5601ea 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/conversion/cpc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/conversion/cpc.ma @@ -18,7 +18,7 @@ include "basic_2/reduction/cpr.ma". (* CONTEXT-SENSITIVE PARALLEL CONVERSION ON TERMS ***************************) definition cpc: lenv → relation term ≝ - λL,T1,T2. L ⊢ T1 ➡ T2 ∨ L ⊢ T2 ➡ T1. + λL,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 ∨ ⦃G, L⦄ ⊢ T2 ➡ T1. interpretation "context-sensitive parallel conversion (term)" @@ -35,6 +35,6 @@ qed. (* Basic forward lemmas *****************************************************) -lemma cpc_fwd_cpr: ∀L,T1,T2. L ⊢ T1 ⬌ T2 → ∃∃T. L ⊢ T1 ➡ T & L ⊢ T2 ➡ T. +lemma cpc_fwd_cpr: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬌ T2 → ∃∃T. ⦃G, L⦄ ⊢ T1 ➡ T & ⦃G, L⦄ ⊢ T2 ➡ T. #L #T1 #T2 * /2 width=3/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/conversion/cpc_cpc.ma b/matita/matita/contribs/lambdadelta/basic_2/conversion/cpc_cpc.ma index dcea07a8b..092d2a9a7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/conversion/cpc_cpc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/conversion/cpc_cpc.ma @@ -18,6 +18,6 @@ include "basic_2/conversion/cpc.ma". (* Main properties **********************************************************) -theorem cpc_conf: ∀L,T0,T1,T2. L ⊢ T0 ⬌ T1 → L ⊢ T0 ⬌ T2 → - ∃∃T. L ⊢ T1 ⬌ T & L ⊢ T2 ⬌ T. +theorem cpc_conf: ∀L,T0,T1,T2. ⦃G, L⦄ ⊢ T0 ⬌ T1 → ⦃G, L⦄ ⊢ T0 ⬌ T2 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ⬌ T & ⦃G, L⦄ ⊢ T2 ⬌ T. /3 width=3/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv.ma index 8d4e9b67a..29d6151fe 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv.ma @@ -22,8 +22,8 @@ inductive lsubsv (h:sh) (g:sd h): relation lenv ≝ | lsubsv_atom: lsubsv h g (⋆) (⋆) | lsubsv_pair: ∀I,L1,L2,V. lsubsv h g L1 L2 → lsubsv h g (L1.ⓑ{I}V) (L2.ⓑ{I}V) -| lsubsv_abbr: ∀L1,L2,W,V,W1,V2,l. ⦃h, L1⦄ ⊢ ⓝW.V ¡[g] → ⦃h, L2⦄ ⊢ W ¡[g] → - ⦃h, L1⦄ ⊢ V •[g] ⦃l+1, W1⦄ → ⦃h, L2⦄ ⊢ W •[g] ⦃l, V2⦄ → +| lsubsv_abbr: ∀L1,L2,W,V,W1,V2,l. ⦃h, L1⦄ ⊢ ⓝW.V ¡[h, g] → ⦃h, L2⦄ ⊢ W ¡[h, g] → + ⦃h, L1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ → ⦃h, L2⦄ ⊢ W •[h, g] ⦃l, V2⦄ → lsubsv h g L1 L2 → lsubsv h g (L1.ⓓⓝW.V) (L2.ⓛW) . @@ -33,7 +33,7 @@ interpretation (* Basic inversion lemmas ***************************************************) -fact lsubsv_inv_atom1_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L1 = ⋆ → L2 = ⋆. +fact lsubsv_inv_atom1_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → L1 = ⋆ → L2 = ⋆. #h #g #L1 #L2 * -L1 -L2 [ // | #I #L1 #L2 #V #_ #H destruct @@ -41,15 +41,15 @@ fact lsubsv_inv_atom1_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L1 = ⋆ → L ] qed-. -lemma lsubsv_inv_atom1: ∀h,g,L2. h ⊢ ⋆ ¡⊑[g] L2 → L2 = ⋆. +lemma lsubsv_inv_atom1: ∀h,g,L2. h ⊢ ⋆ ¡⊑[h, g] L2 → L2 = ⋆. /2 width=5 by lsubsv_inv_atom1_aux/ qed-. -fact lsubsv_inv_pair1_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → +fact lsubsv_inv_pair1_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X → - (∃∃K2. h ⊢ K1 ¡⊑[g] K2 & L2 = K2.ⓑ{I}X) ∨ - ∃∃K2,W,V,W1,V2,l. ⦃h, K1⦄ ⊢ X ¡[g] & ⦃h, K2⦄ ⊢ W ¡[g] & - ⦃h, K1⦄ ⊢ V •[g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[g] ⦃l, V2⦄ & - h ⊢ K1 ¡⊑[g] K2 & + (∃∃K2. h ⊢ K1 ¡⊑[h, g] K2 & L2 = K2.ⓑ{I}X) ∨ + ∃∃K2,W,V,W1,V2,l. ⦃h, K1⦄ ⊢ X ¡[h, g] & ⦃h, K2⦄ ⊢ W ¡[h, g] & + ⦃h, K1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[h, g] ⦃l, V2⦄ & + h ⊢ K1 ¡⊑[h, g] K2 & I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. #h #g #L1 #L2 * -L1 -L2 [ #J #K1 #X #H destruct @@ -58,15 +58,15 @@ fact lsubsv_inv_pair1_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → ] qed-. -lemma lsubsv_inv_pair1: ∀h,g,I,K1,L2,X. h ⊢ K1.ⓑ{I}X ¡⊑[g] L2 → - (∃∃K2. h ⊢ K1 ¡⊑[g] K2 & L2 = K2.ⓑ{I}X) ∨ - ∃∃K2,W,V,W1,V2,l. ⦃h, K1⦄ ⊢ X ¡[g] & ⦃h, K2⦄ ⊢ W ¡[g] & - ⦃h, K1⦄ ⊢ V •[g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[g] ⦃l, V2⦄ & - h ⊢ K1 ¡⊑[g] K2 & +lemma lsubsv_inv_pair1: ∀h,g,I,K1,L2,X. h ⊢ K1.ⓑ{I}X ¡⊑[h, g] L2 → + (∃∃K2. h ⊢ K1 ¡⊑[h, g] K2 & L2 = K2.ⓑ{I}X) ∨ + ∃∃K2,W,V,W1,V2,l. ⦃h, K1⦄ ⊢ X ¡[h, g] & ⦃h, K2⦄ ⊢ W ¡[h, g] & + ⦃h, K1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[h, g] ⦃l, V2⦄ & + h ⊢ K1 ¡⊑[h, g] K2 & I = Abbr & L2 = K2.ⓛW & X = ⓝW.V. /2 width=3 by lsubsv_inv_pair1_aux/ qed-. -fact lsubsv_inv_atom2_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L2 = ⋆ → L1 = ⋆. +fact lsubsv_inv_atom2_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → L2 = ⋆ → L1 = ⋆. #h #g #L1 #L2 * -L1 -L2 [ // | #I #L1 #L2 #V #_ #H destruct @@ -74,15 +74,15 @@ fact lsubsv_inv_atom2_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L2 = ⋆ → L ] qed-. -lemma lsubsv_inv_atom2: ∀h,g,L1. h ⊢ L1 ¡⊑[g] ⋆ → L1 = ⋆. +lemma lsubsv_inv_atom2: ∀h,g,L1. h ⊢ L1 ¡⊑[h, g] ⋆ → L1 = ⋆. /2 width=5 by lsubsv_inv_atom2_aux/ qed-. -fact lsubsv_inv_pair2_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → +fact lsubsv_inv_pair2_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → ∀I,K2,W. L2 = K2.ⓑ{I}W → - (∃∃K1. h ⊢ K1 ¡⊑[g] K2 & L1 = K1.ⓑ{I}W) ∨ - ∃∃K1,V,W1,V2,l. ⦃h, K1⦄ ⊢ ⓝW.V ¡[g] & ⦃h, K2⦄ ⊢ W ¡[g] & - ⦃h, K1⦄ ⊢ V •[g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[g] ⦃l, V2⦄ & - h ⊢ K1 ¡⊑[g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. + (∃∃K1. h ⊢ K1 ¡⊑[h, g] K2 & L1 = K1.ⓑ{I}W) ∨ + ∃∃K1,V,W1,V2,l. ⦃h, K1⦄ ⊢ ⓝW.V ¡[h, g] & ⦃h, K2⦄ ⊢ W ¡[h, g] & + ⦃h, K1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[h, g] ⦃l, V2⦄ & + h ⊢ K1 ¡⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. #h #g #L1 #L2 * -L1 -L2 [ #J #K2 #U #H destruct | #I #L1 #L2 #V #HL12 #J #K2 #U #H destruct /3 width=3/ @@ -90,26 +90,26 @@ fact lsubsv_inv_pair2_aux: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → ] qed-. -lemma lsubsv_inv_pair2: ∀h,g,I,L1,K2,W. h ⊢ L1 ¡⊑[g] K2.ⓑ{I}W → - (∃∃K1. h ⊢ K1 ¡⊑[g] K2 & L1 = K1.ⓑ{I}W) ∨ - ∃∃K1,V,W1,V2,l. ⦃h, K1⦄ ⊢ ⓝW.V ¡[g] & ⦃h, K2⦄ ⊢ W ¡[g] & - ⦃h, K1⦄ ⊢ V •[g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[g] ⦃l, V2⦄ & - h ⊢ K1 ¡⊑[g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. +lemma lsubsv_inv_pair2: ∀h,g,I,L1,K2,W. h ⊢ L1 ¡⊑[h, g] K2.ⓑ{I}W → + (∃∃K1. h ⊢ K1 ¡⊑[h, g] K2 & L1 = K1.ⓑ{I}W) ∨ + ∃∃K1,V,W1,V2,l. ⦃h, K1⦄ ⊢ ⓝW.V ¡[h, g] & ⦃h, K2⦄ ⊢ W ¡[h, g] & + ⦃h, K1⦄ ⊢ V •[h, g] ⦃l+1, W1⦄ & ⦃h, K2⦄ ⊢ W •[h, g] ⦃l, V2⦄ & + h ⊢ K1 ¡⊑[h, g] K2 & I = Abst & L1 = K1. ⓓⓝW.V. /2 width=3 by lsubsv_inv_pair2_aux/ qed-. (* Basic_forward lemmas *****************************************************) -lemma lsubsv_fwd_lsubr: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L1 ⊑ L2. +lemma lsubsv_fwd_lsubr: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → L1 ⊑ L2. #h #g #L1 #L2 #H elim H -L1 -L2 // /2 width=1/ qed-. (* Basic properties *********************************************************) -lemma lsubsv_refl: ∀h,g,L. h ⊢ L ¡⊑[g] L. +lemma lsubsv_refl: ∀h,g,L. h ⊢ L ¡⊑[h, g] L. #h #g #L elim L -L // /2 width=1/ qed. -lemma lsubsv_cprs_trans: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → +lemma lsubsv_cprs_trans: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2. /3 width=5 by lsubsv_fwd_lsubr, lsubr_cprs_trans/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpcs.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpcs.ma index a66322d66..30e42ea09 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpcs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpcs.ma @@ -19,7 +19,7 @@ include "basic_2/dynamic/lsubsv.ma". (* Properties on context-sensitive parallel equivalence for terms ***********) -lemma lsubsv_cpcs_trans: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → +lemma lsubsv_cpcs_trans: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → ∀T1,T2. L2 ⊢ T1 ⬌* T2 → L1 ⊢ T1 ⬌* T2. /3 width=5 by lsubsv_fwd_lsubr, lsubr_cpcs_trans/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpds.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpds.ma index 4ffc91f95..55a7e9a13 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpds.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_cpds.ma @@ -20,13 +20,13 @@ include "basic_2/dynamic/lsubsv_ssta.ma". (* Properties for the preservation results **********************************) fact lsubsv_sstas_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_lsubsv h g L1 T1) → - ∀L2,T. h ⊢ ⦃L0, T0⦄ >[g] ⦃L2, T⦄ → ⦃h, L2⦄ ⊢ T ¡[g] → - ∀L1. h ⊢ L1 ¡⊑[g] L2 → ∀U2. ⦃h, L2⦄ ⊢ T •*[g] U2 → - ∃∃U1. ⦃h, L1⦄ ⊢ T •*[g] U1 & L1 ⊢ U1 ⬌* U2. + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_lsubsv h g L1 T1) → + ∀L2,T. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L2, T⦄ → ⦃h, L2⦄ ⊢ T ¡[h, g] → + ∀L1. h ⊢ L1 ¡⊑[h, g] L2 → ∀U2. ⦃h, L2⦄ ⊢ T •*[h, g] U2 → + ∃∃U1. ⦃h, L1⦄ ⊢ T •*[h, g] U1 & L1 ⊢ U1 ⬌* U2. #h #g #L0 #T0 #IH4 #IH3 #IH2 #IH1 #L2 #T #HLT0 #HT #L1 #HL12 #U2 #H @(sstas_ind … H) -U2 [ /2 width=3/ ] #U2 #W #l #HTU2 #HU2W * #U1 #HTU1 #HU12 lapply (IH1 … HT … HL12) // #H @@ -41,13 +41,13 @@ lapply (cpcs_trans … HW12 … HW2) -W2 /3 width=4/ qed-. fact lsubsv_cpds_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_lsubsv h g L1 T1) → - ∀L2,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L2, T1⦄ → ⦃h, L2⦄ ⊢ T1 ¡[g] → - ∀L1. h ⊢ L1 ¡⊑[g] L2 → ∀T2. ⦃h, L2⦄ ⊢ T1 •*➡*[g] T2 → - ∃∃T. ⦃h, L1⦄ ⊢ T1 •*➡*[g] T & L1 ⊢ T2 ➡* T. + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_lsubsv h g L1 T1) → + ∀L2,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L2, T1⦄ → ⦃h, L2⦄ ⊢ T1 ¡[h, g] → + ∀L1. h ⊢ L1 ¡⊑[h, g] L2 → ∀T2. ⦃h, L2⦄ ⊢ T1 •*➡*[h, g] T2 → + ∃∃T. ⦃h, L1⦄ ⊢ T1 •*➡*[h, g] T & L1 ⊢ T2 ➡* T. #h #g #L0 #T0 #IH4 #IH3 #IH2 #IH1 #L2 #T1 #HLT0 #HT1 #L1 #HL12 #T2 * #T #HT1T #HTT2 lapply (lsubsv_cprs_trans … HL12 … HTT2) -HTT2 #HTT2 elim (lsubsv_sstas_aux … IH4 IH3 IH2 IH1 … HLT0 … HL12 … HT1T) // -L2 -L0 -T0 #T0 #HT10 #HT0 diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ldrop.ma index f8caa3aee..31d47c8c5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ldrop.ma @@ -19,9 +19,9 @@ include "basic_2/dynamic/lsubsv.ma". (* Properties concerning basic local environment slicing ********************) (* Note: the constant 0 cannot be generalized *) -lemma lsubsv_ldrop_O1_conf: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → +lemma lsubsv_ldrop_O1_conf: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → ∀K1,e. ⇩[0, e] L1 ≡ K1 → - ∃∃K2. h ⊢ K1 ¡⊑[g] K2 & ⇩[0, e] L2 ≡ K2. + ∃∃K2. h ⊢ K1 ¡⊑[h, g] K2 & ⇩[0, e] L2 ≡ K2. #h #g #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #K1 #e #H @@ -42,9 +42,9 @@ lemma lsubsv_ldrop_O1_conf: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → qed-. (* Note: the constant 0 cannot be generalized *) -lemma lsubsv_ldrop_O1_trans: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → +lemma lsubsv_ldrop_O1_trans: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 → - ∃∃K1. h ⊢ K1 ¡⊑[g] K2 & ⇩[0, e] L1 ≡ K1. + ∃∃K1. h ⊢ K1 ¡⊑[h, g] K2 & ⇩[0, e] L1 ≡ K1. #h #g #L1 #L2 #H elim H -L1 -L2 [ /2 width=3/ | #I #L1 #L2 #V #_ #IHL12 #K2 #e #H diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_lsuba.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_lsuba.ma index 100c34da2..7d40d4081 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_lsuba.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_lsuba.ma @@ -19,7 +19,7 @@ include "basic_2/dynamic/lsubsv.ma". (* Forward lemmas on lenv refinement for atomic arity assignment ************) -lemma lsubsv_fwd_lsuba: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[g] L2 → L1 ⁝⊑ L2. +lemma lsubsv_fwd_lsuba: ∀h,g,L1,L2. h ⊢ L1 ¡⊑[h, g] L2 → L1 ⁝⊑ L2. #h #g #L1 #L2 #H elim H -L1 -L2 // /2 width=1/ #L1 #L2 #W #V #W1 #V2 #l #HV #HW #_ #_ #_ #IHL12 -W1 -V2 elim (snv_fwd_aaa … HV) -HV #A #HV 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 078b8fa89..67230d980 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_snv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_snv.ma @@ -21,10 +21,10 @@ include "basic_2/dynamic/lsubsv_cpcs.ma". (* Properties concerning stratified native validity *************************) fact snv_lsubsv_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_lsubsv h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_lsubsv h g L1 T1) → ∀L1,T1. L0 = L1 → T0 = T1 → IH_snv_lsubsv h g L1 T1. #h #g #L0 #T0 #IH4 #IH3 #IH2 #IH1 #L2 * * [||||*] // [ #i #HL0 #HT0 #H #L1 #HL12 destruct -IH4 -IH3 -IH2 diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ssta.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ssta.ma index b372dca9c..15f2e14b7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ssta.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ssta.ma @@ -20,9 +20,9 @@ include "basic_2/dynamic/lsubsv_ldrop.ma". (* Properties on stratified static type assignment **************************) -lemma lsubsv_ssta_trans: ∀h,g,L2,T,U2,l. ⦃h, L2⦄ ⊢ T •[g] ⦃l, U2⦄ → - ∀L1. h ⊢ L1 ¡⊑[g] L2 → - ∃∃U1. ⦃h, L1⦄ ⊢ T •[g] ⦃l, U1⦄ & L1 ⊢ U1 ⬌* U2. +lemma lsubsv_ssta_trans: ∀h,g,L2,T,U2,l. ⦃h, L2⦄ ⊢ T •[h, g] ⦃l, U2⦄ → + ∀L1. h ⊢ L1 ¡⊑[h, g] L2 → + ∃∃U1. ⦃h, L1⦄ ⊢ T •[h, g] ⦃l, U1⦄ & L1 ⊢ U1 ⬌* U2. #h #g #L2 #T #U #l #H elim H -L2 -T -U -l [ /3 width=3/ | #L2 #K2 #X #Y #U #i #l #HLK2 #_ #HYU #IHXY #L1 #HL12 diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv.ma index d4e3beda1..3f62cff77 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv.ma @@ -23,10 +23,10 @@ inductive snv (h:sh) (g:sd h): lenv → predicate term ≝ | snv_lref: ∀I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → snv h g K V → snv h g L (#i) | snv_bind: ∀a,I,L,V,T. snv h g L V → snv h g (L.ⓑ{I}V) T → snv h g L (ⓑ{a,I}V.T) | snv_appl: ∀a,L,V,W,W0,T,U,l. snv h g L V → snv h g L T → - ⦃h, L⦄ ⊢ V •[g] ⦃l+1, W⦄ → L ⊢ W ➡* W0 → - ⦃h, L⦄ ⊢ T •*➡*[g] ⓛ{a}W0.U → snv h g L (ⓐV.T) + ⦃G, L⦄ ⊢ V •[h, g] ⦃l+1, W⦄ → ⦃G, L⦄ ⊢ W ➡* W0 → + ⦃G, L⦄ ⊢ T •*➡*[h, g] ⓛ{a}W0.U → snv h g L (ⓐV.T) | snv_cast: ∀L,W,T,U,l. snv h g L W → snv h g L T → - ⦃h, L⦄ ⊢ T •[g] ⦃l+1, U⦄ → L ⊢ U ⬌* W → snv h g L (ⓝW.T) + ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄ → ⦃G, L⦄ ⊢ U ⬌* W → snv h g L (ⓝW.T) . interpretation "stratified native validity (term)" @@ -34,8 +34,8 @@ interpretation "stratified native validity (term)" (* Basic inversion lemmas ***************************************************) -fact snv_inv_lref_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀i. X = #i → - ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V & ⦃h, K⦄ ⊢ V ¡[g]. +fact snv_inv_lref_aux: ∀h,g,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀i. X = #i → + ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V & ⦃h, K⦄ ⊢ V ¡[h, g]. #h #g #L #X * -L -X [ #L #k #i #H destruct | #I #L #K #V #i0 #HLK #HV #i #H destruct /2 width=5/ @@ -45,11 +45,11 @@ fact snv_inv_lref_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀i. X = #i → ] qed. -lemma snv_inv_lref: ∀h,g,L,i. ⦃h, L⦄ ⊢ #i ¡[g] → - ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V & ⦃h, K⦄ ⊢ V ¡[g]. +lemma snv_inv_lref: ∀h,g,L,i. ⦃G, L⦄ ⊢ #i ¡[h, g] → + ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V & ⦃h, K⦄ ⊢ V ¡[h, g]. /2 width=3/ qed-. -fact snv_inv_gref_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀p. X = §p → ⊥. +fact snv_inv_gref_aux: ∀h,g,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀p. X = §p → ⊥. #h #g #L #X * -L -X [ #L #k #p #H destruct | #I #L #K #V #i #_ #_ #p #H destruct @@ -59,11 +59,11 @@ fact snv_inv_gref_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀p. X = §p → ] qed. -lemma snv_inv_gref: ∀h,g,L,p. ⦃h, L⦄ ⊢ §p ¡[g] → ⊥. +lemma snv_inv_gref: ∀h,g,L,p. ⦃G, L⦄ ⊢ §p ¡[h, g] → ⊥. /2 width=7/ qed-. -fact snv_inv_bind_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀a,I,V,T. X = ⓑ{a,I}V.T → - ⦃h, L⦄ ⊢ V ¡[g] ∧ ⦃h, L.ⓑ{I}V⦄ ⊢ T ¡[g]. +fact snv_inv_bind_aux: ∀h,g,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀a,I,V,T. X = ⓑ{a,I}V.T → + ⦃G, L⦄ ⊢ V ¡[h, g] ∧ ⦃h, L.ⓑ{I}V⦄ ⊢ T ¡[h, g]. #h #g #L #X * -L -X [ #L #k #a #I #V #T #H destruct | #I0 #L #K #V0 #i #_ #_ #a #I #V #T #H destruct @@ -73,14 +73,14 @@ fact snv_inv_bind_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀a,I,V,T. X = ] qed. -lemma snv_inv_bind: ∀h,g,a,I,L,V,T. ⦃h, L⦄ ⊢ ⓑ{a,I}V.T ¡[g] → - ⦃h, L⦄ ⊢ V ¡[g] ∧ ⦃h, L.ⓑ{I}V⦄ ⊢ T ¡[g]. +lemma snv_inv_bind: ∀h,g,a,I,L,V,T. ⦃G, L⦄ ⊢ ⓑ{a,I}V.T ¡[h, g] → + ⦃G, L⦄ ⊢ V ¡[h, g] ∧ ⦃h, L.ⓑ{I}V⦄ ⊢ T ¡[h, g]. /2 width=4/ qed-. -fact snv_inv_appl_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀V,T. X = ⓐV.T → - ∃∃a,W,W0,U,l. ⦃h, L⦄ ⊢ V ¡[g] & ⦃h, L⦄ ⊢ T ¡[g] & - ⦃h, L⦄ ⊢ V •[g] ⦃l+1, W⦄ & L ⊢ W ➡* W0 & - ⦃h, L⦄ ⊢ T •*➡*[g] ⓛ{a}W0.U. +fact snv_inv_appl_aux: ∀h,g,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀V,T. X = ⓐV.T → + ∃∃a,W,W0,U,l. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] & + ⦃G, L⦄ ⊢ V •[h, g] ⦃l+1, W⦄ & ⦃G, L⦄ ⊢ W ➡* W0 & + ⦃G, L⦄ ⊢ T •*➡*[h, g] ⓛ{a}W0.U. #h #g #L #X * -L -X [ #L #k #V #T #H destruct | #I #L #K #V0 #i #_ #_ #V #T #H destruct @@ -90,15 +90,15 @@ fact snv_inv_appl_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀V,T. X = ⓐV.T ] qed. -lemma snv_inv_appl: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ ⓐV.T ¡[g] → - ∃∃a,W,W0,U,l. ⦃h, L⦄ ⊢ V ¡[g] & ⦃h, L⦄ ⊢ T ¡[g] & - ⦃h, L⦄ ⊢ V •[g] ⦃l+1, W⦄ & L ⊢ W ➡* W0 & - ⦃h, L⦄ ⊢ T •*➡*[g] ⓛ{a}W0.U. +lemma snv_inv_appl: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ ⓐV.T ¡[h, g] → + ∃∃a,W,W0,U,l. ⦃G, L⦄ ⊢ V ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] & + ⦃G, L⦄ ⊢ V •[h, g] ⦃l+1, W⦄ & ⦃G, L⦄ ⊢ W ➡* W0 & + ⦃G, L⦄ ⊢ T •*➡*[h, g] ⓛ{a}W0.U. /2 width=3/ qed-. -fact snv_inv_cast_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀W,T. X = ⓝW.T → - ∃∃U,l. ⦃h, L⦄ ⊢ W ¡[g] & ⦃h, L⦄ ⊢ T ¡[g] & - ⦃h, L⦄ ⊢ T •[g] ⦃l+1, U⦄ & L ⊢ U ⬌* W. +fact snv_inv_cast_aux: ∀h,g,L,X. ⦃G, L⦄ ⊢ X ¡[h, g] → ∀W,T. X = ⓝW.T → + ∃∃U,l. ⦃G, L⦄ ⊢ W ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] & + ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄ & ⦃G, L⦄ ⊢ U ⬌* W. #h #g #L #X * -L -X [ #L #k #W #T #H destruct | #I #L #K #V #i #_ #_ #W #T #H destruct @@ -108,14 +108,14 @@ fact snv_inv_cast_aux: ∀h,g,L,X. ⦃h, L⦄ ⊢ X ¡[g] → ∀W,T. X = ⓝW.T ] qed. -lemma snv_inv_cast: ∀h,g,L,W,T. ⦃h, L⦄ ⊢ ⓝW.T ¡[g] → - ∃∃U,l. ⦃h, L⦄ ⊢ W ¡[g] & ⦃h, L⦄ ⊢ T ¡[g] & - ⦃h, L⦄ ⊢ T •[g] ⦃l+1, U⦄ & L ⊢ U ⬌* W. +lemma snv_inv_cast: ∀h,g,L,W,T. ⦃G, L⦄ ⊢ ⓝW.T ¡[h, g] → + ∃∃U,l. ⦃G, L⦄ ⊢ W ¡[h, g] & ⦃G, L⦄ ⊢ T ¡[h, g] & + ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄ & ⦃G, L⦄ ⊢ U ⬌* W. /2 width=3/ qed-. (* Basic forward lemmas *****************************************************) -lemma snv_fwd_ssta: ∀h,g,L,T. ⦃h, L⦄ ⊢ T ¡[g] → ∃∃l,U. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄. +lemma snv_fwd_ssta: ∀h,g,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃∃l,U. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄. #h #g #L #T #H elim H -L -T [ #L #k elim (deg_total h g k) /3 width=3/ | * #L #K #V #i #HLK #_ * #l0 #W #HVW diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_aaa.ma index 32807a499..3a161e1f9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_aaa.ma @@ -21,7 +21,7 @@ include "basic_2/dynamic/snv.ma". (* Forward lemmas on atomic arity assignment for terms **********************) -lemma snv_fwd_aaa: ∀h,g,L,T. ⦃h, L⦄ ⊢ T ¡[g] → ∃A. L ⊢ T ⁝ A. +lemma snv_fwd_aaa: ∀h,g,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ∃A. ⦃G, L⦄ ⊢ T ⁝ A. #h #g #L #T #H elim H -L -T [ /2 width=2/ | #I #L #K #V #i #HLK #_ * /3 width=6/ @@ -36,6 +36,6 @@ lemma snv_fwd_aaa: ∀h,g,L,T. ⦃h, L⦄ ⊢ T ¡[g] → ∃A. L ⊢ T ⁝ A. ] qed-. -lemma snv_fwd_csn: ∀h,g,L,T. ⦃h, L⦄ ⊢ T ¡[g] → ⦃h, L⦄ ⊢ ⬊*[g] T. +lemma snv_fwd_csn: ∀h,g,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → ⦃G, L⦄ ⊢ ⬊*[h, g] T. #h #g #L #T #H elim (snv_fwd_aaa … H) -H /2 width=2/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_cpcs.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_cpcs.ma index 707994c62..8aaf788a2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_cpcs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_cpcs.ma @@ -22,41 +22,41 @@ include "basic_2/dynamic/ygt.ma". (* Inductive premises for the preservation results **************************) definition IH_snv_cpr_lpr: ∀h:sh. sd h → relation2 lenv term ≝ - λh,g,L1,T1. ⦃h, L1⦄ ⊢ T1 ¡[g] → - ∀T2. L1 ⊢ T1 ➡ T2 → ∀L2. L1 ⊢ ➡ L2 → ⦃h, L2⦄ ⊢ T2 ¡[g]. + λh,g,L1,T1. ⦃h, L1⦄ ⊢ T1 ¡[h, g] → + ∀T2. L1 ⊢ T1 ➡ T2 → ∀L2. L1 ⊢ ➡ L2 → ⦃h, L2⦄ ⊢ T2 ¡[h, g]. definition IH_ssta_cpr_lpr: ∀h:sh. sd h → relation2 lenv term ≝ - λh,g,L1,T1. ⦃h, L1⦄ ⊢ T1 ¡[g] → - ∀U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ → + λh,g,L1,T1. ⦃h, L1⦄ ⊢ T1 ¡[h, g] → + ∀U1,l. ⦃h, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ → ∀T2. L1 ⊢ T1 ➡ T2 → ∀L2. L1 ⊢ ➡ L2 → - ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄ & L2 ⊢ U1 ⬌* U2. + ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄ & L2 ⊢ U1 ⬌* U2. definition IH_snv_ssta: ∀h:sh. sd h → relation2 lenv term ≝ - λh,g,L,T. ⦃h, L⦄ ⊢ T ¡[g] → - ∀U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l+1, U⦄ → ⦃h, L⦄ ⊢ U ¡[g]. + λh,g,L,T. ⦃G, L⦄ ⊢ T ¡[h, g] → + ∀U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄ → ⦃G, L⦄ ⊢ U ¡[h, g]. definition IH_snv_lsubsv: ∀h:sh. sd h → relation2 lenv term ≝ - λh,g,L2,T. ⦃h, L2⦄ ⊢ T ¡[g] → - ∀L1. h ⊢ L1 ¡⊑[g] L2 → ⦃h, L1⦄ ⊢ T ¡[g]. + λh,g,L2,T. ⦃h, L2⦄ ⊢ T ¡[h, g] → + ∀L1. h ⊢ L1 ¡⊑[h, g] L2 → ⦃h, L1⦄ ⊢ T ¡[h, g]. (* Properties for the preservation results **********************************) fact snv_cprs_lpr_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - ∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → ⦃h, L1⦄ ⊢ T1 ¡[g] → - ∀T2. L1 ⊢ T1 ➡* T2 → ∀L2. L1 ⊢ ➡ L2 → ⦃h, L2⦄ ⊢ T2 ¡[g]. + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + ∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → ⦃h, L1⦄ ⊢ T1 ¡[h, g] → + ∀T2. L1 ⊢ T1 ➡* T2 → ∀L2. L1 ⊢ ➡ L2 → ⦃h, L2⦄ ⊢ T2 ¡[h, g]. #h #g #L0 #T0 #IH #L1 #T1 #HLT0 #HT1 #T2 #H elim H -T2 [ /2 width=6/ ] -HT1 /4 width=6 by ygt_yprs_trans, cprs_yprs/ qed-. fact ssta_cprs_lpr_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - ∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → ⦃h, L1⦄ ⊢ T1 ¡[g] → - ∀U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + ∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → ⦃h, L1⦄ ⊢ T1 ¡[h, g] → + ∀U1,l. ⦃h, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ → ∀T2. L1 ⊢ T1 ➡* T2 → ∀L2. L1 ⊢ ➡ L2 → - ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄ & L2 ⊢ U1 ⬌* U2. + ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄ & L2 ⊢ U1 ⬌* U2. #h #g #L0 #T0 #IH2 #IH1 #L1 #T1 #H01 #HT1 #U1 #l #HTU1 #T2 #H elim H -T2 [ /2 width=7/ ] #T #T2 #HT1T #HTT2 #IHT1 #L2 #HL12 @@ -70,12 +70,12 @@ lapply (cpcs_trans … HU1 … HU2) -U /2 width=3/ qed-. fact ssta_cpcs_lpr_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - ∀L1,T1,T2. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T2⦄ → - ⦃h, L1⦄ ⊢ T1 ¡[g] → ⦃h, L1⦄ ⊢ T2 ¡[g] → - ∀U1,l1. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l1, U1⦄ → - ∀U2,l2. ⦃h, L1⦄ ⊢ T2 •[g] ⦃l2, U2⦄ → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + ∀L1,T1,T2. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T2⦄ → + ⦃h, L1⦄ ⊢ T1 ¡[h, g] → ⦃h, L1⦄ ⊢ T2 ¡[h, g] → + ∀U1,l1. ⦃h, L1⦄ ⊢ T1 •[h, g] ⦃l1, U1⦄ → + ∀U2,l2. ⦃h, L1⦄ ⊢ T2 •[h, g] ⦃l2, U2⦄ → L1 ⊢ T1 ⬌* T2 → ∀L2. L1 ⊢ ➡ L2 → l1 = l2 ∧ L2 ⊢ U1 ⬌* U2. #h #g #L0 #T0 #IH2 #IH1 #L1 #T1 #T2 #H01 #H02 #HT1 #HT2 #U1 #l1 #HTU1 #U2 #l2 #HTU2 #H #L2 #HL12 @@ -87,28 +87,28 @@ lapply (cpcs_canc_dx … HUW1 … HUW2) -W2 /2 width=1/ qed-. fact snv_sstas_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - ∀L,T. h ⊢ ⦃L0, T0⦄ >[g] ⦃L, T⦄ → ⦃h, L⦄ ⊢ T ¡[g] → - ∀U. ⦃h, L⦄ ⊢ T •*[g] U → ⦃h, L⦄ ⊢ U ¡[g]. + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + ∀L,T. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L, T⦄ → ⦃G, L⦄ ⊢ T ¡[h, g] → + ∀U. ⦃G, L⦄ ⊢ T •*[h, g] U → ⦃G, L⦄ ⊢ U ¡[h, g]. #h #g #L0 #T0 #IH #L #T #H01 #HT #U #H @(sstas_ind … H) -U // -HT /4 width=5 by ygt_yprs_trans, sstas_yprs/ qed-. fact snv_sstas_lpr_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - ∀L1,T. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T⦄ → ⦃h, L1⦄ ⊢ T ¡[g] → - ∀U. ⦃h, L1⦄ ⊢ T •*[g] U → ∀L2. L1 ⊢ ➡ L2 → ⦃h, L2⦄ ⊢ U ¡[g]. + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + ∀L1,T. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T⦄ → ⦃h, L1⦄ ⊢ T ¡[h, g] → + ∀U. ⦃h, L1⦄ ⊢ T •*[h, g] U → ∀L2. L1 ⊢ ➡ L2 → ⦃h, L2⦄ ⊢ U ¡[h, g]. /4 width=7 by snv_sstas_aux, ygt_yprs_trans, sstas_yprs/ qed-. fact sstas_cprs_lpr_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - ∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → ⦃h, L1⦄ ⊢ T1 ¡[g] → - ∀U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 → ∀T2. L1 ⊢ T1 ➡* T2 → ∀L2. L1 ⊢ ➡ L2 → - ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 & L2 ⊢ U1 ⬌* U2. + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + ∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → ⦃h, L1⦄ ⊢ T1 ¡[h, g] → + ∀U1. ⦃h, L1⦄ ⊢ T1 •*[h, g] U1 → ∀T2. L1 ⊢ T1 ➡* T2 → ∀L2. L1 ⊢ ➡ L2 → + ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*[h, g] U2 & L2 ⊢ U1 ⬌* U2. #h #g #L0 #T0 #IH3 #IH2 #IH1 #L1 #T1 #H01 #HT1 #U1 #H @(sstas_ind … H) -U1 [ /3 width=5 by lpr_cprs_conf, ex2_intro/ ] #U1 #W1 #l1 #HTU1 #HUW1 #IHTU1 #T2 #HT12 #L2 #HL12 @@ -128,13 +128,13 @@ lapply (cpcs_trans … HW1 … HW12) -W /3 width=4/ qed-. fact cpds_cprs_lpr_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - ∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → ⦃h, L1⦄ ⊢ T1 ¡[g] → - ∀U1. ⦃h, L1⦄ ⊢ T1 •*➡*[g] U1 → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + ∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → ⦃h, L1⦄ ⊢ T1 ¡[h, g] → + ∀U1. ⦃h, L1⦄ ⊢ T1 •*➡*[h, g] U1 → ∀T2. L1 ⊢ T1 ➡* T2 → ∀L2. L1 ⊢ ➡ L2 → - ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*➡*[g] U2 & L2 ⊢ U1 ➡* U2. + ∃∃U2. ⦃h, L2⦄ ⊢ T2 •*➡*[h, g] U2 & L2 ⊢ U1 ➡* U2. #h #g #L0 #T0 #IH3 #IH2 #IH1 #L1 #T1 #H01 #HT1 #U1 * #W1 #HTW1 #HWU1 #T2 #HT12 #L2 #HL12 elim (sstas_cprs_lpr_aux … IH3 IH2 IH1 … H01 … HTW1 … HT12 … HL12) // -L0 -T0 -T1 #W2 #HTW2 #HW12 lapply (lpr_cprs_conf … HL12 … HWU1) -L1 #HWU1 @@ -143,11 +143,11 @@ elim (cpcs_inv_cprs … H) -H /3 width=3/ qed-. fact ssta_cpds_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - ∀L,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L, T1⦄ → ⦃h, L⦄ ⊢ T1 ¡[g] → - ∀l,U1. ⦃h, L⦄ ⊢ T1 •[g] ⦃l+1, U1⦄ → ∀T2. ⦃h, L⦄ ⊢ T1 •*➡*[g] T2 → - ∃∃U,U2. ⦃h, L⦄ ⊢ U1 •*[g] U & ⦃h, L⦄ ⊢ T2 •*[g] U2 & L ⊢ U ⬌* U2. + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + ∀L,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L, T1⦄ → ⦃G, L⦄ ⊢ T1 ¡[h, g] → + ∀l,U1. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, U1⦄ → ∀T2. ⦃G, L⦄ ⊢ T1 •*➡*[h, g] T2 → + ∃∃U,U2. ⦃G, L⦄ ⊢ U1 •*[h, g] U & ⦃G, L⦄ ⊢ T2 •*[h, g] U2 & ⦃G, L⦄ ⊢ U ⬌* U2. #h #g #L0 #T0 #IH2 #IH1 #L #T1 #H01 #HT1 #l #U1 #HTU1 #T2 * #T #HT1T #HTT2 elim (sstas_strip … HT1T … HTU1) #HU1T destruct [ -HT1T | -L0 -T0 -T1 ] [ elim (ssta_cprs_lpr_aux … IH2 IH1 … HTU1 … HTT2 L) // -L0 -T0 -T /3 width=5/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lift.ma index eb698c599..9b9513f0a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lift.ma @@ -20,8 +20,8 @@ include "basic_2/dynamic/snv.ma". (* Relocation properties ****************************************************) -lemma snv_lift: ∀h,g,K,T. ⦃h, K⦄ ⊢ T ¡[g] → ∀L,d,e. ⇩[d, e] L ≡ K → - ∀U. ⇧[d, e] T ≡ U → ⦃h, L⦄ ⊢ U ¡[g]. +lemma snv_lift: ∀h,g,K,T. ⦃h, K⦄ ⊢ T ¡[h, g] → ∀L,d,e. ⇩[d, e] L ≡ K → + ∀U. ⇧[d, e] T ≡ U → ⦃G, L⦄ ⊢ U ¡[h, g]. #h #g #K #T #H elim H -K -T [ #K #k #L #d #e #_ #X #H >(lift_inv_sort1 … H) -X -K -d -e // @@ -50,8 +50,8 @@ lemma snv_lift: ∀h,g,K,T. ⦃h, K⦄ ⊢ T ¡[g] → ∀L,d,e. ⇩[d, e] L ≡ ] qed. -lemma snv_inv_lift: ∀h,g,L,U. ⦃h, L⦄ ⊢ U ¡[g] → ∀K,d,e. ⇩[d, e] L ≡ K → - ∀T. ⇧[d, e] T ≡ U → ⦃h, K⦄ ⊢ T ¡[g]. +lemma snv_inv_lift: ∀h,g,L,U. ⦃G, L⦄ ⊢ U ¡[h, g] → ∀K,d,e. ⇩[d, e] L ≡ K → + ∀T. ⇧[d, e] T ≡ U → ⦃h, K⦄ ⊢ T ¡[h, g]. #h #g #L #U #H elim H -L -U [ #L #k #K #d #e #_ #X #H >(lift_inv_sort2 … H) -X -L -d -e // @@ -81,7 +81,7 @@ qed-. (* Advanced properties ******************************************************) lemma snv_fsup_conf: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → - ⦃h, L1⦄ ⊢ T1 ¡[g] → ⦃h, L2⦄ ⊢ T2 ¡[g]. + ⦃h, L1⦄ ⊢ T1 ¡[h, g] → ⦃h, L2⦄ ⊢ T2 ¡[h, g]. #h #g #L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [ #I1 #L1 #V1 #H elim (snv_inv_lref … H) -H #I2 #L2 #V2 #H #HV2 diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lpr.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lpr.ma index 39a60a97a..3007d7d67 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_lpr.ma @@ -21,10 +21,10 @@ include "basic_2/dynamic/snv_cpcs.ma". (* Properties on context-free parallel reduction for local environments *****) fact snv_cpr_lpr_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_lsubsv h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_lsubsv h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → ∀L1,T1. L0 = L1 → T0 = T1 → IH_snv_cpr_lpr h g L1 T1. #h #g #L0 #T0 #IH4 #IH3 #IH2 #IH1 #L1 * * [||||*] [ #k #HL0 #HT0 #H1 #X #H2 #L2 #_ destruct -IH4 -IH3 -IH2 -IH1 -H1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ssta.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ssta.ma index 766e88393..3ba30937e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ssta.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ssta.ma @@ -20,9 +20,9 @@ include "basic_2/dynamic/snv_cpcs.ma". (* Properties on stratified static type assignment for terms ****************) fact snv_ssta_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → ∀L1,T1. L0 = L1 → T0 = T1 → IH_snv_ssta h g L1 T1. #h #g #L0 #T0 #IH3 #IH2 #IH1 #L1 * * [||||*] [ #k #HL0 #HT0 #_ #X #l #H2 destruct -IH3 -IH2 -IH1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ssta_lpr.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ssta_lpr.ma index 6c5e0f6a5..bb07981b0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ssta_lpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_ssta_lpr.ma @@ -21,9 +21,9 @@ include "basic_2/dynamic/lsubsv_ssta.ma". (* Properties on sn parallel reduction for local environments ***************) fact ssta_cpr_lpr_aux: ∀h,g,L0,T0. - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → - (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_ssta h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_snv_cpr_lpr h g L1 T1) → + (∀L1,T1. h ⊢ ⦃L0, T0⦄ >[h, g] ⦃L1, T1⦄ → IH_ssta_cpr_lpr h g L1 T1) → ∀L1,T1. L0 = L1 → T0 = T1 → IH_ssta_cpr_lpr h g L1 T1. #h #g #L0 #T0 #IH3 #IH2 #IH1 #L1 * * [|||| *] [ #k #_ #_ #_ #X2 #l #H2 #X3 #H3 #L2 #HL12 -IH3 -IH2 -IH1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_sstas.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_sstas.ma index a663fecfc..bc2d74497 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_sstas.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/snv_sstas.ma @@ -19,8 +19,8 @@ include "basic_2/dynamic/snv.ma". (* Forward_lemmas on iterated stratified static type assignment for terms ***) -lemma snv_sstas_fwd_correct: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 ¡[g] → ⦃h, L⦄ ⊢ T1 •* [g] T2 → - ∃∃U2,l. ⦃h, L⦄ ⊢ T2 •[g] ⦃l, U2⦄. +lemma snv_sstas_fwd_correct: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ¡[h, g] → ⦃G, L⦄ ⊢ T1 •* [h, g] T2 → + ∃∃U2,l. ⦃G, L⦄ ⊢ T2 •[h, g] ⦃l, U2⦄. #h #g #L #T1 #T2 #HT1 #HT12 elim (snv_fwd_ssta … HT1) -HT1 /2 width=5 by sstas_fwd_correct/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/ygt.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/ygt.ma index b593d637d..7b50bfdcf 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/ygt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/ygt.ma @@ -19,9 +19,9 @@ include "basic_2/dynamic/yprs.ma". (* "BIG TREE" PROPER PARALLEL COMPUTATION FOR CLOSURES **********************) inductive ygt (h) (g) (L1) (T1): relation2 lenv term ≝ -| ygt_inj : ∀L,L2,T,T2. h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L, T⦄ → h ⊢ ⦃L, T⦄ ≻[g] ⦃L2, T2⦄ → +| ygt_inj : ∀L,L2,T,T2. h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L, T⦄ → h ⊢ ⦃L, T⦄ ≻[h, g] ⦃L2, T2⦄ → ygt h g L1 T1 L2 T2 -| ygt_step: ∀L,L2,T. ygt h g L1 T1 L T → L ⊢ ➡ L2 → ygt h g L1 T1 L2 T +| ygt_step: ∀L,L2,T. ygt h g L1 T1 L T → ⦃G, L⦄ ⊢ ➡ L2 → ygt h g L1 T1 L2 T . interpretation "'big tree' proper parallel computation (closure)" @@ -29,49 +29,49 @@ interpretation "'big tree' proper parallel computation (closure)" (* Basic forvard lemmas *****************************************************) -lemma ygt_fwd_yprs: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄ → - h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄. +lemma ygt_fwd_yprs: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L2, T2⦄ → + h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L2, T2⦄. #h #g #L1 #L2 #T1 #T2 #H elim H -L2 -T2 /3 width=4 by yprs_strap1, ysc_ypr, ypr_lpr/ qed-. (* Basic properties *********************************************************) -lemma ysc_ygt: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ ≻[g] ⦃L2, T2⦄ → - h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄. +lemma ysc_ygt: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ ≻[h, g] ⦃L2, T2⦄ → + h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L2, T2⦄. /3 width=4/ qed. -lemma ygt_strap1: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ >[g] ⦃L, T⦄ → - h ⊢ ⦃L, T⦄ ≽[g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄. +lemma ygt_strap1: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L, T⦄ → + h ⊢ ⦃L, T⦄ ≽[h, g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L2, T2⦄. #h #g #L1 #L #L2 #T1 #T #T2 #H1 #H2 lapply (ygt_fwd_yprs … H1) #H0 elim (ypr_inv_ysc … H2) -H2 [| * #HL2 #H destruct ] /2 width=4/ qed-. -lemma ygt_strap2: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ ≽[g] ⦃L, T⦄ → - h ⊢ ⦃L, T⦄ >[g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄. +lemma ygt_strap2: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ ≽[h, g] ⦃L, T⦄ → + h ⊢ ⦃L, T⦄ >[h, g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L2, T2⦄. #h #g #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim H2 -L2 -T2 [ /3 width=4 by ygt_inj, yprs_strap2/ | /2 width=3/ ] qed-. -lemma ygt_yprs_trans: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ >[g] ⦃L, T⦄ → - h ⊢ ⦃L, T⦄ ≥[g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄. +lemma ygt_yprs_trans: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L, T⦄ → + h ⊢ ⦃L, T⦄ ≥[h, g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L2, T2⦄. #h #g #L1 #L #L2 #T1 #T #T2 #HT1 #HT2 @(yprs_ind … HT2) -L2 -T2 // /2 width=4 by ygt_strap1/ qed-. -lemma yprs_ygt_trans: ∀h,g,L1,L,T1,T. h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L, T⦄ → - ∀L2,T2. h ⊢ ⦃L, T⦄ >[g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄. +lemma yprs_ygt_trans: ∀h,g,L1,L,T1,T. h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L, T⦄ → + ∀L2,T2. h ⊢ ⦃L, T⦄ >[h, g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L2, T2⦄. #h #g #L1 #L #T1 #T #HT1 @(yprs_ind … HT1) -L -T // /3 width=4 by ygt_strap2/ qed-. -lemma fsupp_ygt: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃+ ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄. +lemma fsupp_ygt: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃+ ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[h, g] ⦃L2, T2⦄. #h #g #L1 #L2 #T1 #T2 #H @(fsupp_ind … L2 T2 H) -L2 -T2 /3 width=1/ /3 width=4/ qed. -lemma cprs_ygt: ∀h,g,L,T1,T2. L ⊢ T1 ➡* T2 → (T1 = T2 → ⊥) → - h ⊢ ⦃L, T1⦄ >[g] ⦃L, T2⦄. +lemma cprs_ygt: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → (T1 = T2 → ⊥) → + h ⊢ ⦃L, T1⦄ >[h, g] ⦃L, T2⦄. #h #g #L #T1 #T2 #H @(cprs_ind … H) -T2 [ #H elim H // | #T #T2 #_ #HT2 #IHT1 #HT12 @@ -83,8 +83,8 @@ lemma cprs_ygt: ∀h,g,L,T1,T2. L ⊢ T1 ➡* T2 → (T1 = T2 → ⊥) → ] qed. -lemma sstas_ygt: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 •*[g] T2 → (T1 = T2 → ⊥) → - h ⊢ ⦃L, T1⦄ >[g] ⦃L, T2⦄. +lemma sstas_ygt: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 → (T1 = T2 → ⊥) → + h ⊢ ⦃L, T1⦄ >[h, g] ⦃L, T2⦄. #h #g #L #T1 #T2 #H @(sstas_ind … H) -T2 [ #H elim H // | #T #T2 #l #_ #HT2 #IHT1 #HT12 -HT12 @@ -96,6 +96,6 @@ lemma sstas_ygt: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 •*[g] T2 → (T1 = T2 → ] qed. -lemma lsubsv_ygt: ∀h,g,L1,L2,T. h ⊢ L2 ¡⊑[g] L1 → (L1 = L2 → ⊥) → - h ⊢ ⦃L1, T⦄ >[g] ⦃L2, T⦄. +lemma lsubsv_ygt: ∀h,g,L1,L2,T. h ⊢ L2 ¡⊑[h, g] L1 → (L1 = L2 → ⊥) → + h ⊢ ⦃L1, T⦄ >[h, g] ⦃L2, T⦄. /4 width=1/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/ypr.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/ypr.ma index a471fe949..84879b10a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/ypr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/ypr.ma @@ -23,8 +23,8 @@ inductive ypr (h) (g) (L1) (T1): relation2 lenv term ≝ | ypr_fsup : ∀L2,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ypr h g L1 T1 L2 T2 | ypr_lpr : ∀L2. L1 ⊢ ➡ L2 → ypr h g L1 T1 L2 T1 | ypr_cpr : ∀T2. L1 ⊢ T1 ➡ T2 → ypr h g L1 T1 L1 T2 -| ypr_ssta : ∀T2,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l+1, T2⦄ → ypr h g L1 T1 L1 T2 -| ypr_lsubsv: ∀L2. h ⊢ L2 ¡⊑[g] L1 → ypr h g L1 T1 L2 T1 +| ypr_ssta : ∀T2,l. ⦃h, L1⦄ ⊢ T1 •[h, g] ⦃l+1, T2⦄ → ypr h g L1 T1 L1 T2 +| ypr_lsubsv: ∀L2. h ⊢ L2 ¡⊑[h, g] L1 → ypr h g L1 T1 L2 T1 . interpretation diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/yprs.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/yprs.ma index cc94dff44..a8be72e72 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/yprs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/yprs.ma @@ -28,13 +28,13 @@ interpretation "'big tree' parallel computation (closure)" (* Basic eliminators ********************************************************) lemma yprs_ind: ∀h,g,L1,T1. ∀R:relation2 lenv term. R L1 T1 → - (∀L,L2,T,T2. h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L, T⦄ → h ⊢ ⦃L, T⦄ ≽[g] ⦃L2, T2⦄ → R L T → R L2 T2) → - ∀L2,T2. h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄ → R L2 T2. + (∀L,L2,T,T2. h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L, T⦄ → h ⊢ ⦃L, T⦄ ≽[h, g] ⦃L2, T2⦄ → R L T → R L2 T2) → + ∀L2,T2. h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L2, T2⦄ → R L2 T2. /3 width=7 by bi_TC_star_ind/ qed-. lemma yprs_ind_dx: ∀h,g,L2,T2. ∀R:relation2 lenv term. R L2 T2 → - (∀L1,L,T1,T. h ⊢ ⦃L1, T1⦄ ≽[g] ⦃L, T⦄ → h ⊢ ⦃L, T⦄ ≥[g] ⦃L2, T2⦄ → R L T → R L1 T1) → - ∀L1,T1. h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄ → R L1 T1. + (∀L1,L,T1,T. h ⊢ ⦃L1, T1⦄ ≽[h, g] ⦃L, T⦄ → h ⊢ ⦃L, T⦄ ≥[h, g] ⦃L2, T2⦄ → R L T → R L1 T1) → + ∀L1,T1. h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L2, T2⦄ → R L1 T1. /3 width=7 by bi_TC_star_ind_dx/ qed-. (* Basic properties *********************************************************) @@ -42,41 +42,41 @@ lemma yprs_ind_dx: ∀h,g,L2,T2. ∀R:relation2 lenv term. R L2 T2 → lemma yprs_refl: ∀h,g. bi_reflexive … (yprs h g). /2 width=1/ qed. -lemma ypr_yprs: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ ≽[g] ⦃L2, T2⦄ → - h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄. +lemma ypr_yprs: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ ≽[h, g] ⦃L2, T2⦄ → + h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L2, T2⦄. /2 width=1/ qed. -lemma yprs_strap1: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L, T⦄ → - h ⊢ ⦃L, T⦄ ≽[g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄. +lemma yprs_strap1: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L, T⦄ → + h ⊢ ⦃L, T⦄ ≽[h, g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L2, T2⦄. /2 width=4/ qed-. -lemma yprs_strap2: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ ≽[g] ⦃L, T⦄ → - h ⊢ ⦃L, T⦄ ≥[g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄. +lemma yprs_strap2: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ ≽[h, g] ⦃L, T⦄ → + h ⊢ ⦃L, T⦄ ≥[h, g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L2, T2⦄. /2 width=4/ qed-. (* Note: this is a general property of bi_TC *) lemma fsupp_yprs: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃+ ⦃L2, T2⦄ → - h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄. + h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L2, T2⦄. #h #g #L1 #L2 #T1 #T2 #H @(fsupp_ind … L2 T2 H) -L2 -T2 /3 width=1/ /3 width=4/ qed. -lemma cprs_yprs: ∀h,g,L,T1,T2. L ⊢ T1 ➡* T2 → h ⊢ ⦃L, T1⦄ ≥[g] ⦃L, T2⦄. +lemma cprs_yprs: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → h ⊢ ⦃L, T1⦄ ≥[h, g] ⦃L, T2⦄. #h #g #L #T1 #T2 #H @(cprs_ind … H) -T2 // /3 width=4 by ypr_cpr, yprs_strap1/ qed. -lemma lprs_yprs: ∀h,g,L1,L2,T. L1 ⊢ ➡* L2 → h ⊢ ⦃L1, T⦄ ≥[g] ⦃L2, T⦄. +lemma lprs_yprs: ∀h,g,L1,L2,T. L1 ⊢ ➡* L2 → h ⊢ ⦃L1, T⦄ ≥[h, g] ⦃L2, T⦄. #h #g #L1 #L2 #T #H @(lprs_ind … H) -L2 // /3 width=4 by ypr_lpr, yprs_strap1/ qed. -lemma sstas_yprs: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 •*[g] T2 → - h ⊢ ⦃L, T1⦄ ≥[g] ⦃L, T2⦄. +lemma sstas_yprs: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 •*[h, g] T2 → + h ⊢ ⦃L, T1⦄ ≥[h, g] ⦃L, T2⦄. #h #g #L #T1 #T2 #H @(sstas_ind … H) -T2 // /3 width=4 by ypr_ssta, yprs_strap1/ qed. -lemma lsubsv_yprs: ∀h,g,L1,L2,T. h ⊢ L2 ¡⊑[g] L1 → h ⊢ ⦃L1, T⦄ ≥[g] ⦃L2, T⦄. +lemma lsubsv_yprs: ∀h,g,L1,L2,T. h ⊢ L2 ¡⊑[h, g] L1 → h ⊢ ⦃L1, T⦄ ≥[h, g] ⦃L2, T⦄. /3 width=1/ qed. lemma cprs_lpr_yprs: ∀h,g,L1,L2,T1,T2. L1 ⊢ T1 ➡* T2 → L1 ⊢ ➡ L2 → - h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄. + h ⊢ ⦃L1, T1⦄ ≥[h, g] ⦃L2, T2⦄. /3 width=4 by yprs_strap1, ypr_lpr, cprs_yprs/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/ysc.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/ysc.ma index 87ca83fb1..09785875c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/dynamic/ysc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/ysc.ma @@ -20,8 +20,8 @@ include "basic_2/dynamic/ypr.ma". inductive ysc (h) (g) (L1) (T1): relation2 lenv term ≝ | ysc_fsup : ∀L2,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ysc h g L1 T1 L2 T2 | ysc_cpr : ∀T2. L1 ⊢ T1 ➡ T2 → (T1 = T2 → ⊥) → ysc h g L1 T1 L1 T2 -| ysc_ssta : ∀T2,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l+1, T2⦄ → ysc h g L1 T1 L1 T2 -| ysc_lsubsv: ∀L2. h ⊢ L2 ¡⊑[g] L1 → (L1 = L2 → ⊥) → ysc h g L1 T1 L2 T1 +| ysc_ssta : ∀T2,l. ⦃h, L1⦄ ⊢ T1 •[h, g] ⦃l+1, T2⦄ → ysc h g L1 T1 L1 T2 +| ysc_lsubsv: ∀L2. h ⊢ L2 ¡⊑[h, g] L1 → (L1 = L2 → ⊥) → ysc h g L1 T1 L2 T1 . interpretation @@ -30,15 +30,15 @@ interpretation (* Basic properties *********************************************************) -lemma ysc_ypr: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ ≻[g] ⦃L2, T2⦄ → - h ⊢ ⦃L1, T1⦄ ≽[g] ⦃L2, T2⦄. +lemma ysc_ypr: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ ≻[h, g] ⦃L2, T2⦄ → + h ⊢ ⦃L1, T1⦄ ≽[h, g] ⦃L2, T2⦄. #h #g #L1 #L2 #T1 #T2 * -L2 -T2 /2 width=1/ /2 width=2/ qed. (* Inversion lemmas on "big tree" parallel reduction for closures ***********) -lemma ypr_inv_ysc: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ ≽[g] ⦃L2, T2⦄ → - h ⊢ ⦃L1, T1⦄ ≻[g] ⦃L2, T2⦄ ∨ (L1 ⊢ ➡ L2 ∧ T1 = T2). +lemma ypr_inv_ysc: ∀h,g,L1,L2,T1,T2. h ⊢ ⦃L1, T1⦄ ≽[h, g] ⦃L2, T2⦄ → + h ⊢ ⦃L1, T1⦄ ≻[h, g] ⦃L2, T2⦄ ∨ (L1 ⊢ ➡ L2 ∧ T1 = T2). #h #g #L1 #L2 #T1 #T2 * -L2 -T2 /3 width=1/ /3 width=2/ [ #T2 #HT12 elim (term_eq_dec T1 T2) #H destruct /3 width=1/ /4 width=1/ | #L2 #HL21 elim (lenv_eq_dec L1 L2) #H destruct /3 width=1/ /4 width=1/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs.ma b/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs.ma index d26648839..a7e5b7514 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs.ma @@ -25,14 +25,14 @@ interpretation "context-sensitive parallel equivalence (term)" (* Basic eliminators ********************************************************) lemma cpcs_ind: ∀L,T1. ∀R:predicate term. R T1 → - (∀T,T2. L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → R T → R T2) → - ∀T2. L ⊢ T1 ⬌* T2 → R T2. + (∀T,T2. ⦃G, L⦄ ⊢ T1 ⬌* T → ⦃G, L⦄ ⊢ T ⬌ T2 → R T → R T2) → + ∀T2. ⦃G, L⦄ ⊢ T1 ⬌* T2 → R T2. #L #T1 #R #HT1 #IHT1 #T2 #HT12 @(TC_star_ind … HT1 IHT1 … HT12) // qed-. lemma cpcs_ind_dx: ∀L,T2. ∀R:predicate term. R T2 → - (∀T1,T. L ⊢ T1 ⬌ T → L ⊢ T ⬌* T2 → R T → R T1) → - ∀T1. L ⊢ T1 ⬌* T2 → R T1. + (∀T1,T. ⦃G, L⦄ ⊢ T1 ⬌ T → ⦃G, L⦄ ⊢ T ⬌* T2 → R T → R T1) → + ∀T1. ⦃G, L⦄ ⊢ T1 ⬌* T2 → R T1. #L #T2 #R #HT2 #IHT2 #T1 #HT12 @(TC_star_ind_dx … HT2 IHT2 … HT12) // qed-. @@ -47,38 +47,38 @@ lemma cpcs_refl: ∀L. reflexive … (cpcs L). lemma cpcs_sym: ∀L. symmetric … (cpcs L). #L @TC_symmetric // qed. -lemma cpc_cpcs: ∀L,T1,T2. L ⊢ T1 ⬌ T2 → L ⊢ T2 ⬌* T2. +lemma cpc_cpcs: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬌ T2 → ⦃G, L⦄ ⊢ T2 ⬌* T2. /2 width=1/ qed. -lemma cpcs_strap1: ∀L,T1,T,T2. L ⊢ T1 ⬌* T → L ⊢ T ⬌ T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_strap1: ∀L,T1,T,T2. ⦃G, L⦄ ⊢ T1 ⬌* T → ⦃G, L⦄ ⊢ T ⬌ T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L @step qed. -lemma cpcs_strap2: ∀L,T1,T,T2. L ⊢ T1 ⬌ T → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_strap2: ∀L,T1,T,T2. ⦃G, L⦄ ⊢ T1 ⬌ T → ⦃G, L⦄ ⊢ T ⬌* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L @TC_strap qed. (* Basic_1: was: pc3_pr2_r *) -lemma cpr_cpcs_dx: ∀L,T1,T2. L ⊢ T1 ➡ T2 → L ⊢ T1 ⬌* T2. +lemma cpr_cpcs_dx: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=1/ qed. (* Basic_1: was: pc3_pr2_x *) -lemma cpr_cpcs_sn: ∀L,T1,T2. L ⊢ T2 ➡ T1 → L ⊢ T1 ⬌* T2. +lemma cpr_cpcs_sn: ∀L,T1,T2. ⦃G, L⦄ ⊢ T2 ➡ T1 → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=1/ qed. -lemma cpcs_cpr_strap1: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T ➡ T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_cpr_strap1: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ⬌* T → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=3/ qed. (* Basic_1: was: pc3_pr2_u *) -lemma cpcs_cpr_strap2: ∀L,T1,T. L ⊢ T1 ➡ T → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_cpr_strap2: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ➡ T → ∀T2. ⦃G, L⦄ ⊢ T ⬌* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=3/ qed. -lemma cpcs_cpr_div: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T2 ➡ T → L ⊢ T1 ⬌* T2. +lemma cpcs_cpr_div: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ⬌* T → ∀T2. ⦃G, L⦄ ⊢ T2 ➡ T → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=3/ qed. -lemma cpr_div: ∀L,T1,T. L ⊢ T1 ➡ T → ∀T2. L ⊢ T2 ➡ T → L ⊢ T1 ⬌* T2. +lemma cpr_div: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ➡ T → ∀T2. ⦃G, L⦄ ⊢ T2 ➡ T → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=3/ qed-. (* Basic_1: was: pc3_pr2_u2 *) -lemma cpcs_cpr_conf: ∀L,T1,T. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_cpr_conf: ∀L,T1,T. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ⬌* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=3/ qed. (* Basic_1: removed theorems 9: diff --git a/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_aaa.ma index 190092f86..1a5e48687 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_aaa.ma @@ -19,8 +19,8 @@ include "basic_2/equivalence/cpcs_cpcs.ma". (* Main properties about atomic arity assignment on terms *******************) -theorem aaa_cpcs_mono: ∀L,T1,T2. L ⊢ T1 ⬌* T2 → - ∀A1. L ⊢ T1 ⁝ A1 → ∀A2. L ⊢ T2 ⁝ A2 → +theorem aaa_cpcs_mono: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬌* T2 → + ∀A1. ⦃G, L⦄ ⊢ T1 ⁝ A1 → ∀A2. ⦃G, L⦄ ⊢ T2 ⁝ A2 → A1 = A2. #L #T1 #T2 #HT12 #A1 #HA1 #A2 #HA2 elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2 diff --git a/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cpcs.ma b/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cpcs.ma index 44cbd7491..125cc70fe 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cpcs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cpcs.ma @@ -20,8 +20,8 @@ include "basic_2/equivalence/cpcs_cprs.ma". (* Advanced inversion lemmas ************************************************) -lemma cpcs_inv_cprs: ∀L,T1,T2. L ⊢ T1 ⬌* T2 → - ∃∃T. L ⊢ T1 ➡* T & L ⊢ T2 ➡* T. +lemma cpcs_inv_cprs: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬌* T2 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ➡* T & ⦃G, L⦄ ⊢ T2 ➡* T. #L #T1 #T2 #H @(cpcs_ind … H) -T2 [ /3 width=3/ | #T #T2 #_ #HT2 * #T0 #HT10 elim HT2 -HT2 #HT2 #HT0 @@ -33,27 +33,27 @@ lemma cpcs_inv_cprs: ∀L,T1,T2. L ⊢ T1 ⬌* T2 → qed-. (* Basic_1: was: pc3_gen_sort *) -lemma cpcs_inv_sort: ∀L,k1,k2. L ⊢ ⋆k1 ⬌* ⋆k2 → k1 = k2. +lemma cpcs_inv_sort: ∀L,k1,k2. ⦃G, L⦄ ⊢ ⋆k1 ⬌* ⋆k2 → k1 = k2. #L #k1 #k2 #H elim (cpcs_inv_cprs … H) -H #T #H1 >(cprs_inv_sort1 … H1) -T #H2 lapply (cprs_inv_sort1 … H2) -L #H destruct // qed-. -lemma cpcs_inv_abst1: ∀a,L,W1,T1,T. L ⊢ ⓛ{a}W1.T1 ⬌* T → - ∃∃W2,T2. L ⊢ T ➡* ⓛ{a}W2.T2 & L ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2. +lemma cpcs_inv_abst1: ∀a,L,W1,T1,T. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ⬌* T → + ∃∃W2,T2. ⦃G, L⦄ ⊢ T ➡* ⓛ{a}W2.T2 & ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2. #a #L #W1 #T1 #T #H elim (cpcs_inv_cprs … H) -H #X #H1 #H2 elim (cprs_inv_abst1 … H1) -H1 #W2 #T2 #HW12 #HT12 #H destruct @(ex2_2_intro … H2) -H2 /2 width=2/ (**) (* explicit constructor, /3 width=6/ is slow *) qed-. -lemma cpcs_inv_abst2: ∀a,L,W1,T1,T. L ⊢ T ⬌* ⓛ{a}W1.T1 → - ∃∃W2,T2. L ⊢ T ➡* ⓛ{a}W2.T2 & L ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2. +lemma cpcs_inv_abst2: ∀a,L,W1,T1,T. ⦃G, L⦄ ⊢ T ⬌* ⓛ{a}W1.T1 → + ∃∃W2,T2. ⦃G, L⦄ ⊢ T ➡* ⓛ{a}W2.T2 & ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2. /3 width=1 by cpcs_inv_abst1, cpcs_sym/ qed-. (* Basic_1: was: pc3_gen_sort_abst *) -lemma cpcs_inv_sort_abst: ∀a,L,W,T,k. L ⊢ ⋆k ⬌* ⓛ{a}W.T → ⊥. +lemma cpcs_inv_sort_abst: ∀a,L,W,T,k. ⦃G, L⦄ ⊢ ⋆k ⬌* ⓛ{a}W.T → ⊥. #a #L #W #T #k #H elim (cpcs_inv_cprs … H) -H #X #H1 >(cprs_inv_sort1 … H1) -X #H2 @@ -63,7 +63,7 @@ qed-. (* Basic_1: was: pc3_gen_lift *) lemma cpcs_inv_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀T2,U2. ⇧[d, e] T2 ≡ U2 → - L ⊢ U1 ⬌* U2 → K ⊢ T1 ⬌* T2. + ⦃G, L⦄ ⊢ U1 ⬌* U2 → K ⊢ T1 ⬌* T2. #L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12 elim (cpcs_inv_cprs … HU12) -HU12 #U #HU1 #HU2 elim (cprs_inv_lift1 … HU1 … HLK … HTU1) -U1 #T #HTU #HT1 @@ -87,17 +87,17 @@ lapply (lprs_cprs_trans … HT1 … HL12) -HT1 lapply (lprs_cprs_trans … HT2 … HL12) -L2 /2 width=3/ qed-. -lemma cpr_cprs_conf_cpcs: ∀L,T,T1,T2. L ⊢ T ➡* T1 → L ⊢ T ➡ T2 → L ⊢ T1 ⬌* T2. +lemma cpr_cprs_conf_cpcs: ∀L,T,T1,T2. ⦃G, L⦄ ⊢ T ➡* T1 → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T #T1 #T2 #HT1 #HT2 elim (cprs_strip … HT1 … HT2) /2 width=3 by cpr_cprs_div/ qed-. -lemma cprs_cpr_conf_cpcs: ∀L,T,T1,T2. L ⊢ T ➡* T1 → L ⊢ T ➡ T2 → L ⊢ T2 ⬌* T1. +lemma cprs_cpr_conf_cpcs: ∀L,T,T1,T2. ⦃G, L⦄ ⊢ T ➡* T1 → ⦃G, L⦄ ⊢ T ➡ T2 → ⦃G, L⦄ ⊢ T2 ⬌* T1. #L #T #T1 #T2 #HT1 #HT2 elim (cprs_strip … HT1 … HT2) /2 width=3 by cprs_cpr_div/ qed-. -lemma cprs_conf_cpcs: ∀L,T,T1,T2. L ⊢ T ➡* T1 → L ⊢ T ➡* T2 → L ⊢ T1 ⬌* T2. +lemma cprs_conf_cpcs: ∀L,T,T1,T2. ⦃G, L⦄ ⊢ T ➡* T1 → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T #T1 #T2 #HT1 #HT2 elim (cprs_conf … HT1 … HT2) /2 width=3/ qed-. @@ -118,24 +118,24 @@ lemma lpr_cpr_conf: ∀L1,L2. L1 ⊢ ➡ L2 → ∀T1,T2. L1 ⊢ T1 ➡ T2 → L /3 width=5 by lpr_cprs_conf, cpr_cprs/ qed-. (* Basic_1: was only: pc3_thin_dx *) -lemma cpcs_flat: ∀L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L ⊢ T1 ⬌* T2 → - ∀I. L ⊢ ⓕ{I}V1. T1 ⬌* ⓕ{I}V2. T2. +lemma cpcs_flat: ∀L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬌* V2 → ∀T1,T2. ⦃G, L⦄ ⊢ T1 ⬌* T2 → + ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1. T1 ⬌* ⓕ{I}V2. T2. #L #V1 #V2 #HV12 #T1 #T2 #HT12 #I elim (cpcs_inv_cprs … HV12) -HV12 #V #HV1 #HV2 elim (cpcs_inv_cprs … HT12) -HT12 /3 width=5 by cprs_flat, cprs_div/ (**) (* /3 width=5/ is too slow *) qed. -lemma cpcs_flat_dx_cpr_rev: ∀L,V1,V2. L ⊢ V2 ➡ V1 → ∀T1,T2. L ⊢ T1 ⬌* T2 → - ∀I. L ⊢ ⓕ{I}V1. T1 ⬌* ⓕ{I}V2. T2. +lemma cpcs_flat_dx_cpr_rev: ∀L,V1,V2. ⦃G, L⦄ ⊢ V2 ➡ V1 → ∀T1,T2. ⦃G, L⦄ ⊢ T1 ⬌* T2 → + ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1. T1 ⬌* ⓕ{I}V2. T2. /3 width=1/ qed. lemma cpcs_bind_dx: ∀a,I,L,V,T1,T2. L.ⓑ{I}V ⊢ T1 ⬌* T2 → - L ⊢ ⓑ{a,I}V. T1 ⬌* ⓑ{a,I}V. T2. + ⦃G, L⦄ ⊢ ⓑ{a,I}V. T1 ⬌* ⓑ{a,I}V. T2. #a #I #L #V #T1 #T2 #HT12 elim (cpcs_inv_cprs … HT12) -HT12 /3 width=5 by cprs_div, cprs_bind/ (**) (* /3 width=5/ is a bit slow *) qed. -lemma cpcs_bind_sn: ∀a,I,L,V1,V2,T. L ⊢ V1 ⬌* V2 → L ⊢ ⓑ{a,I}V1. T ⬌* ⓑ{a,I}V2. T. +lemma cpcs_bind_sn: ∀a,I,L,V1,V2,T. ⦃G, L⦄ ⊢ V1 ⬌* V2 → ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T ⬌* ⓑ{a,I}V2. T. #a #I #L #V1 #V2 #T #HV12 elim (cpcs_inv_cprs … HV12) -HV12 /3 width=5 by cprs_div, cprs_bind/ (**) (* /3 width=5/ is a bit slow *) qed. @@ -150,7 +150,7 @@ qed-. (* Basic_1: was: pc3_lift *) lemma cpcs_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K → ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀T2,U2. ⇧[d, e] T2 ≡ U2 → - K ⊢ T1 ⬌* T2 → L ⊢ U1 ⬌* U2. + K ⊢ T1 ⬌* T2 → ⦃G, L⦄ ⊢ U1 ⬌* U2. #L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12 elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2 elim (lift_total T d e) #U #HTU @@ -158,14 +158,14 @@ lapply (cprs_lift … HT1 … HLK … HTU1 … HTU) -T1 #HU1 lapply (cprs_lift … HT2 … HLK … HTU2 … HTU) -K -T2 -T -d -e /2 width=3/ qed. -lemma cpcs_strip: ∀L,T1,T. L ⊢ T ⬌* T1 → ∀T2. L ⊢ T ⬌ T2 → - ∃∃T0. L ⊢ T1 ⬌ T0 & L ⊢ T2 ⬌* T0. +lemma cpcs_strip: ∀L,T1,T. ⦃G, L⦄ ⊢ T ⬌* T1 → ∀T2. ⦃G, L⦄ ⊢ T ⬌ T2 → + ∃∃T0. ⦃G, L⦄ ⊢ T1 ⬌ T0 & ⦃G, L⦄ ⊢ T2 ⬌* T0. #L #T1 #T @TC_strip1 /2 width=3/ qed-. (* More inversion lemmas ****************************************************) -lemma cpcs_inv_abst_sn: ∀a1,a2,L,W1,W2,T1,T2. L ⊢ ⓛ{a1}W1.T1 ⬌* ⓛ{a2}W2.T2 → - ∧∧ L ⊢ W1 ⬌* W2 & L.ⓛW1 ⊢ T1 ⬌* T2 & a1 = a2. +lemma cpcs_inv_abst_sn: ∀a1,a2,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a1}W1.T1 ⬌* ⓛ{a2}W2.T2 → + ∧∧ ⦃G, L⦄ ⊢ W1 ⬌* W2 & L.ⓛW1 ⊢ T1 ⬌* T2 & a1 = a2. #a1 #a2 #L #W1 #W2 #T1 #T2 #H elim (cpcs_inv_cprs … H) -H #T #H1 #H2 elim (cprs_inv_abst1 … H1) -H1 #W0 #T0 #HW10 #HT10 #H destruct @@ -175,8 +175,8 @@ lapply (lprs_cpcs_trans … (L.ⓛW1) … HT2) /2 width=1/ -HT2 #HT2 /4 width=3 by and3_intro, cprs_div, cpcs_cprs_div, cpcs_sym/ qed-. -lemma cpcs_inv_abst_dx: ∀a1,a2,L,W1,W2,T1,T2. L ⊢ ⓛ{a1}W1.T1 ⬌* ⓛ{a2}W2.T2 → - ∧∧ L ⊢ W1 ⬌* W2 & L. ⓛW2 ⊢ T1 ⬌* T2 & a1 = a2. +lemma cpcs_inv_abst_dx: ∀a1,a2,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a1}W1.T1 ⬌* ⓛ{a2}W2.T2 → + ∧∧ ⦃G, L⦄ ⊢ W1 ⬌* W2 & L. ⓛW2 ⊢ T1 ⬌* T2 & a1 = a2. #a1 #a2 #L #W1 #W2 #T1 #T2 #HT12 lapply (cpcs_sym … HT12) -HT12 #HT12 elim (cpcs_inv_abst_sn … HT12) -HT12 /3 width=1/ @@ -185,23 +185,23 @@ qed-. (* Main properties **********************************************************) (* Basic_1: was pc3_t *) -theorem cpcs_trans: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. +theorem cpcs_trans: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ⬌* T → ∀T2. ⦃G, L⦄ ⊢ T ⬌* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T1 #T #HT1 #T2 @(trans_TC … HT1) qed-. -theorem cpcs_canc_sn: ∀L,T,T1,T2. L ⊢ T ⬌* T1 → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. +theorem cpcs_canc_sn: ∀L,T,T1,T2. ⦃G, L⦄ ⊢ T ⬌* T1 → ⦃G, L⦄ ⊢ T ⬌* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=3 by cpcs_trans, cpcs_sym/ qed-. (**) (* /3 width=3/ is too slow *) -theorem cpcs_canc_dx: ∀L,T,T1,T2. L ⊢ T1 ⬌* T → L ⊢ T2 ⬌* T → L ⊢ T1 ⬌* T2. +theorem cpcs_canc_dx: ∀L,T,T1,T2. ⦃G, L⦄ ⊢ T1 ⬌* T → ⦃G, L⦄ ⊢ T2 ⬌* T → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=3 by cpcs_trans, cpcs_sym/ qed-. (**) (* /3 width=3/ is too slow *) -lemma cpcs_bind1: ∀a,I,L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓑ{I}V1 ⊢ T1 ⬌* T2 → - L ⊢ ⓑ{a,I}V1. T1 ⬌* ⓑ{a,I}V2. T2. +lemma cpcs_bind1: ∀a,I,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓑ{I}V1 ⊢ T1 ⬌* T2 → + ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ⬌* ⓑ{a,I}V2. T2. #a #I #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cpcs_trans … (ⓑ{a,I}V1.T2)) /2 width=1/ qed. -lemma cpcs_bind2: ∀a,I,L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓑ{I}V2 ⊢ T1 ⬌* T2 → - L ⊢ ⓑ{a,I}V1. T1 ⬌* ⓑ{a,I}V2. T2. +lemma cpcs_bind2: ∀a,I,L,V1,V2. ⦃G, L⦄ ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓑ{I}V2 ⊢ T1 ⬌* T2 → + ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ⬌* ⓑ{a,I}V2. T2. #a #I #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cpcs_trans … (ⓑ{a,I}V2.T1)) /2 width=1/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cprs.ma b/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cprs.ma index 8c1fec2e1..30ee9fb9d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cprs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cprs.ma @@ -20,40 +20,40 @@ include "basic_2/equivalence/cpcs.ma". (* Properties about context sensitive computation on terms ******************) (* Basic_1: was: pc3_pr3_r *) -lemma cpcs_cprs_dx: ∀L,T1,T2. L ⊢ T1 ➡* T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_cprs_dx: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T1 #T2 #H @(cprs_ind … H) -T2 /width=1/ /3 width=3/ qed. (* Basic_1: was: pc3_pr3_x *) -lemma cpcs_cprs_sn: ∀L,T1,T2. L ⊢ T2 ➡* T1 → L ⊢ T1 ⬌* T2. +lemma cpcs_cprs_sn: ∀L,T1,T2. ⦃G, L⦄ ⊢ T2 ➡* T1 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T1 #T2 #H @(cprs_ind_dx … H) -T2 /width=1/ /3 width=3/ qed. -lemma cpcs_cprs_strap1: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T ➡* T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_cprs_strap1: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ⬌* T → ∀T2. ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T1 #T #HT1 #T2 #H @(cprs_ind … H) -T2 /width=1/ /2 width=3/ qed. -lemma cpcs_cprs_strap2: ∀L,T1,T. L ⊢ T1 ➡* T → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_cprs_strap2: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ➡* T → ∀T2. ⦃G, L⦄ ⊢ T ⬌* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T1 #T #H #T2 #HT2 @(cprs_ind_dx … H) -T1 /width=1/ /2 width=3/ qed. -lemma cpcs_cprs_div: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T2 ➡* T → L ⊢ T1 ⬌* T2. +lemma cpcs_cprs_div: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ⬌* T → ∀T2. ⦃G, L⦄ ⊢ T2 ➡* T → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T1 #T #HT1 #T2 #H @(cprs_ind_dx … H) -T2 /width=1/ /2 width=3/ qed. (* Basic_1: was: pc3_pr3_conf *) -lemma cpcs_cprs_conf: ∀L,T1,T. L ⊢ T ➡* T1 → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2. +lemma cpcs_cprs_conf: ∀L,T1,T. ⦃G, L⦄ ⊢ T ➡* T1 → ∀T2. ⦃G, L⦄ ⊢ T ⬌* T2 → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T1 #T #H #T2 #HT2 @(cprs_ind … H) -T1 /width=1/ /2 width=3/ qed. (* Basic_1: was: pc3_pr3_t *) (* Basic_1: note: pc3_pr3_t should be renamed *) -lemma cprs_div: ∀L,T1,T. L ⊢ T1 ➡* T → ∀T2. L ⊢ T2 ➡* T → L ⊢ T1 ⬌* T2. +lemma cprs_div: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ➡* T → ∀T2. ⦃G, L⦄ ⊢ T2 ➡* T → ⦃G, L⦄ ⊢ T1 ⬌* T2. #L #T1 #T #HT1 #T2 #H @(cprs_ind_dx … H) -T2 /2 width=1/ /2 width=3/ qed. -lemma cprs_cpr_div: ∀L,T1,T. L ⊢ T1 ➡* T → ∀T2. L ⊢ T2 ➡ T → L ⊢ T1 ⬌* T2. +lemma cprs_cpr_div: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ➡* T → ∀T2. ⦃G, L⦄ ⊢ T2 ➡ T → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=5 by cpr_cprs, cprs_div/ qed-. -lemma cpr_cprs_div: ∀L,T1,T. L ⊢ T1 ➡ T → ∀T2. L ⊢ T2 ➡* T → L ⊢ T1 ⬌* T2. +lemma cpr_cprs_div: ∀L,T1,T. ⦃G, L⦄ ⊢ T1 ➡ T → ∀T2. ⦃G, L⦄ ⊢ T2 ➡* T → ⦃G, L⦄ ⊢ T1 ⬌* T2. /3 width=3 by cpr_cprs, cprs_div/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/cl_weight.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/cl_weight.ma index 66c8d0fb7..1188e613c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/cl_weight.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/cl_weight.ma @@ -12,32 +12,34 @@ (* *) (**************************************************************************) -include "basic_2/notation/functions/weight_2.ma". +include "basic_2/notation/functions/weight_3.ma". +include "basic_2/grammar/genv.ma". (**) (* including genv after lenv shows a disambiguation bug: only the last interpretation is considered *) include "basic_2/grammar/lenv_weight.ma". include "basic_2/grammar/cl_shift.ma". (* WEIGHT OF A CLOSURE ******************************************************) -definition fw: lenv → term → ? ≝ λL,T. ♯{L} + ♯{T}. +(* activate genv *) +definition fw: genv → lenv → term → ? ≝ λG,L,T. ♯{L} + ♯{T}. -interpretation "weight (closure)" 'Weight L T = (fw L T). +interpretation "weight (closure)" 'Weight G L T = (fw G L T). (* Basic properties *********************************************************) (* Basic_1: was: flt_shift *) -lemma fw_shift: ∀a,K,I,V,T. ♯{K. ⓑ{I} V, T} < ♯{K, ⓑ{a,I} V. T}. +lemma fw_shift: ∀a,I,G,K,V,T. ♯{G, K.ⓑ{I}V, T} < ♯{G, K, ⓑ{a,I}V.T}. normalize // qed. -lemma fw_tpair_sn: ∀I,L,V,T. ♯{L, V} < ♯{L, ②{I}V.T}. -normalize in ⊢ (?→?→?→?→?%%); // +lemma fw_tpair_sn: ∀I,G,L,V,T. ♯{G, L, V} < ♯{G, L, ②{I}V.T}. +normalize in ⊢ (?→?→?→?→?→?%%); // qed. -lemma fw_tpair_dx: ∀I,L,V,T. ♯{L, T} < ♯{L, ②{I}V.T}. -normalize in ⊢ (?→?→?→?→?%%); // +lemma fw_tpair_dx: ∀I,G,L,V,T. ♯{G, L, T} < ♯{G, L, ②{I}V.T}. +normalize in ⊢ (?→?→?→?→?→?%%); // qed. -lemma fw_lpair_sn: ∀I,L,V,T. ♯{L, V} < ♯{L.ⓑ{I}V, T}. +lemma fw_lpair_sn: ∀I,G,L,V,T. ♯{G, L, V} < ♯{G, L.ⓑ{I}V, T}. normalize /3 width=1 by monotonic_lt_plus_l, monotonic_le_plus_r/ (**) (* auto too slow without trace *) qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/genv.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/genv.ma index 2cf0f6d3d..38757ca16 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/genv.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/genv.ma @@ -38,4 +38,4 @@ interpretation "abstraction (global environment)" (* Basic properties *********************************************************) -axiom genv_eq_dec: ∀T1,T2:genv. Decidable (T1 = T2). +axiom genv_eq_dec: ∀G1,G2:genv. Decidable (G1 = G2). diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/functions/weight_2.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/functions/weight_2.ma deleted file mode 100644 index 161b66bd0..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/functions/weight_2.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 x , break term 46 y } )" - non associative with precedence 90 - for @{ 'Weight $x $y }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/functions/weight_3.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/functions/weight_3.ma new file mode 100644 index 000000000..42b3c60ec --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/functions/weight_3.ma @@ -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 T } )" + non associative with precedence 90 + for @{ 'Weight $G $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/notation.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/notation.ma index d131ad872..bc904846c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/notation.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/notation.ma @@ -14,7 +14,7 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ break ⌘ ⦃ term 46 T ⦄ ≡ break term 46 k )" +notation "hvbox( ⦃G, L⦄ ⊢ break ⌘ ⦃ term 46 T ⦄ ≡ break term 46 k )" non associative with precedence 45 for @{ 'ICM $L $T $k }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/atomicarity_3.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/atomicarity_3.ma index b92527537..42f9d584d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/atomicarity_3.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/atomicarity_3.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ break term 46 T ⁝ break term 46 A )" +notation "hvbox( ⦃G, L⦄ ⊢ break term 46 T ⁝ break term 46 A )" non associative with precedence 45 for @{ 'AtomicArity $L $T $A }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/normal_2.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/normal_2.ma index 519b5a9d2..92cfc51d9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/normal_2.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/normal_2.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ 𝐍 break ⦃ term 46 T ⦄ )" +notation "hvbox( ⦃G, L⦄ ⊢ 𝐍 break ⦃ term 46 T ⦄ )" non associative with precedence 45 for @{ 'Normal $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/notreducible_2.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/notreducible_2.ma index 2048df022..9659a6e23 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/notreducible_2.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/notreducible_2.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ 𝐈 break ⦃ term 46 T ⦄ )" +notation "hvbox( ⦃G, L⦄ ⊢ 𝐈 break ⦃ term 46 T ⦄ )" non associative with precedence 45 for @{ 'NotReducible $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pconv_3.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pconv_3.ma index 374586eb2..c4a789ab8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pconv_3.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pconv_3.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ break term 46 T1 ⬌ break term 46 T2 )" +notation "hvbox( ⦃G, L⦄ ⊢ break term 46 T1 ⬌ break term 46 T2 )" non associative with precedence 45 for @{ 'PConv $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pconvstar_3.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pconvstar_3.ma index 8e604fa53..2088bc68d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pconvstar_3.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pconvstar_3.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ break term 46 T1 ⬌* break term 46 T2 )" +notation "hvbox( ⦃G, L⦄ ⊢ break term 46 T1 ⬌* break term 46 T2 )" non associative with precedence 45 for @{ 'PConvStar $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/peval_3.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/peval_3.ma index 6740f7c03..b2d785e83 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/peval_3.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/peval_3.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ break term 46 T1 ➡ * break 𝐍 ⦃ term 46 T2 ⦄ )" +notation "hvbox( ⦃G, L⦄ ⊢ break term 46 T1 ➡ * break 𝐍 ⦃ term 46 T2 ⦄ )" non associative with precedence 45 for @{ 'PEval $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pred_3.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pred_3.ma index 969791be8..318edbab6 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pred_3.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/pred_3.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ break term 46 T1 ➡ break term 46 T2 )" +notation "hvbox( ⦃G, L⦄ ⊢ break term 46 T1 ➡ break term 46 T2 )" non associative with precedence 45 for @{ 'PRed $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predstar_3.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predstar_3.ma index d6c5062f8..8053f7f7c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predstar_3.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/predstar_3.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ break term 46 T1 ➡ * break term 46 T2 )" +notation "hvbox( ⦃G, L⦄ ⊢ break term 46 T1 ➡ * break term 46 T2 )" non associative with precedence 45 for @{ 'PRedStar $L $T1 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/reducible_2.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/reducible_2.ma index e6844a99f..f5cadce5f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/reducible_2.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/reducible_2.ma @@ -14,6 +14,6 @@ (* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) -notation "hvbox( L ⊢ 𝐑 break ⦃ term 46 T ⦄ )" +notation "hvbox( ⦃G, L⦄ ⊢ 𝐑 break ⦃ term 46 T ⦄ )" non associative with precedence 45 for @{ 'Reducible $L $T }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/supterm_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/supterm_4.ma deleted file mode 100644 index b2114a050..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/supterm_4.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 L1, break term 46 T1 ⦄ ⊃ break ⦃ term 46 L2 , break term 46 T2 ⦄ )" - non associative with precedence 45 - for @{ 'SupTerm $L1 $T1 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/supterm_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/supterm_6.ma new file mode 100644 index 000000000..23c966912 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/supterm_6.ma @@ -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 G1, break term 46 L1, break term 46 T1 ⦄ ⊃ break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" + non associative with precedence 45 + for @{ 'SupTerm $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermopt_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermopt_4.ma deleted file mode 100644 index 3e55a3143..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermopt_4.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 L1, break term 46 T1 ⦄ ⊃⸮ break ⦃ term 46 L2 , break term 46 T2 ⦄ )" - non associative with precedence 45 - for @{ 'SupTermOpt $L1 $T1 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermopt_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermopt_6.ma new file mode 100644 index 000000000..ff1e3b551 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermopt_6.ma @@ -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 G1, break term 46 L1, break term 46 T1 ⦄ ⊃⸮ break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" + non associative with precedence 45 + for @{ 'SupTermOpt $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermoptalt_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermoptalt_4.ma deleted file mode 100644 index 1035ce49a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermoptalt_4.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 L1, break term 46 T1 ⦄ ⊃⊃⸮ break ⦃ term 46 L2 , break term 46 T2 ⦄ )" - non associative with precedence 45 - for @{ 'SupTermOptAlt $L1 $T1 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermoptalt_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermoptalt_6.ma new file mode 100644 index 000000000..f5b529bc3 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermoptalt_6.ma @@ -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 G1, break term 46 L1, break term 46 T1 ⦄ ⊃⊃⸮ break ⦃ term 46 G2, break term 46 L2 , break term 46 T2 ⦄ )" + non associative with precedence 45 + for @{ 'SupTermOptAlt $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermplus_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermplus_4.ma deleted file mode 100644 index 020d64ba3..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermplus_4.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 L1, break term 46 T1 ⦄ ⊃ + break ⦃ term 46 L2 , break term 46 T2 ⦄ )" - non associative with precedence 45 - for @{ 'SupTermPlus $L1 $T1 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermplus_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermplus_6.ma new file mode 100644 index 000000000..558c62e83 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermplus_6.ma @@ -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 G1, term 46 L1, break term 46 T1 ⦄ ⊃ + break ⦃ term 46 G2, term 46 L2 , break term 46 T2 ⦄ )" + non associative with precedence 45 + for @{ 'SupTermPlus $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermstar_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermstar_4.ma deleted file mode 100644 index 53b3bce5e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermstar_4.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 L1, break term 46 T1 ⦄ ⊃ * break ⦃ term 46 L2 , break term 46 T2 ⦄ )" - non associative with precedence 45 - for @{ 'SupTermStar $L1 $T1 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermstar_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermstar_6.ma new file mode 100644 index 000000000..e4c64a6b6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/suptermstar_6.ma @@ -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 G1, term 46 L1, break term 46 T1 ⦄ ⊃ * break ⦃ term 46 G2, term 46 L2 , break term 46 T2 ⦄ )" + non associative with precedence 45 + for @{ 'SupTermStar $G1 $L1 $T1 $G2 $L2 $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma index 0aa31fe72..d9509e17a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir.ma @@ -17,39 +17,39 @@ include "basic_2/reduction/crr.ma". (* CONTEXT-SENSITIVE IRREDUCIBLE TERMS **************************************) -definition cir: lenv → predicate term ≝ λL,T. L ⊢ 𝐑⦃T⦄ → ⊥. +definition cir: lenv → predicate term ≝ λL,T. ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → ⊥. interpretation "context-sensitive irreducibility (term)" 'NotReducible L T = (cir L T). (* Basic inversion lemmas ***************************************************) -lemma cir_inv_delta: ∀L,K,V,i. ⇩[0, i] L ≡ K.ⓓV → L ⊢ 𝐈⦃#i⦄ → ⊥. +lemma cir_inv_delta: ∀L,K,V,i. ⇩[0, i] L ≡ K.ⓓV → ⦃G, L⦄ ⊢ 𝐈⦃#i⦄ → ⊥. /3 width=3/ qed-. -lemma cir_inv_ri2: ∀I,L,V,T. ri2 I → L ⊢ 𝐈⦃②{I}V.T⦄ → ⊥. +lemma cir_inv_ri2: ∀I,L,V,T. ri2 I → ⦃G, L⦄ ⊢ 𝐈⦃②{I}V.T⦄ → ⊥. /3 width=1/ qed-. -lemma cir_inv_ib2: ∀a,I,L,V,T. ib2 a I → L ⊢ 𝐈⦃ⓑ{a,I}V.T⦄ → - L ⊢ 𝐈⦃V⦄ ∧ L.ⓑ{I}V ⊢ 𝐈⦃T⦄. +lemma cir_inv_ib2: ∀a,I,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐈⦃ⓑ{a,I}V.T⦄ → + ⦃G, L⦄ ⊢ 𝐈⦃V⦄ ∧ L.ⓑ{I}V ⊢ 𝐈⦃T⦄. /4 width=1/ qed-. -lemma cir_inv_bind: ∀a,I,L,V,T. L ⊢ 𝐈⦃ⓑ{a,I}V.T⦄ → - ∧∧ L ⊢ 𝐈⦃V⦄ & L.ⓑ{I}V ⊢ 𝐈⦃T⦄ & ib2 a I. +lemma cir_inv_bind: ∀a,I,L,V,T. ⦃G, L⦄ ⊢ 𝐈⦃ⓑ{a,I}V.T⦄ → + ∧∧ ⦃G, 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 (cir_inv_ib2 … H) -H /2 width=1/ /3 width=1/ ] qed-. -lemma cir_inv_appl: ∀L,V,T. L ⊢ 𝐈⦃ⓐV.T⦄ → ∧∧ L ⊢ 𝐈⦃V⦄ & L ⊢ 𝐈⦃T⦄ & 𝐒⦃T⦄. +lemma cir_inv_appl: ∀L,V,T. ⦃G, L⦄ ⊢ 𝐈⦃ⓐV.T⦄ → ∧∧ ⦃G, L⦄ ⊢ 𝐈⦃V⦄ & ⦃G, 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 cir_inv_flat: ∀I,L,V,T. L ⊢ 𝐈⦃ⓕ{I}V.T⦄ → - ∧∧ L ⊢ 𝐈⦃V⦄ & L ⊢ 𝐈⦃T⦄ & 𝐒⦃T⦄ & I = Appl. +lemma cir_inv_flat: ∀I,L,V,T. ⦃G, L⦄ ⊢ 𝐈⦃ⓕ{I}V.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ 𝐈⦃V⦄ & ⦃G, L⦄ ⊢ 𝐈⦃T⦄ & 𝐒⦃T⦄ & I = Appl. * #L #V #T #H [ elim (cir_inv_appl … H) -H /2 width=1/ | elim (cir_inv_ri2 … H) -H /2 width=1/ @@ -58,21 +58,21 @@ qed-. (* Basic properties *********************************************************) -lemma cir_sort: ∀L,k. L ⊢ 𝐈⦃⋆k⦄. +lemma cir_sort: ∀L,k. ⦃G, L⦄ ⊢ 𝐈⦃⋆k⦄. /2 width=3 by crr_inv_sort/ qed. -lemma cir_gref: ∀L,p. L ⊢ 𝐈⦃§p⦄. +lemma cir_gref: ∀L,p. ⦃G, L⦄ ⊢ 𝐈⦃§p⦄. /2 width=3 by crr_inv_gref/ qed. lemma tir_atom: ∀I. ⋆ ⊢ 𝐈⦃⓪{I}⦄. /2 width=2 by trr_inv_atom/ qed. -lemma cir_ib2: ∀a,I,L,V,T. ib2 a I → L ⊢ 𝐈⦃V⦄ → L.ⓑ{I}V ⊢ 𝐈⦃T⦄ → L ⊢ 𝐈⦃ⓑ{a,I}V.T⦄. +lemma cir_ib2: ∀a,I,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐈⦃V⦄ → L.ⓑ{I}V ⊢ 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃ⓑ{a,I}V.T⦄. #a #I #L #V #T #HI #HV #HT #H elim (crr_inv_ib2 … HI H) -HI -H /2 width=1/ qed. -lemma cir_appl: ∀L,V,T. L ⊢ 𝐈⦃V⦄ → L ⊢ 𝐈⦃T⦄ → 𝐒⦃T⦄ → L ⊢ 𝐈⦃ⓐV.T⦄. +lemma cir_appl: ∀L,V,T. ⦃G, L⦄ ⊢ 𝐈⦃V⦄ → ⦃G, L⦄ ⊢ 𝐈⦃T⦄ → 𝐒⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃ⓐV.T⦄. #L #V #T #HV #HT #H1 #H2 elim (crr_inv_appl … H2) -H2 /2 width=1/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_append.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_append.ma index fb88321ef..b91465fc0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_append.ma @@ -19,7 +19,7 @@ include "basic_2/reduction/cir.ma". (* Advanved properties ******************************************************) -lemma cir_labst_last: ∀L,T,W. L ⊢ 𝐈⦃T⦄ → ⋆.ⓛW @@ L ⊢ 𝐈⦃T⦄. +lemma cir_labst_last: ∀L,T,W. ⦃G, L⦄ ⊢ 𝐈⦃T⦄ → ⋆.ⓛW @@ ⦃G, L⦄ ⊢ 𝐈⦃T⦄. /3 width=2 by crr_inv_labst_last/ qed. lemma cir_tif: ∀T,W. ⋆ ⊢ 𝐈⦃T⦄ → ⋆.ⓛW ⊢ 𝐈⦃T⦄. @@ -27,7 +27,7 @@ lemma cir_tif: ∀T,W. ⋆ ⊢ 𝐈⦃T⦄ → ⋆.ⓛW ⊢ 𝐈⦃T⦄. (* Advanced inversion lemmas ************************************************) -lemma cir_inv_append_sn: ∀L,K,T. K @@ L ⊢ 𝐈⦃T⦄ → L ⊢ 𝐈⦃T⦄. +lemma cir_inv_append_sn: ∀L,K,T. K @@ ⦃G, L⦄ ⊢ 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃T⦄. /3 width=1/ qed-. lemma cir_inv_tir: ∀T,W. ⋆.ⓛW ⊢ 𝐈⦃T⦄ → ⋆ ⊢ 𝐈⦃T⦄. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_lift.ma index 9539a6c03..0dcd1cb81 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cir_lift.ma @@ -20,9 +20,9 @@ include "basic_2/reduction/cir.ma". (* Properties on relocation *************************************************) lemma cir_lift: ∀K,T. K ⊢ 𝐈⦃T⦄ → ∀L,d,e. ⇩[d, e] L ≡ K → - ∀U. ⇧[d, e] T ≡ U → L ⊢ 𝐈⦃U⦄. + ∀U. ⇧[d, e] T ≡ U → ⦃G, L⦄ ⊢ 𝐈⦃U⦄. /3 width=7 by crr_inv_lift/ qed. -lemma cir_inv_lift: ∀L,U. L ⊢ 𝐈⦃U⦄ → ∀K,d,e. ⇩[d, e] L ≡ K → +lemma cir_inv_lift: ∀L,U. ⦃G, L⦄ ⊢ 𝐈⦃U⦄ → ∀K,d,e. ⇩[d, e] L ≡ K → ∀T. ⇧[d, e] T ≡ U → K ⊢ 𝐈⦃T⦄. /3 width=7/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma index eec73fb8e..f89a5038d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix.ma @@ -18,43 +18,43 @@ include "basic_2/reduction/crx.ma". (* CONTEXT-SENSITIVE EXTENDED IRREDUCIBLE TERMS *****************************) -definition cix: ∀h. sd h → lenv → predicate term ≝ λh,g,L,T. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → ⊥. +definition cix: ∀h. sd h → lenv → predicate term ≝ λh,g,L,T. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → ⊥. interpretation "context-sensitive extended irreducibility (term)" 'NotReducible h g L T = (cix h g L T). (* Basic inversion lemmas ***************************************************) -lemma cix_inv_sort: ∀h,g,L,k,l. deg h g k (l+1) → ⦃h, L⦄ ⊢ 𝐈[g]⦃⋆k⦄ → ⊥. +lemma cix_inv_sort: ∀h,g,L,k,l. deg h g k (l+1) → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃⋆k⦄ → ⊥. /3 width=2/ qed-. -lemma cix_inv_delta: ∀h,g,I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃h, L⦄ ⊢ 𝐈[g]⦃#i⦄ → ⊥. +lemma cix_inv_delta: ∀h,g,I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃#i⦄ → ⊥. /3 width=4/ qed-. -lemma cix_inv_ri2: ∀h,g,I,L,V,T. ri2 I → ⦃h, L⦄ ⊢ 𝐈[g]⦃②{I}V.T⦄ → ⊥. +lemma cix_inv_ri2: ∀h,g,I,L,V,T. ri2 I → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃②{I}V.T⦄ → ⊥. /3 width=1/ qed-. -lemma cix_inv_ib2: ∀h,g,a,I,L,V,T. ib2 a I → ⦃h, L⦄ ⊢ 𝐈[g]⦃ⓑ{a,I}V.T⦄ → - ⦃h, L⦄ ⊢ 𝐈[g]⦃V⦄ ∧ ⦃h, L.ⓑ{I}V⦄ ⊢ 𝐈[g]⦃T⦄. +lemma cix_inv_ib2: ∀h,g,a,I,L,V,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓑ{a,I}V.T⦄ → + ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ ∧ ⦃h, L.ⓑ{I}V⦄ ⊢ 𝐈[h, g]⦃T⦄. /4 width=1/ qed-. -lemma cix_inv_bind: ∀h,g,a,I,L,V,T. ⦃h, L⦄ ⊢ 𝐈[g]⦃ⓑ{a,I}V.T⦄ → - ∧∧ ⦃h, L⦄ ⊢ 𝐈[g]⦃V⦄ & ⦃h, L.ⓑ{I}V⦄ ⊢ 𝐈[g]⦃T⦄ & ib2 a I. +lemma cix_inv_bind: ∀h,g,a,I,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓑ{a,I}V.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ & ⦃h, L.ⓑ{I}V⦄ ⊢ 𝐈[h, g]⦃T⦄ & ib2 a I. #h #g #a * [ elim a -a ] [ #L #V #T #H elim H -H /3 width=1/ |*: #L #V #T #H elim (cix_inv_ib2 … H) -H /2 width=1/ /3 width=1/ ] qed-. -lemma cix_inv_appl: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐈[g]⦃ⓐV.T⦄ → - ∧∧ ⦃h, L⦄ ⊢ 𝐈[g]⦃V⦄ & ⦃h, L⦄ ⊢ 𝐈[g]⦃T⦄ & 𝐒⦃T⦄. +lemma cix_inv_appl: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓐV.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ & ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ & 𝐒⦃T⦄. #h #g #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 cix_inv_flat: ∀h,g,I,L,V,T. ⦃h, L⦄ ⊢ 𝐈[g]⦃ⓕ{I}V.T⦄ → - ∧∧ ⦃h, L⦄ ⊢ 𝐈[g]⦃V⦄ & ⦃h, L⦄ ⊢ 𝐈[g]⦃T⦄ & 𝐒⦃T⦄ & I = Appl. +lemma cix_inv_flat: ∀h,g,I,L,V,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃ⓕ{I}V.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ 𝐈[h, g]⦃V⦄ & ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ & 𝐒⦃T⦄ & I = Appl. #h #g * #L #V #T #H [ elim (cix_inv_appl … H) -H /2 width=1/ | elim (cix_inv_ri2 … H) -H /2 width=1/ @@ -63,31 +63,31 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma cix_inv_cir: ∀h,g,L,T. ⦃h, L⦄ ⊢ 𝐈[g]⦃T⦄ → L ⊢ 𝐈⦃T⦄. +lemma cix_inv_cir: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃T⦄. /3 width=1/ qed-. (* Basic properties *********************************************************) -lemma cix_sort: ∀h,g,L,k. deg h g k 0 → ⦃h, L⦄ ⊢ 𝐈[g]⦃⋆k⦄. +lemma cix_sort: ∀h,g,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃⋆k⦄. #h #g #L #k #Hk #H elim (crx_inv_sort … H) -L #l #Hkl lapply (deg_mono … Hk Hkl) -h -k (cpr_inv_sort1 … H) // qed. (* Basic_1: was: nf2_abst *) -lemma cnr_abst: ∀a,L,W,T. L ⊢ 𝐍⦃W⦄ → L.ⓛW ⊢ 𝐍⦃T⦄ → L ⊢ 𝐍⦃ⓛ{a}W.T⦄. +lemma cnr_abst: ∀a,L,W,T. ⦃G, L⦄ ⊢ 𝐍⦃W⦄ → L.ⓛW ⊢ 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍⦃ⓛ{a}W.T⦄. #a #L #W #T #HW #HT #X #H elim (cpr_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct >(HW … HW0) -W0 >(HT … HT0) -T0 // qed. (* Basic_1: was only: nf2_appl_lref *) -lemma cnr_appl_simple: ∀L,V,T. L ⊢ 𝐍⦃V⦄ → L ⊢ 𝐍⦃T⦄ → 𝐒⦃T⦄ → L ⊢ 𝐍⦃ⓐV.T⦄. +lemma cnr_appl_simple: ∀L,V,T. ⦃G, L⦄ ⊢ 𝐍⦃V⦄ → ⦃G, L⦄ ⊢ 𝐍⦃T⦄ → 𝐒⦃T⦄ → ⦃G, 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 cnr_dec: ∀L,T1. L ⊢ 𝐍⦃T1⦄ ∨ - ∃∃T2. L ⊢ T1 ➡ T2 & (T1 = T2 → ⊥). +axiom cnr_dec: ∀L,T1. ⦃G, L⦄ ⊢ 𝐍⦃T1⦄ ∨ + ∃∃T2. ⦃G, L⦄ ⊢ T1 ➡ T2 & (T1 = T2 → ⊥). (* Basic_1: removed theorems 1: nf2_abst_shift *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_cir.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_cir.ma index 50f30ab60..269cd6561 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_cir.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_cir.ma @@ -19,10 +19,10 @@ include "basic_2/reduction/cnr_crr.ma". (* Main properties on context-sensitive irreducible terms *******************) -theorem cir_cnr: ∀L,T. L ⊢ 𝐈⦃T⦄ → L ⊢ 𝐍⦃T⦄. +theorem cir_cnr: ∀L,T. ⦃G, L⦄ ⊢ 𝐈⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍⦃T⦄. /2 width=3 by cpr_fwd_cir/ qed. (* Main inversion lemmas on context-sensitive irreducible terms *************) -theorem cnr_inv_cir: ∀L,T. L ⊢ 𝐍⦃T⦄ → L ⊢ 𝐈⦃T⦄. +theorem cnr_inv_cir: ∀L,T. ⦃G, L⦄ ⊢ 𝐍⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈⦃T⦄. /2 width=4 by cnr_inv_crr/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_crr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_crr.ma index 6d68cf9e0..c52165e3e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_crr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_crr.ma @@ -20,7 +20,7 @@ include "basic_2/reduction/cnr.ma". (* Advanced inversion lemmas on context-sensitive reducible terms ***********) (* Note: this property is unusual *) -lemma cnr_inv_crr: ∀L,T. L ⊢ 𝐑⦃T⦄ → L ⊢ 𝐍⦃T⦄ → ⊥. +lemma cnr_inv_crr: ∀L,T. ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍⦃T⦄ → ⊥. #L #T #H elim H -L -T [ #L #K #V #i #HLK #H elim (cnr_inv_delta … HLK H) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_lift.ma index 5e3c2c419..8d0267c47 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnr_lift.ma @@ -20,7 +20,7 @@ include "basic_2/reduction/cnr.ma". (* Advanced properties ******************************************************) (* Basic_1: was only: nf2_csort_lref *) -lemma cnr_lref_atom: ∀L,i. ⇩[0, i] L ≡ ⋆ → L ⊢ 𝐍⦃#i⦄. +lemma cnr_lref_atom: ∀L,i. ⇩[0, i] L ≡ ⋆ → ⦃G, L⦄ ⊢ 𝐍⦃#i⦄. #L #i #HL #X #H elim (cpr_inv_lref1 … H) -H // * #K #V1 #V2 #HLK #_ #_ @@ -28,7 +28,7 @@ lapply (ldrop_mono … HL … HLK) -L #H destruct qed. (* Basic_1: was: nf2_lref_abst *) -lemma cnr_lref_abst: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓛV → L ⊢ 𝐍⦃#i⦄. +lemma cnr_lref_abst: ∀L,K,V,i. ⇩[0, i] L ≡ K. ⓛV → ⦃G, L⦄ ⊢ 𝐍⦃#i⦄. #L #K #V #i #HLK #X #H elim (cpr_inv_lref1 … H) -H // * #K0 #V1 #V2 #HLK0 #_ #_ @@ -39,7 +39,7 @@ qed. (* Basic_1: was: nf2_lift *) lemma cnr_lift: ∀L0,L,T,T0,d,e. - L ⊢ 𝐍⦃T⦄ → ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → L0 ⊢ 𝐍⦃T0⦄. + ⦃G, 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 @@ -48,7 +48,7 @@ qed. (* Note: this was missing in basic_1 *) lemma cnr_inv_lift: ∀L0,L,T,T0,d,e. - L0 ⊢ 𝐍⦃T0⦄ → ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → L ⊢ 𝐍⦃T⦄. + L0 ⊢ 𝐍⦃T0⦄ → ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → ⦃G, L⦄ ⊢ 𝐍⦃T⦄. #L0 #L #T #T0 #d #e #HLT0 #HL0 #HT0 #X #H elim (lift_total X d e) #X0 #HX0 lapply (cpr_lift … H … HL0 … HT0 … HX0) -L #HTX0 diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma index 8e6dd66ba..5c97081f5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma @@ -27,37 +27,37 @@ interpretation (* Basic inversion lemmas ***************************************************) -lemma cnx_inv_sort: ∀h,g,L,k. ⦃h, L⦄ ⊢ 𝐍[g]⦃⋆k⦄ → deg h g k 0. +lemma cnx_inv_sort: ∀h,g,L,k. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃⋆k⦄ → deg h g k 0. #h #g #L #k #H elim (deg_total h g k) #l @(nat_ind_plus … l) -l // #l #_ #Hkl lapply (H (⋆(next h k)) ?) -H /2 width=2/ -L -l #H destruct -H -e0 (**) (* destruct does not remove some premises *) lapply (next_lt h k) >e1 -e1 #H elim (lt_refl_false … H) qed-. -lemma cnx_inv_delta: ∀h,g,I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃h, L⦄ ⊢ 𝐍[g]⦃#i⦄ → ⊥. +lemma cnx_inv_delta: ∀h,g,I,L,K,V,i. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃#i⦄ → ⊥. #h #g #I #L #K #V #i #HLK #H elim (lift_total V 0 (i+1)) #W #HVW lapply (H W ?) -H [ /3 width=7/ ] -HLK #H destruct elim (lift_inv_lref2_be … HVW) -HVW // qed-. -lemma cnx_inv_abst: ∀h,g,a,L,V,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃ⓛ{a}V.T⦄ → - ⦃h, L⦄ ⊢ 𝐍[g]⦃V⦄ ∧ ⦃h, L.ⓛV⦄ ⊢ 𝐍[g]⦃T⦄. +lemma cnx_inv_abst: ∀h,g,a,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃ⓛ{a}V.T⦄ → + ⦃G, L⦄ ⊢ 𝐍[h, g]⦃V⦄ ∧ ⦃h, L.ⓛV⦄ ⊢ 𝐍[h, g]⦃T⦄. #h #g #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 cnx_inv_abbr: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃-ⓓV.T⦄ → - ⦃h, L⦄ ⊢ 𝐍[g]⦃V⦄ ∧ ⦃h, L.ⓓV⦄ ⊢ 𝐍[g]⦃T⦄. +lemma cnx_inv_abbr: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃-ⓓV.T⦄ → + ⦃G, L⦄ ⊢ 𝐍[h, g]⦃V⦄ ∧ ⦃h, L.ⓓV⦄ ⊢ 𝐍[h, g]⦃T⦄. #h #g #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-. -lemma cnx_inv_zeta: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃+ⓓV.T⦄ → ⊥. +lemma cnx_inv_zeta: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃+ⓓV.T⦄ → ⊥. #h #g #L #V #T #H elim (is_lift_dec T 0 1) [ * #U #HTU lapply (H U ?) -H /2 width=3/ #H destruct @@ -68,8 +68,8 @@ lemma cnx_inv_zeta: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃+ⓓV.T⦄ → ⊥. ] qed-. -lemma cnx_inv_appl: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃ⓐV.T⦄ → - ∧∧ ⦃h, L⦄ ⊢ 𝐍[g]⦃V⦄ & ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ & 𝐒⦃T⦄. +lemma cnx_inv_appl: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃ⓐV.T⦄ → + ∧∧ ⦃G, L⦄ ⊢ 𝐍[h, g]⦃V⦄ & ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ & 𝐒⦃T⦄. #h #g #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 // @@ -80,43 +80,43 @@ lemma cnx_inv_appl: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃ⓐV.T⦄ → ] qed-. -lemma cnx_inv_tau: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃ⓝV.T⦄ → ⊥. +lemma cnx_inv_tau: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃ⓝV.T⦄ → ⊥. #h #g #L #V #T #H lapply (H T ?) -H /2 width=1/ #H @discr_tpair_xy_y // qed-. (* Basic forward lemmas *****************************************************) -lemma cnx_fwd_cnr: ∀h,g,L,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → L ⊢ 𝐍⦃T⦄. +lemma cnx_fwd_cnr: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍⦃T⦄. #h #g #L #T #H #U #HTU @H /2 width=1/ (**) (* auto fails because a δ-expansion gets in the way *) qed-. (* Basic properties *********************************************************) -lemma cnx_sort: ∀h,g,L,k. deg h g k 0 → ⦃h, L⦄ ⊢ 𝐍[g]⦃⋆k⦄. +lemma cnx_sort: ∀h,g,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃⋆k⦄. #h #g #L #k #Hk #X #H elim (cpx_inv_sort1 … H) -H // * #l #Hkl #_ lapply (deg_mono … Hkl Hk) -h -L (HW … HW0) -W0 >(HT … HT0) -T0 // qed. -lemma cnx_appl_simple: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃V⦄ → ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → 𝐒⦃T⦄ → - ⦃h, L⦄ ⊢ 𝐍[g]⦃ⓐV.T⦄. +lemma cnx_appl_simple: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃V⦄ → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → 𝐒⦃T⦄ → + ⦃G, L⦄ ⊢ 𝐍[h, g]⦃ⓐV.T⦄. #h #g #L #V #T #HV #HT #HS #X #H elim (cpx_inv_appl1_simple … H) -H // #V0 #T0 #HV0 #HT0 #H destruct >(HV … HV0) -V0 >(HT … HT0) -T0 // qed. -axiom cnx_dec: ∀h,g,L,T1. ⦃h, L⦄ ⊢ 𝐍[g]⦃T1⦄ ∨ - ∃∃T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 & (T1 = T2 → ⊥). +axiom cnx_dec: ∀h,g,L,T1. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T1⦄ ∨ + ∃∃T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 & (T1 = T2 → ⊥). diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_cix.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_cix.ma index 05f1acedb..f45801a25 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_cix.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_cix.ma @@ -19,10 +19,10 @@ include "basic_2/reduction/cnx_crx.ma". (* Main properties on context-sensitive extended irreducible terms **********) -theorem cix_cnr: ∀h,g,L,T. ⦃h, L⦄ ⊢ 𝐈[g]⦃T⦄ → ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄. +theorem cix_cnr: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄. /2 width=5 by cpx_fwd_cix/ qed. (* Main inversion lemmas on context-sensitive extended irreducible terms ****) -theorem cnr_inv_cix: ∀h,g,L,T. ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → ⦃h, L⦄ ⊢ 𝐈[g]⦃T⦄. +theorem cnr_inv_cix: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T⦄. /2 width=6 by cnx_inv_crx/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_crx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_crx.ma index 99ca32c13..84951db34 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_crx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_crx.ma @@ -20,7 +20,7 @@ include "basic_2/reduction/cnx.ma". (* Advanced inversion lemmas on context-sensitive reducible terms ***********) (* Note: this property is unusual *) -lemma cnx_inv_crx: ∀h,g,L,T. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → ⊥. +lemma cnx_inv_crx: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → ⊥. #h #g #L #T #H elim H -L -T [ #L #k #l #Hkl #H lapply (cnx_inv_sort … H) -H #H diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma index 750d23995..d26c41cb7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma @@ -19,7 +19,7 @@ include "basic_2/reduction/cnx.ma". (* Advanced properties ******************************************************) -lemma cnx_lref_atom: ∀h,g,L,i. ⇩[0, i] L ≡ ⋆ → ⦃h, L⦄ ⊢ 𝐍[g]⦃#i⦄. +lemma cnx_lref_atom: ∀h,g,L,i. ⇩[0, i] L ≡ ⋆ → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃#i⦄. #h #g #L #i #HL #X #H elim (cpx_inv_lref1 … H) -H // * #I #K #V1 #V2 #HLK #_ #_ @@ -28,16 +28,16 @@ qed. (* Relocation properties ****************************************************) -lemma cnx_lift: ∀h,g,L0,L,T,T0,d,e. ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄ → ⇩[d, e] L0 ≡ L → - ⇧[d, e] T ≡ T0 → ⦃h, L0⦄ ⊢ 𝐍[g]⦃T0⦄. +lemma cnx_lift: ∀h,g,L0,L,T,T0,d,e. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → ⇩[d, e] L0 ≡ L → + ⇧[d, e] T ≡ T0 → ⦃h, L0⦄ ⊢ 𝐍[h, g]⦃T0⦄. #h #g #L0 #L #T #T0 #d #e #HLT #HL0 #HT0 #X #H elim (cpx_inv_lift1 … H … HL0 … HT0) -L0 #T1 #HT10 #HT1 <(HLT … HT1) in HT0; -L #HT0 >(lift_mono … HT10 … HT0) -T1 -X // qed. -lemma cnx_inv_lift: ∀h,g,L0,L,T,T0,d,e. ⦃h, L0⦄ ⊢ 𝐍[g]⦃T0⦄ → ⇩[d, e] L0 ≡ L → - ⇧[d, e] T ≡ T0 → ⦃h, L⦄ ⊢ 𝐍[g]⦃T⦄. +lemma cnx_inv_lift: ∀h,g,L0,L,T,T0,d,e. ⦃h, L0⦄ ⊢ 𝐍[h, g]⦃T0⦄ → ⇩[d, e] L0 ≡ L → + ⇧[d, e] T ≡ T0 → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄. #h #g #L0 #L #T #T0 #d #e #HLT0 #HL0 #HT0 #X #H elim (lift_total X d e) #X0 #HX0 lapply (cpx_lift … H … HL0 … HT0 … HX0) -L #HTX0 diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma index 2f263f76c..69b591178 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma @@ -60,23 +60,23 @@ lemma lsubr_cpr_trans: lsub_trans … cpr lsubr. qed-. (* Basic_1: was by definition: pr2_free *) -lemma tpr_cpr: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2. +lemma tpr_cpr: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → ∀L. ⦃G, L⦄ ⊢ T1 ➡ T2. #T1 #T2 #HT12 #L lapply (lsubr_cpr_trans … HT12 L ?) // qed. (* Basic_1: includes by definition: pr0_refl *) -lemma cpr_refl: ∀T,L. L ⊢ T ➡ T. +lemma cpr_refl: ∀T,L. ⦃G, L⦄ ⊢ T ➡ T. #T elim T -T // * /2 width=1/ qed. (* Basic_1: was: pr2_head_1 *) -lemma cpr_pair_sn: ∀I,L,V1,V2. L ⊢ V1 ➡ V2 → - ∀T. L ⊢ ②{I}V1.T ➡ ②{I}V2.T. +lemma cpr_pair_sn: ∀I,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → + ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡ ②{I}V2.T. * /2 width=1/ qed. lemma cpr_delift: ∀K,V,T1,L,d. ⇩[0, d] L ≡ (K.ⓓV) → - ∃∃T2,T. L ⊢ T1 ➡ T2 & ⇧[d, 1] T ≡ T2. + ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡ T2 & ⇧[d, 1] T ≡ T2. #K #V #T1 elim T1 -T1 [ * #i #L #d #HLK /2 width=4/ elim (lt_or_eq_or_gt i d) #Hid [1,3: /3 width=4/ ] @@ -101,7 +101,7 @@ qed. (* Basic inversion lemmas ***************************************************) -fact cpr_inv_atom1_aux: ∀L,T1,T2. L ⊢ T1 ➡ T2 → ∀I. T1 = ⓪{I} → +fact cpr_inv_atom1_aux: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ∀I. T1 = ⓪{I} → T2 = ⓪{I} ∨ ∃∃K,V,V2,i. ⇩[O, i] L ≡ K. ⓓV & K ⊢ V ➡ V2 & @@ -119,7 +119,7 @@ fact cpr_inv_atom1_aux: ∀L,T1,T2. L ⊢ T1 ➡ T2 → ∀I. T1 = ⓪{I} → ] qed-. -lemma cpr_inv_atom1: ∀I,L,T2. L ⊢ ⓪{I} ➡ T2 → +lemma cpr_inv_atom1: ∀I,L,T2. ⦃G, L⦄ ⊢ ⓪{I} ➡ T2 → T2 = ⓪{I} ∨ ∃∃K,V,V2,i. ⇩[O, i] L ≡ K. ⓓV & K ⊢ V ➡ V2 & @@ -128,14 +128,14 @@ lemma cpr_inv_atom1: ∀I,L,T2. L ⊢ ⓪{I} ➡ T2 → /2 width=3 by cpr_inv_atom1_aux/ qed-. (* Basic_1: includes: pr0_gen_sort pr2_gen_sort *) -lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k. +lemma cpr_inv_sort1: ∀L,T2,k. ⦃G, L⦄ ⊢ ⋆k ➡ T2 → T2 = ⋆k. #L #T2 #k #H elim (cpr_inv_atom1 … H) -H // * #K #V #V2 #i #_ #_ #_ #H destruct qed-. (* Basic_1: includes: pr0_gen_lref pr2_gen_lref *) -lemma cpr_inv_lref1: ∀L,T2,i. L ⊢ #i ➡ T2 → +lemma cpr_inv_lref1: ∀L,T2,i. ⦃G, L⦄ ⊢ #i ➡ T2 → T2 = #i ∨ ∃∃K,V,V2. ⇩[O, i] L ≡ K. ⓓV & K ⊢ V ➡ V2 & @@ -145,15 +145,15 @@ elim (cpr_inv_atom1 … H) -H /2 width=1/ * #K #V #V2 #j #HLK #HV2 #HVT2 #H destruct /3 width=6/ qed-. -lemma cpr_inv_gref1: ∀L,T2,p. L ⊢ §p ➡ T2 → T2 = §p. +lemma cpr_inv_gref1: ∀L,T2,p. ⦃G, L⦄ ⊢ §p ➡ T2 → T2 = §p. #L #T2 #p #H elim (cpr_inv_atom1 … H) -H // * #K #V #V2 #i #_ #_ #_ #H destruct qed-. -fact cpr_inv_bind1_aux: ∀L,U1,U2. L ⊢ U1 ➡ U2 → +fact cpr_inv_bind1_aux: ∀L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡ U2 → ∀a,I,V1,T1. U1 = ⓑ{a,I}V1. T1 → ( - ∃∃V2,T2. L ⊢ V1 ➡ V2 & + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & L. ⓑ{I}V1 ⊢ T1 ➡ T2 & U2 = ⓑ{a,I}V2.T2 ) ∨ @@ -170,8 +170,8 @@ fact cpr_inv_bind1_aux: ∀L,U1,U2. L ⊢ U1 ➡ U2 → ] qed-. -lemma cpr_inv_bind1: ∀a,I,L,V1,T1,U2. L ⊢ ⓑ{a,I}V1.T1 ➡ U2 → ( - ∃∃V2,T2. L ⊢ V1 ➡ V2 & +lemma cpr_inv_bind1: ∀a,I,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡ U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & L. ⓑ{I}V1 ⊢ T1 ➡ T2 & U2 = ⓑ{a,I}V2.T2 ) ∨ @@ -179,8 +179,8 @@ lemma cpr_inv_bind1: ∀a,I,L,V1,T1,U2. L ⊢ ⓑ{a,I}V1.T1 ➡ U2 → ( /2 width=3 by cpr_inv_bind1_aux/ qed-. (* Basic_1: includes: pr0_gen_abbr pr2_gen_abbr *) -lemma cpr_inv_abbr1: ∀a,L,V1,T1,U2. L ⊢ ⓓ{a}V1.T1 ➡ U2 → ( - ∃∃V2,T2. L ⊢ V1 ➡ V2 & +lemma cpr_inv_abbr1: ∀a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡ U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & L. ⓓV1 ⊢ T1 ➡ T2 & U2 = ⓓ{a}V2.T2 ) ∨ @@ -190,8 +190,8 @@ elim (cpr_inv_bind1 … H) -H * /3 width=3/ /3 width=5/ qed-. (* Basic_1: includes: pr0_gen_abst pr2_gen_abst *) -lemma cpr_inv_abst1: ∀a,L,V1,T1,U2. L ⊢ ⓛ{a}V1.T1 ➡ U2 → - ∃∃V2,T2. L ⊢ V1 ➡ V2 & L.ⓛV1 ⊢ T1 ➡ T2 & +lemma cpr_inv_abst1: ∀a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡ U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & L.ⓛV1 ⊢ T1 ➡ T2 & U2 = ⓛ{a}V2.T2. #a #L #V1 #T1 #U2 #H elim (cpr_inv_bind1 … H) -H * @@ -200,16 +200,16 @@ elim (cpr_inv_bind1 … H) -H * ] qed-. -fact cpr_inv_flat1_aux: ∀L,U,U2. L ⊢ U ➡ U2 → +fact cpr_inv_flat1_aux: ∀L,U,U2. ⦃G, L⦄ ⊢ U ➡ U2 → ∀I,V1,U1. U = ⓕ{I}V1.U1 → - ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U1 ➡ T2 & + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 & U2 = ⓕ{I} V2. T2 - | (L ⊢ U1 ➡ U2 ∧ I = Cast) - | ∃∃a,V2,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & + | (⦃G, L⦄ ⊢ U1 ➡ U2 ∧ I = Cast) + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 & L.ⓛW1 ⊢ T1 ➡ T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl - | ∃∃a,V,V2,W1,W2,T1,T2. L ⊢ V1 ➡ V & ⇧[0,1] V ≡ V2 & - L ⊢ W1 ➡ W2 & L.ⓓW1 ⊢ T1 ➡ T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⇧[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡ W2 & L.ⓓW1 ⊢ T1 ➡ T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl. #L #U #U2 * -L -U -U2 @@ -224,28 +224,28 @@ fact cpr_inv_flat1_aux: ∀L,U,U2. L ⊢ U ➡ U2 → ] qed-. -lemma cpr_inv_flat1: ∀I,L,V1,U1,U2. L ⊢ ⓕ{I}V1.U1 ➡ U2 → - ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U1 ➡ T2 & +lemma cpr_inv_flat1: ∀I,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡ U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 & U2 = ⓕ{I}V2.T2 - | (L ⊢ U1 ➡ U2 ∧ I = Cast) - | ∃∃a,V2,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & + | (⦃G, L⦄ ⊢ U1 ➡ U2 ∧ I = Cast) + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 & L.ⓛW1 ⊢ T1 ➡ T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl - | ∃∃a,V,V2,W1,W2,T1,T2. L ⊢ V1 ➡ V & ⇧[0,1] V ≡ V2 & - L ⊢ W1 ➡ W2 & L.ⓓW1 ⊢ T1 ➡ T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⇧[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡ W2 & L.ⓓW1 ⊢ T1 ➡ T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl. /2 width=3 by cpr_inv_flat1_aux/ qed-. (* Basic_1: includes: pr0_gen_appl pr2_gen_appl *) -lemma cpr_inv_appl1: ∀L,V1,U1,U2. L ⊢ ⓐV1.U1 ➡ U2 → - ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U1 ➡ T2 & +lemma cpr_inv_appl1: ∀L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐV1.U1 ➡ U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 & U2 = ⓐV2.T2 - | ∃∃a,V2,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ W1 ➡ W2 & L.ⓛW1 ⊢ T1 ➡ T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 - | ∃∃a,V,V2,W1,W2,T1,T2. L ⊢ V1 ➡ V & ⇧[0,1] V ≡ V2 & - L ⊢ W1 ➡ W2 & L.ⓓW1 ⊢ T1 ➡ T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡ V & ⇧[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡ W2 & L.ⓓW1 ⊢ T1 ➡ T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2. #L #V1 #U1 #U2 #H elim (cpr_inv_flat1 … H) -H * [ /3 width=5/ @@ -256,8 +256,8 @@ lemma cpr_inv_appl1: ∀L,V1,U1,U2. L ⊢ ⓐV1.U1 ➡ U2 → qed-. (* Note: the main property of simple terms *) -lemma cpr_inv_appl1_simple: ∀L,V1,T1,U. L ⊢ ⓐV1. T1 ➡ U → 𝐒⦃T1⦄ → - ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 & +lemma cpr_inv_appl1_simple: ∀L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1. T1 ➡ U → 𝐒⦃T1⦄ → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ T1 ➡ T2 & U = ⓐV2. T2. #L #V1 #T1 #U #H #HT1 elim (cpr_inv_appl1 … H) -H * @@ -270,11 +270,11 @@ elim (cpr_inv_appl1 … H) -H * qed-. (* Basic_1: includes: pr0_gen_cast pr2_gen_cast *) -lemma cpr_inv_cast1: ∀L,V1,U1,U2. L ⊢ ⓝ V1. U1 ➡ U2 → ( - ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U1 ➡ T2 & +lemma cpr_inv_cast1: ∀L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝ V1. U1 ➡ U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡ V2 & ⦃G, L⦄ ⊢ U1 ➡ T2 & U2 = ⓝ V2. T2 ) ∨ - L ⊢ U1 ➡ U2. + ⦃G, L⦄ ⊢ U1 ➡ U2. #L #V1 #U1 #U2 #H elim (cpr_inv_flat1 … H) -H * [ /3 width=5/ | /2 width=1/ @@ -285,8 +285,8 @@ 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 & +lemma cpr_fwd_bind1_minus: ∀I,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡ T → ∀b. + ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡ ⓑ{b,I}V2.T2 & T = -ⓑ{I}V2.T2. #I #L #V1 #T1 #T #H #b elim (cpr_inv_bind1 … H) -H * @@ -295,7 +295,7 @@ elim (cpr_inv_bind1 … H) -H * ] qed-. -lemma cpr_fwd_shift1: ∀L1,L,T1,T. L ⊢ L1 @@ T1 ➡ T → +lemma cpr_fwd_shift1: ∀L1,L,T1,T. ⦃G, L⦄ ⊢ L1 @@ T1 ➡ T → ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. #L1 @(lenv_ind_dx … L1) -L1 normalize [ #L #T1 #T #HT1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cir.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cir.ma index ee49679b7..a54e9c420 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cir.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cir.ma @@ -19,7 +19,7 @@ include "basic_2/reduction/cpr.ma". (* Advanced forward lemmas on context-sensitive irreducible terms ***********) -lemma cpr_fwd_cir: ∀L,T1,T2. L ⊢ T1 ➡ T2 → L ⊢ 𝐈⦃T1⦄ → T2 = T1. +lemma cpr_fwd_cir: ∀L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ 𝐈⦃T1⦄ → T2 = T1. #L #T1 #T2 #H elim H -L -T1 -T2 [ // | #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #H diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma index 6e1a50a65..6225e896b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx.ma @@ -63,28 +63,28 @@ lemma lsubr_cpx_trans: ∀h,g. lsub_trans … (cpx h g) lsubr. qed-. (* Note: this is "∀h,g,L. reflexive … (cpx h g L)" *) -lemma cpx_refl: ∀h,g,T,L. ⦃h, L⦄ ⊢ T ➡[g] T. +lemma cpx_refl: ∀h,g,T,L. ⦃G, L⦄ ⊢ T ➡[h, g] T. #h #g #T elim T -T // * /2 width=1/ qed. -lemma cpr_cpx: ∀h,g,L,T1,T2. L ⊢ T1 ➡ T2 → ⦃h, L⦄ ⊢ T1 ➡[g] T2. +lemma cpr_cpx: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2. #h #g #L #T1 #T2 #H elim H -L -T1 -T2 // /2 width=1/ /2 width=3/ /2 width=7/ qed. -fact ssta_cpx_aux: ∀h,g,L,T1,T2,l0. ⦃h, L⦄ ⊢ T1 •[g] ⦃l0, T2⦄ → - ∀l. l0 = l+1 → ⦃h, L⦄ ⊢ T1 ➡[g] T2. +fact ssta_cpx_aux: ∀h,g,L,T1,T2,l0. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l0, T2⦄ → + ∀l. l0 = l+1 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2. #h #g #L #T1 #T2 #l0 #H elim H -L -T1 -T2 -l0 /2 width=2/ /2 width=7/ /3 width=2/ /3 width=7/ qed-. -lemma ssta_cpx: ∀h,g,L,T1,T2,l. ⦃h, L⦄ ⊢ T1 •[g] ⦃l+1, T2⦄ → ⦃h, L⦄ ⊢ T1 ➡[g] T2. +lemma ssta_cpx: ∀h,g,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •[h, g] ⦃l+1, T2⦄ → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2. /2 width=4 by ssta_cpx_aux/ qed. -lemma cpx_pair_sn: ∀h,g,I,L,V1,V2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 → - ∀T. ⦃h, L⦄ ⊢ ②{I}V1.T ➡[g] ②{I}V2.T. +lemma cpx_pair_sn: ∀h,g,I,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 → + ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡[h, g] ②{I}V2.T. #h #g * /2 width=1/ qed. lemma cpx_delift: ∀h,g,I,K,V,T1,L,d. ⇩[0, d] L ≡ (K.ⓑ{I}V) → - ∃∃T2,T. ⦃h, L⦄ ⊢ T1 ➡[g] T2 & ⇧[d, 1] T ≡ T2. + ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 & ⇧[d, 1] T ≡ T2. #h #g #I #K #V #T1 elim T1 -T1 [ * #i #L #d #HLK /2 width=4/ elim (lt_or_eq_or_gt i d) #Hid [1,3: /3 width=4/ ] @@ -109,10 +109,10 @@ qed. (* Basic inversion lemmas ***************************************************) -fact cpx_inv_atom1_aux: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → ∀J. T1 = ⓪{J} → +fact cpx_inv_atom1_aux: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ∀J. T1 = ⓪{J} → ∨∨ T2 = ⓪{J} | ∃∃k,l. deg h g k (l+1) & T2 = ⋆(next h k) & J = Sort k - | ∃∃I,K,V,V2,i. ⇩[O, i] L ≡ K.ⓑ{I}V & ⦃h, K⦄ ⊢ V ➡[g] V2 & + | ∃∃I,K,V,V2,i. ⇩[O, i] L ≡ K.ⓑ{I}V & ⦃h, K⦄ ⊢ V ➡[h, g] V2 & ⇧[O, i + 1] V2 ≡ T2 & J = LRef i. #h #g #L #T1 #T2 * -L -T1 -T2 [ #I #L #J #H destruct /2 width=1/ @@ -128,14 +128,14 @@ fact cpx_inv_atom1_aux: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → ∀J. T1 ] qed-. -lemma cpx_inv_atom1: ∀h,g,J,L,T2. ⦃h, L⦄ ⊢ ⓪{J} ➡[g] T2 → +lemma cpx_inv_atom1: ∀h,g,J,L,T2. ⦃G, L⦄ ⊢ ⓪{J} ➡[h, g] T2 → ∨∨ T2 = ⓪{J} | ∃∃k,l. deg h g k (l+1) & T2 = ⋆(next h k) & J = Sort k - | ∃∃I,K,V,V2,i. ⇩[O, i] L ≡ K.ⓑ{I}V & ⦃h, K⦄ ⊢ V ➡[g] V2 & + | ∃∃I,K,V,V2,i. ⇩[O, i] L ≡ K.ⓑ{I}V & ⦃h, K⦄ ⊢ V ➡[h, g] V2 & ⇧[O, i + 1] V2 ≡ T2 & J = LRef i. /2 width=3 by cpx_inv_atom1_aux/ qed-. -lemma cpx_inv_sort1: ∀h,g,L,T2,k. ⦃h, L⦄ ⊢ ⋆k ➡[g] T2 → T2 = ⋆k ∨ +lemma cpx_inv_sort1: ∀h,g,L,T2,k. ⦃G, L⦄ ⊢ ⋆k ➡[h, g] T2 → T2 = ⋆k ∨ ∃∃l. deg h g k (l+1) & T2 = ⋆(next h k). #h #g #L #T2 #k #H elim (cpx_inv_atom1 … H) -H /2 width=1/ * @@ -144,9 +144,9 @@ elim (cpx_inv_atom1 … H) -H /2 width=1/ * ] qed-. -lemma cpx_inv_lref1: ∀h,g,L,T2,i. ⦃h, L⦄ ⊢ #i ➡[g] T2 → +lemma cpx_inv_lref1: ∀h,g,L,T2,i. ⦃G, L⦄ ⊢ #i ➡[h, g] T2 → T2 = #i ∨ - ∃∃I,K,V,V2. ⇩[O, i] L ≡ K. ⓑ{I}V & ⦃h, K⦄ ⊢ V ➡[g] V2 & + ∃∃I,K,V,V2. ⇩[O, i] L ≡ K. ⓑ{I}V & ⦃h, K⦄ ⊢ V ➡[h, g] V2 & ⇧[O, i + 1] V2 ≡ T2. #h #g #L #T2 #i #H elim (cpx_inv_atom1 … H) -H /2 width=1/ * @@ -155,7 +155,7 @@ elim (cpx_inv_atom1 … H) -H /2 width=1/ * ] qed-. -lemma cpx_inv_gref1: ∀h,g,L,T2,p. ⦃h, L⦄ ⊢ §p ➡[g] T2 → T2 = §p. +lemma cpx_inv_gref1: ∀h,g,L,T2,p. ⦃G, L⦄ ⊢ §p ➡[h, g] T2 → T2 = §p. #h #g #L #T2 #p #H elim (cpx_inv_atom1 … H) -H // * [ #k #l #_ #_ #H destruct @@ -163,12 +163,12 @@ elim (cpx_inv_atom1 … H) -H // * ] qed-. -fact cpx_inv_bind1_aux: ∀h,g,L,U1,U2. ⦃h, L⦄ ⊢ U1 ➡[g] U2 → +fact cpx_inv_bind1_aux: ∀h,g,L,U1,U2. ⦃G, L⦄ ⊢ U1 ➡[h, g] U2 → ∀a,J,V1,T1. U1 = ⓑ{a,J}V1.T1 → ( - ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L.ⓑ{J}V1⦄ ⊢ T1 ➡[g] T2 & + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃h, L.ⓑ{J}V1⦄ ⊢ T1 ➡[h, g] T2 & U2 = ⓑ{a,J}V2.T2 ) ∨ - ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 ➡[g] T & ⇧[0, 1] U2 ≡ T & + ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 ➡[h, g] T & ⇧[0, 1] U2 ≡ T & a = true & J = Abbr. #h #g #L #U1 #U2 * -L -U1 -U2 [ #I #L #b #J #W #U1 #H destruct @@ -184,25 +184,25 @@ fact cpx_inv_bind1_aux: ∀h,g,L,U1,U2. ⦃h, L⦄ ⊢ U1 ➡[g] U2 → ] qed-. -lemma cpx_inv_bind1: ∀h,g,a,I,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡[g] U2 → ( - ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L.ⓑ{I}V1⦄ ⊢ T1 ➡[g] T2 & +lemma cpx_inv_bind1: ∀h,g,a,I,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡[h, g] U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃h, L.ⓑ{I}V1⦄ ⊢ T1 ➡[h, g] T2 & U2 = ⓑ{a,I} V2. T2 ) ∨ - ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 ➡[g] T & ⇧[0, 1] U2 ≡ T & + ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 ➡[h, g] T & ⇧[0, 1] U2 ≡ T & a = true & I = Abbr. /2 width=3 by cpx_inv_bind1_aux/ qed-. -lemma cpx_inv_abbr1: ∀h,g,a,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓓ{a}V1.T1 ➡[g] U2 → ( - ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L.ⓓV1⦄ ⊢ T1 ➡[g] T2 & +lemma cpx_inv_abbr1: ∀h,g,a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡[h, g] U2 → ( + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃h, L.ⓓV1⦄ ⊢ T1 ➡[h, g] T2 & U2 = ⓓ{a} V2. T2 ) ∨ - ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 ➡[g] T & ⇧[0, 1] U2 ≡ T & a = true. + ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 ➡[h, g] T & ⇧[0, 1] U2 ≡ T & a = true. #h #g #a #L #V1 #T1 #U2 #H elim (cpx_inv_bind1 … H) -H * /3 width=3/ /3 width=5/ qed-. -lemma cpx_inv_abst1: ∀h,g,a,L,V1,T1,U2. ⦃h, L⦄ ⊢ ⓛ{a}V1.T1 ➡[g] U2 → - ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L.ⓛV1⦄ ⊢ T1 ➡[g] T2 & +lemma cpx_inv_abst1: ∀h,g,a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 ➡[h, g] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃h, L.ⓛV1⦄ ⊢ T1 ➡[h, g] T2 & U2 = ⓛ{a} V2. T2. #h #g #a #L #V1 #T1 #U2 #H elim (cpx_inv_bind1 … H) -H * @@ -211,18 +211,18 @@ elim (cpx_inv_bind1 … H) -H * ] qed-. -fact cpx_inv_flat1_aux: ∀h,g,L,U,U2. ⦃h, L⦄ ⊢ U ➡[g] U2 → +fact cpx_inv_flat1_aux: ∀h,g,L,U,U2. ⦃G, L⦄ ⊢ U ➡[h, g] U2 → ∀J,V1,U1. U = ⓕ{J}V1.U1 → - ∨∨ ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L⦄ ⊢ U1 ➡[g] T2 & + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, g] T2 & U2 = ⓕ{J}V2.T2 - | (⦃h, L⦄ ⊢ U1 ➡[g] U2 ∧ J = Cast) - | (⦃h, L⦄ ⊢ V1 ➡[g] U2 ∧ J = Cast) - | ∃∃a,V2,W1,W2,T1,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L⦄ ⊢ W1 ➡[g] W2 & - ⦃h, L.ⓛW1⦄ ⊢ T1 ➡[g] T2 & + | (⦃G, L⦄ ⊢ U1 ➡[h, g] U2 ∧ J = Cast) + | (⦃G, L⦄ ⊢ V1 ➡[h, g] U2 ∧ J = Cast) + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & + ⦃h, L.ⓛW1⦄ ⊢ T1 ➡[h, g] T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 & J = Appl - | ∃∃a,V,V2,W1,W2,T1,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V & ⇧[0,1] V ≡ V2 & - ⦃h, L⦄ ⊢ W1 ➡[g] W2 & ⦃h, L.ⓓW1⦄ ⊢ T1 ➡[g] T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V & ⇧[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & ⦃h, L.ⓓW1⦄ ⊢ T1 ➡[h, g] T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2 & J = Appl. #h #g #L #U #U2 * -L -U -U2 @@ -239,29 +239,29 @@ fact cpx_inv_flat1_aux: ∀h,g,L,U,U2. ⦃h, L⦄ ⊢ U ➡[g] U2 → ] qed-. -lemma cpx_inv_flat1: ∀h,g,I,L,V1,U1,U2. ⦃h, L⦄ ⊢ ⓕ{I}V1.U1 ➡[g] U2 → - ∨∨ ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L⦄ ⊢ U1 ➡[g] T2 & +lemma cpx_inv_flat1: ∀h,g,I,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓕ{I}V1.U1 ➡[h, g] U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, g] T2 & U2 = ⓕ{I} V2. T2 - | (⦃h, L⦄ ⊢ U1 ➡[g] U2 ∧ I = Cast) - | (⦃h, L⦄ ⊢ V1 ➡[g] U2 ∧ I = Cast) - | ∃∃a,V2,W1,W2,T1,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L⦄ ⊢ W1 ➡[g] W2 & - ⦃h, L.ⓛW1⦄ ⊢ T1 ➡[g] T2 & + | (⦃G, L⦄ ⊢ U1 ➡[h, g] U2 ∧ I = Cast) + | (⦃G, L⦄ ⊢ V1 ➡[h, g] U2 ∧ I = Cast) + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & + ⦃h, L.ⓛW1⦄ ⊢ T1 ➡[h, g] T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 & I = Appl - | ∃∃a,V,V2,W1,W2,T1,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V & ⇧[0,1] V ≡ V2 & - ⦃h, L⦄ ⊢ W1 ➡[g] W2 & ⦃h, L.ⓓW1⦄ ⊢ T1 ➡[g] T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V & ⇧[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & ⦃h, L.ⓓW1⦄ ⊢ T1 ➡[h, g] T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2.ⓐV2.T2 & I = Appl. /2 width=3 by cpx_inv_flat1_aux/ qed-. -lemma cpx_inv_appl1: ∀h,g,L,V1,U1,U2. ⦃h, L⦄ ⊢ ⓐ V1.U1 ➡[g] U2 → - ∨∨ ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L⦄ ⊢ U1 ➡[g] T2 & +lemma cpx_inv_appl1: ∀h,g,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓐ V1.U1 ➡[h, g] U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, g] T2 & U2 = ⓐ V2. T2 - | ∃∃a,V2,W1,W2,T1,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L⦄ ⊢ W1 ➡[g] W2 & - ⦃h, L.ⓛW1⦄ ⊢ T1 ➡[g] T2 & + | ∃∃a,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & + ⦃h, L.ⓛW1⦄ ⊢ T1 ➡[h, g] T2 & U1 = ⓛ{a}W1.T1 & U2 = ⓓ{a}ⓝW2.V2.T2 - | ∃∃a,V,V2,W1,W2,T1,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V & ⇧[0,1] V ≡ V2 & - ⦃h, L⦄ ⊢ W1 ➡[g] W2 & ⦃h, L.ⓓW1⦄ ⊢ T1 ➡[g] T2 & + | ∃∃a,V,V2,W1,W2,T1,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V & ⇧[0,1] V ≡ V2 & + ⦃G, L⦄ ⊢ W1 ➡[h, g] W2 & ⦃h, L.ⓓW1⦄ ⊢ T1 ➡[h, g] T2 & U1 = ⓓ{a}W1.T1 & U2 = ⓓ{a}W2. ⓐV2. T2. #h #g #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H * [ /3 width=5/ @@ -272,8 +272,8 @@ lemma cpx_inv_appl1: ∀h,g,L,V1,U1,U2. ⦃h, L⦄ ⊢ ⓐ V1.U1 ➡[g] U2 → qed-. (* Note: the main property of simple terms *) -lemma cpx_inv_appl1_simple: ∀h,g,L,V1,T1,U. ⦃h, L⦄ ⊢ ⓐV1.T1 ➡[g] U → 𝐒⦃T1⦄ → - ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L⦄ ⊢ T1 ➡[g] T2 & +lemma cpx_inv_appl1_simple: ∀h,g,L,V1,T1,U. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡[h, g] U → 𝐒⦃T1⦄ → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 & U = ⓐV2.T2. #h #g #L #V1 #T1 #U #H #HT1 elim (cpx_inv_appl1 … H) -H * @@ -285,11 +285,11 @@ elim (cpx_inv_appl1 … H) -H * ] qed-. -lemma cpx_inv_cast1: ∀h,g,L,V1,U1,U2. ⦃h, L⦄ ⊢ ⓝV1.U1 ➡[g] U2 → - ∨∨ ∃∃V2,T2. ⦃h, L⦄ ⊢ V1 ➡[g] V2 & ⦃h, L⦄ ⊢ U1 ➡[g] T2 & +lemma cpx_inv_cast1: ∀h,g,L,V1,U1,U2. ⦃G, L⦄ ⊢ ⓝV1.U1 ➡[h, g] U2 → + ∨∨ ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡[h, g] V2 & ⦃G, L⦄ ⊢ U1 ➡[h, g] T2 & U2 = ⓝ V2. T2 - | ⦃h, L⦄ ⊢ U1 ➡[g] U2 - | ⦃h, L⦄ ⊢ V1 ➡[g] U2. + | ⦃G, L⦄ ⊢ U1 ➡[h, g] U2 + | ⦃G, L⦄ ⊢ V1 ➡[h, g] U2. #h #g #L #V1 #U1 #U2 #H elim (cpx_inv_flat1 … H) -H * [ /3 width=5/ |2,3: /2 width=1/ @@ -300,8 +300,8 @@ qed-. (* Basic forward lemmas *****************************************************) -lemma cpx_fwd_bind1_minus: ∀h,g,I,L,V1,T1,T. ⦃h, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[g] T → ∀b. - ∃∃V2,T2. ⦃h, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[g] ⓑ{b,I}V2.T2 & +lemma cpx_fwd_bind1_minus: ∀h,g,I,L,V1,T1,T. ⦃G, L⦄ ⊢ -ⓑ{I}V1.T1 ➡[h, g] T → ∀b. + ∃∃V2,T2. ⦃G, L⦄ ⊢ ⓑ{b,I}V1.T1 ➡[h, g] ⓑ{b,I}V2.T2 & T = -ⓑ{I}V2.T2. #h #g #I #L #V1 #T1 #T #H #b elim (cpx_inv_bind1 … H) -H * @@ -310,7 +310,7 @@ elim (cpx_inv_bind1 … H) -H * ] qed-. -lemma cpx_fwd_shift1: ∀h,g,L1,L,T1,T. ⦃h, L⦄ ⊢ L1 @@ T1 ➡[g] T → +lemma cpx_fwd_shift1: ∀h,g,L1,L,T1,T. ⦃G, L⦄ ⊢ L1 @@ T1 ➡[h, g] T → ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. #h #g #L1 @(lenv_ind_dx … L1) -L1 normalize [ #L #T1 #T #HT1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_cix.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_cix.ma index 7567f9f9c..9aa2d62a8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_cix.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_cix.ma @@ -19,7 +19,7 @@ include "basic_2/reduction/cpx.ma". (* Advanced forward lemmas on context-sensitive extended irreducible terms **) -lemma cpx_fwd_cix: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → ⦃h, L⦄ ⊢ 𝐈[g]⦃T1⦄ → T2 = T1. +lemma cpx_fwd_cix: ∀h,g,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃T1⦄ → T2 = T1. #h #g #L #T1 #T2 #H elim H -L -T1 -T2 [ // | #L #k #l #Hkl #H elim (cix_inv_sort … Hkl H) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_lift.ma index 30eef6262..40f6e1e97 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_lift.ma @@ -120,8 +120,8 @@ qed-. (* Properties on supclosure *************************************************) lemma fsupq_cpx_trans: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → - ∀U2. ⦃h, L2⦄ ⊢ T2 ➡[g] U2 → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. + ∀U2. ⦃h, L2⦄ ⊢ T2 ➡[h, g] U2 → + ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. #h #g #L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [1: /2 width=3/ |3,4,5: /3 width=3/ ] [ #I #L1 #V2 #U2 #HVU2 elim (lift_total U2 0 1) /4 width=9/ @@ -133,16 +133,16 @@ lemma fsupq_cpx_trans: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → qed-. lemma fsupq_ssta_trans: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → - ∀U2,l. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l+1, U2⦄ → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. + ∀U2,l. ⦃h, L2⦄ ⊢ T2 •[h, g] ⦃l+1, U2⦄ → + ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. /3 width=4 by fsupq_cpx_trans, ssta_cpx/ qed-. lemma fsup_cpx_trans: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → - ∀U2. ⦃h, L2⦄ ⊢ T2 ➡[g] U2 → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. + ∀U2. ⦃h, L2⦄ ⊢ T2 ➡[h, g] U2 → + ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. /3 width=3 by fsupq_cpx_trans, fsup_fsupq/ qed-. lemma fsup_ssta_trans: ∀h,g,L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → - ∀U2,l. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l+1, U2⦄ → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. + ∀U2,l. ⦃h, L2⦄ ⊢ T2 •[h, g] ⦃l+1, U2⦄ → + ∃∃U1. ⦃h, L1⦄ ⊢ T1 ➡[h, g] U1 & ⦃L1, U1⦄ ⊃⸮ ⦃L2, U2⦄. /3 width=4 by fsupq_ssta_trans, fsup_fsupq/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crr.ma index b9a51d862..9fb127750 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crr.ma @@ -43,7 +43,7 @@ interpretation (* Basic inversion lemmas ***************************************************) -fact crr_inv_sort_aux: ∀L,T,k. L ⊢ 𝐑⦃T⦄ → T = ⋆k → ⊥. +fact crr_inv_sort_aux: ∀L,T,k. ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → T = ⋆k → ⊥. #L #T #k0 * -L -T [ #L #K #V #i #HLK #H destruct | #L #V #T #_ #H destruct @@ -56,10 +56,10 @@ fact crr_inv_sort_aux: ∀L,T,k. L ⊢ 𝐑⦃T⦄ → T = ⋆k → ⊥. ] qed-. -lemma crr_inv_sort: ∀L,k. L ⊢ 𝐑⦃⋆k⦄ → ⊥. +lemma crr_inv_sort: ∀L,k. ⦃G, L⦄ ⊢ 𝐑⦃⋆k⦄ → ⊥. /2 width=5 by crr_inv_sort_aux/ qed-. -fact crr_inv_lref_aux: ∀L,T,i. L ⊢ 𝐑⦃T⦄ → T = #i → ∃∃K,V. ⇩[0, i] L ≡ K.ⓓV. +fact crr_inv_lref_aux: ∀L,T,i. ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → T = #i → ∃∃K,V. ⇩[0, i] L ≡ K.ⓓV. #L #T #j * -L -T [ #L #K #V #i #HLK #H destruct /2 width=3/ | #L #V #T #_ #H destruct @@ -72,10 +72,10 @@ fact crr_inv_lref_aux: ∀L,T,i. L ⊢ 𝐑⦃T⦄ → T = #i → ∃∃K,V. ⇩ ] qed-. -lemma crr_inv_lref: ∀L,i. L ⊢ 𝐑⦃#i⦄ → ∃∃K,V. ⇩[0, i] L ≡ K.ⓓV. +lemma crr_inv_lref: ∀L,i. ⦃G, L⦄ ⊢ 𝐑⦃#i⦄ → ∃∃K,V. ⇩[0, i] L ≡ K.ⓓV. /2 width=3 by crr_inv_lref_aux/ qed-. -fact crr_inv_gref_aux: ∀L,T,p. L ⊢ 𝐑⦃T⦄ → T = §p → ⊥. +fact crr_inv_gref_aux: ∀L,T,p. ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → T = §p → ⊥. #L #T #q * -L -T [ #L #K #V #i #HLK #H destruct | #L #V #T #_ #H destruct @@ -88,7 +88,7 @@ fact crr_inv_gref_aux: ∀L,T,p. L ⊢ 𝐑⦃T⦄ → T = §p → ⊥. ] qed-. -lemma crr_inv_gref: ∀L,p. L ⊢ 𝐑⦃§p⦄ → ⊥. +lemma crr_inv_gref: ∀L,p. ⦃G, L⦄ ⊢ 𝐑⦃§p⦄ → ⊥. /2 width=5 by crr_inv_gref_aux/ qed-. lemma trr_inv_atom: ∀I. ⋆ ⊢ 𝐑⦃⓪{I}⦄ → ⊥. @@ -100,8 +100,8 @@ lemma trr_inv_atom: ∀I. ⋆ ⊢ 𝐑⦃⓪{I}⦄ → ⊥. ] qed-. -fact crr_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⦄. +fact crr_inv_ib2_aux: ∀a,I,L,W,U,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → T = ⓑ{a,I}W.U → + ⦃G, L⦄ ⊢ 𝐑⦃W⦄ ∨ L.ⓑ{I}W ⊢ 𝐑⦃U⦄. #b #J #L #W0 #U #T #HI * -L -T [ #L #K #V #i #_ #H destruct | #L #V #T #_ #H destruct @@ -116,12 +116,12 @@ fact crr_inv_ib2_aux: ∀a,I,L,W,U,T. ib2 a I → L ⊢ 𝐑⦃T⦄ → T = ⓑ{ ] qed-. -lemma crr_inv_ib2: ∀a,I,L,W,T. ib2 a I → L ⊢ 𝐑⦃ⓑ{a,I}W.T⦄ → - L ⊢ 𝐑⦃W⦄ ∨ L.ⓑ{I}W ⊢ 𝐑⦃T⦄. +lemma crr_inv_ib2: ∀a,I,L,W,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐑⦃ⓑ{a,I}W.T⦄ → + ⦃G, L⦄ ⊢ 𝐑⦃W⦄ ∨ L.ⓑ{I}W ⊢ 𝐑⦃T⦄. /2 width=5 by crr_inv_ib2_aux/ qed-. -fact crr_inv_appl_aux: ∀L,W,U,T. L ⊢ 𝐑⦃T⦄ → T = ⓐW.U → - ∨∨ L ⊢ 𝐑⦃W⦄ | L ⊢ 𝐑⦃U⦄ | (𝐒⦃U⦄ → ⊥). +fact crr_inv_appl_aux: ∀L,W,U,T. ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → T = ⓐW.U → + ∨∨ ⦃G, L⦄ ⊢ 𝐑⦃W⦄ | ⦃G, L⦄ ⊢ 𝐑⦃U⦄ | (𝐒⦃U⦄ → ⊥). #L #W0 #U #T * -L -T [ #L #K #V #i #_ #H destruct | #L #V #T #HV #H destruct /2 width=1/ @@ -137,5 +137,5 @@ fact crr_inv_appl_aux: ∀L,W,U,T. L ⊢ 𝐑⦃T⦄ → T = ⓐW.U → ] qed-. -lemma crr_inv_appl: ∀L,V,T. L ⊢ 𝐑⦃ⓐV.T⦄ → ∨∨ L ⊢ 𝐑⦃V⦄ | L ⊢ 𝐑⦃T⦄ | (𝐒⦃T⦄ → ⊥). +lemma crr_inv_appl: ∀L,V,T. ⦃G, L⦄ ⊢ 𝐑⦃ⓐV.T⦄ → ∨∨ ⦃G, L⦄ ⊢ 𝐑⦃V⦄ | ⦃G, L⦄ ⊢ 𝐑⦃T⦄ | (𝐒⦃T⦄ → ⊥). /2 width=3 by crr_inv_appl_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crr_append.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crr_append.ma index fdb29e2af..e0861ae55 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crr_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crr_append.ma @@ -19,14 +19,14 @@ include "basic_2/reduction/crr.ma". (* Advanved properties ******************************************************) -lemma crr_append_sn: ∀L,K,T. L ⊢ 𝐑⦃T⦄ → K @@ L ⊢ 𝐑⦃T⦄. +lemma crr_append_sn: ∀L,K,T. ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → K @@ ⦃G, L⦄ ⊢ 𝐑⦃T⦄. #L #K0 #T #H elim H -L -T /2 width=1/ #L #K #V #i #HLK lapply (ldrop_fwd_length_lt2 … HLK) #Hi lapply (ldrop_O1_append_sn_le … HLK … K0) -HLK /2 width=2/ -Hi /2 width=3/ qed. -lemma trr_crr: ∀L,T. ⋆ ⊢ 𝐑⦃T⦄ → L ⊢ 𝐑⦃T⦄. +lemma trr_crr: ∀L,T. ⋆ ⊢ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ 𝐑⦃T⦄. #L #T #H lapply (crr_append_sn … H) // qed. @@ -49,7 +49,7 @@ fact crr_inv_labst_last_aux: ∀L1,T,W. L1 ⊢ 𝐑⦃T⦄ → ] qed. -lemma crr_inv_labst_last: ∀L,T,W. ⋆.ⓛW @@ L ⊢ 𝐑⦃T⦄ → L ⊢ 𝐑⦃T⦄. +lemma crr_inv_labst_last: ∀L,T,W. ⋆.ⓛW @@ ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ 𝐑⦃T⦄. /2 width=4/ qed-. lemma crr_inv_trr: ∀T,W. ⋆.ⓛW ⊢ 𝐑⦃T⦄ → ⋆ ⊢ 𝐑⦃T⦄. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crr_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crr_lift.ma index 731d00bb2..5d923559d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crr_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crr_lift.ma @@ -20,7 +20,7 @@ include "basic_2/reduction/crr.ma". (* Properties on relocation *************************************************) lemma crr_lift: ∀K,T. K ⊢ 𝐑⦃T⦄ → ∀L,d,e. ⇩[d, e] L ≡ K → - ∀U. ⇧[d, e] T ≡ U → L ⊢ 𝐑⦃U⦄. + ∀U. ⇧[d, e] T ≡ U → ⦃G, L⦄ ⊢ 𝐑⦃U⦄. #K #T #H elim H -K -T [ #K #K0 #V #i #HK0 #L #d #e #HLK #X #H elim (lift_inv_lref1 … H) -H * #Hid #H destruct @@ -46,7 +46,7 @@ lemma crr_lift: ∀K,T. K ⊢ 𝐑⦃T⦄ → ∀L,d,e. ⇩[d, e] L ≡ K → ] qed. -lemma crr_inv_lift: ∀L,U. L ⊢ 𝐑⦃U⦄ → ∀K,d,e. ⇩[d, e] L ≡ K → +lemma crr_inv_lift: ∀L,U. ⦃G, L⦄ ⊢ 𝐑⦃U⦄ → ∀K,d,e. ⇩[d, e] L ≡ K → ∀T. ⇧[d, e] T ≡ U → K ⊢ 𝐑⦃T⦄. #L #U #H elim H -L -U [ #L #L0 #W #i #HK0 #K #d #e #HLK #X #H diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx.ma index 059e62151..922dcd4fe 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx.ma @@ -37,13 +37,13 @@ interpretation (* Basic properties *********************************************************) -lemma crr_crx: ∀h,g,L,T. L ⊢ 𝐑⦃T⦄ → ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄. +lemma crr_crx: ∀h,g,L,T. ⦃G, L⦄ ⊢ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄. #h #g #L #T #H elim H -L -T // /2 width=1/ /2 width=4/ qed. (* Basic inversion lemmas ***************************************************) -fact crx_inv_sort_aux: ∀h,g,L,T,k. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → T = ⋆k → +fact crx_inv_sort_aux: ∀h,g,L,T,k. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → T = ⋆k → ∃l. deg h g k (l+1). #h #g #L #T #k0 * -L -T [ #L #k #l #Hkl #H destruct /2 width=2/ @@ -58,10 +58,10 @@ fact crx_inv_sort_aux: ∀h,g,L,T,k. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → T = ⋆k ] qed-. -lemma crx_inv_sort: ∀h,g,L,k. ⦃h, L⦄ ⊢ 𝐑[g]⦃⋆k⦄ → ∃l. deg h g k (l+1). +lemma crx_inv_sort: ∀h,g,L,k. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃⋆k⦄ → ∃l. deg h g k (l+1). /2 width=4 by crx_inv_sort_aux/ qed-. -fact crx_inv_lref_aux: ∀h,g,L,T,i. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → T = #i → +fact crx_inv_lref_aux: ∀h,g,L,T,i. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → T = #i → ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V. #h #g #L #T #j * -L -T [ #L #k #l #_ #H destruct @@ -76,10 +76,10 @@ fact crx_inv_lref_aux: ∀h,g,L,T,i. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → T = #i ] qed-. -lemma crx_inv_lref: ∀h,g,L,i. ⦃h, L⦄ ⊢ 𝐑[g]⦃#i⦄ → ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V. +lemma crx_inv_lref: ∀h,g,L,i. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃#i⦄ → ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V. /2 width=5 by crx_inv_lref_aux/ qed-. -fact crx_inv_gref_aux: ∀h,g,L,T,p. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → T = §p → ⊥. +fact crx_inv_gref_aux: ∀h,g,L,T,p. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → T = §p → ⊥. #h #g #L #T #q * -L -T [ #L #k #l #_ #H destruct | #I #L #K #V #i #HLK #H destruct @@ -93,10 +93,10 @@ fact crx_inv_gref_aux: ∀h,g,L,T,p. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → T = §p ] qed-. -lemma crx_inv_gref: ∀h,g,L,p. ⦃h, L⦄ ⊢ 𝐑[g]⦃§p⦄ → ⊥. +lemma crx_inv_gref: ∀h,g,L,p. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃§p⦄ → ⊥. /2 width=7 by crx_inv_gref_aux/ qed-. -lemma trx_inv_atom: ∀h,g,I. ⦃h, ⋆⦄ ⊢ 𝐑[g]⦃⓪{I}⦄ → +lemma trx_inv_atom: ∀h,g,I. ⦃h, ⋆⦄ ⊢ 𝐑[h, g]⦃⓪{I}⦄ → ∃∃k,l. deg h g k (l+1) & I = Sort k. #h #g * #i #H [ elim (crx_inv_sort … H) -H /2 width=4/ @@ -106,8 +106,8 @@ lemma trx_inv_atom: ∀h,g,I. ⦃h, ⋆⦄ ⊢ 𝐑[g]⦃⓪{I}⦄ → ] qed-. -fact crx_inv_ib2_aux: ∀h,g,a,I,L,W,U,T. ib2 a I → ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → - T = ⓑ{a,I}W.U → ⦃h, L⦄ ⊢ 𝐑[g]⦃W⦄ ∨ ⦃h, L.ⓑ{I}W⦄ ⊢ 𝐑[g]⦃U⦄. +fact crx_inv_ib2_aux: ∀h,g,a,I,L,W,U,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → + T = ⓑ{a,I}W.U → ⦃G, L⦄ ⊢ 𝐑[h, g]⦃W⦄ ∨ ⦃h, L.ⓑ{I}W⦄ ⊢ 𝐑[h, g]⦃U⦄. #h #g #b #J #L #W0 #U #T #HI * -L -T [ #L #k #l #_ #H destruct | #I #L #K #V #i #_ #H destruct @@ -123,12 +123,12 @@ fact crx_inv_ib2_aux: ∀h,g,a,I,L,W,U,T. ib2 a I → ⦃h, L⦄ ⊢ 𝐑[g]⦃T ] qed-. -lemma crx_inv_ib2: ∀h,g,a,I,L,W,T. ib2 a I → ⦃h, L⦄ ⊢ 𝐑[g]⦃ⓑ{a,I}W.T⦄ → - ⦃h, L⦄ ⊢ 𝐑[g]⦃W⦄ ∨ ⦃h, L.ⓑ{I}W⦄ ⊢ 𝐑[g]⦃T⦄. +lemma crx_inv_ib2: ∀h,g,a,I,L,W,T. ib2 a I → ⦃G, L⦄ ⊢ 𝐑[h, g]⦃ⓑ{a,I}W.T⦄ → + ⦃G, L⦄ ⊢ 𝐑[h, g]⦃W⦄ ∨ ⦃h, L.ⓑ{I}W⦄ ⊢ 𝐑[h, g]⦃T⦄. /2 width=5 by crx_inv_ib2_aux/ qed-. -fact crx_inv_appl_aux: ∀h,g,L,W,U,T. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → T = ⓐW.U → - ∨∨ ⦃h, L⦄ ⊢ 𝐑[g]⦃W⦄ | ⦃h, L⦄ ⊢ 𝐑[g]⦃U⦄ | (𝐒⦃U⦄ → ⊥). +fact crx_inv_appl_aux: ∀h,g,L,W,U,T. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → T = ⓐW.U → + ∨∨ ⦃G, L⦄ ⊢ 𝐑[h, g]⦃W⦄ | ⦃G, L⦄ ⊢ 𝐑[h, g]⦃U⦄ | (𝐒⦃U⦄ → ⊥). #h #g #L #W0 #U #T * -L -T [ #L #k #l #_ #H destruct | #I #L #K #V #i #_ #H destruct @@ -145,6 +145,6 @@ fact crx_inv_appl_aux: ∀h,g,L,W,U,T. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → T = ⓐ ] qed-. -lemma crx_inv_appl: ∀h,g,L,V,T. ⦃h, L⦄ ⊢ 𝐑[g]⦃ⓐV.T⦄ → - ∨∨ ⦃h, L⦄ ⊢ 𝐑[g]⦃V⦄ | ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ | (𝐒⦃T⦄ → ⊥). +lemma crx_inv_appl: ∀h,g,L,V,T. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃ⓐV.T⦄ → + ∨∨ ⦃G, L⦄ ⊢ 𝐑[h, g]⦃V⦄ | ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ | (𝐒⦃T⦄ → ⊥). /2 width=3 by crx_inv_appl_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_append.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_append.ma index 9d5dca4a6..bd91ad1d5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_append.ma @@ -19,13 +19,13 @@ include "basic_2/reduction/crx.ma". (* Advanved properties ******************************************************) -lemma crx_append_sn: ∀h,g,L,K,T. ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄ → ⦃h, K @@ L⦄ ⊢ 𝐑[g]⦃T⦄. +lemma crx_append_sn: ∀h,g,L,K,T. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄ → ⦃h, K @@ L⦄ ⊢ 𝐑[h, g]⦃T⦄. #h #g #L #K0 #T #H elim H -L -T /2 width=1/ /2 width=2/ #I #L #K #V #i #HLK lapply (ldrop_fwd_length_lt2 … HLK) #Hi lapply (ldrop_O1_append_sn_le … HLK … K0) -HLK /2 width=2/ -Hi /2 width=4/ qed. -lemma trx_crx: ∀h,g,L,T. ⦃h, ⋆⦄ ⊢ 𝐑[g]⦃T⦄ → ⦃h, L⦄ ⊢ 𝐑[g]⦃T⦄. +lemma trx_crx: ∀h,g,L,T. ⦃h, ⋆⦄ ⊢ 𝐑[h, g]⦃T⦄ → ⦃G, L⦄ ⊢ 𝐑[h, g]⦃T⦄. #h #g #L #T #H lapply (crx_append_sn … H) // qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_lift.ma index c1989c282..5a5a286ff 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx_lift.ma @@ -19,8 +19,8 @@ include "basic_2/reduction/crx.ma". (* Properties on relocation *************************************************) -lemma crx_lift: ∀h,g,K,T. ⦃h, K⦄ ⊢ 𝐑[g]⦃T⦄ → ∀L,d,e. ⇩[d, e] L ≡ K → - ∀U. ⇧[d, e] T ≡ U → ⦃h, L⦄ ⊢ 𝐑[g]⦃U⦄. +lemma crx_lift: ∀h,g,K,T. ⦃h, K⦄ ⊢ 𝐑[h, g]⦃T⦄ → ∀L,d,e. ⇩[d, e] L ≡ K → + ∀U. ⇧[d, e] T ≡ U → ⦃G, L⦄ ⊢ 𝐑[h, g]⦃U⦄. #h #g #K #T #H elim H -K -T [ #K #k #l #Hkl #L #d #e #_ #X #H >(lift_inv_sort1 … H) -X /2 width=2/ @@ -48,8 +48,8 @@ lemma crx_lift: ∀h,g,K,T. ⦃h, K⦄ ⊢ 𝐑[g]⦃T⦄ → ∀L,d,e. ⇩[d, e ] qed. -lemma crx_inv_lift: ∀h,g,L,U. ⦃h, L⦄ ⊢ 𝐑[g]⦃U⦄ → ∀K,d,e. ⇩[d, e] L ≡ K → - ∀T. ⇧[d, e] T ≡ U → ⦃h, K⦄ ⊢ 𝐑[g]⦃T⦄. +lemma crx_inv_lift: ∀h,g,L,U. ⦃G, L⦄ ⊢ 𝐑[h, g]⦃U⦄ → ∀K,d,e. ⇩[d, e] L ≡ K → + ∀T. ⇧[d, e] T ≡ U → ⦃h, K⦄ ⊢ 𝐑[h, g]⦃T⦄. #h #g #L #U #H elim H -L -U [ #L #k #l #Hkl #K #d #e #_ #X #H >(lift_inv_sort2 … H) -X /2 width=2/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr.ma index afab7f20f..9fa1327fa 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr.ma @@ -44,7 +44,7 @@ lemma lpr_inv_pair2: ∀I,L1,K2,V2. L1 ⊢ ➡ K2. ⓑ{I} V2 → (* Basic properties *********************************************************) (* Note: lemma 250 *) -lemma lpr_refl: ∀L. L ⊢ ➡ L. +lemma lpr_refl: ∀L. ⦃G, L⦄ ⊢ ➡ L. /2 width=1 by lpx_sn_refl/ qed. lemma lpr_pair: ∀I,K1,K2,V1,V2. K1 ⊢ ➡ K2 → K1 ⊢ V1 ➡ V2 → @@ -66,7 +66,7 @@ lemma lpr_fwd_append1: ∀K1,L1,L. K1 @@ L1 ⊢ ➡ L → ∃∃K2,L2. K1 ⊢ ➡ K2 & L = K2 @@ L2. /2 width=2 by lpx_sn_fwd_append1/ qed-. -lemma lpr_fwd_append2: ∀L,K2,L2. L ⊢ ➡ K2 @@ L2 → +lemma lpr_fwd_append2: ∀L,K2,L2. ⦃G, L⦄ ⊢ ➡ K2 @@ L2 → ∃∃K1,L1. K1 ⊢ ➡ K2 & L = K1 @@ L1. /2 width=2 by lpx_sn_fwd_append2/ 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 index fd7b613d6..120e5e60e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_ldrop.ma @@ -35,7 +35,7 @@ lemma lpr_ldrop_trans_O1: dropable_dx lpr. (* Properties on context-sensitive parallel reduction for terms *************) 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⦄. + ∃∃L,U1. L1 ⊢ ➡ L & ⦃G, 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/ ] [ #L #K #U #T #d #e #HLK #HUT #He #U2 #HU2 elim (lift_total U2 d e) #T2 #HUT2 diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpr.ma index d87242057..2cc08f081 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpr_lpr.ma @@ -27,8 +27,8 @@ fact cpr_conf_lpr_atom_atom: fact cpr_conf_lpr_atom_delta: ∀L0,i. ( ∀L,T. ⦃L0, #i⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀K0,V0. ⇩[O, i] L0 ≡ K0.ⓓV0 → @@ -51,8 +51,8 @@ qed-. fact cpr_conf_lpr_delta_delta: ∀L0,i. ( ∀L,T. ⦃L0, #i⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀K0,V0. ⇩[O, i] L0 ≡ K0.ⓓV0 → @@ -80,8 +80,8 @@ qed-. fact cpr_conf_lpr_bind_bind: ∀a,I,L0,V0,T0. ( ∀L,T. ⦃L0,ⓑ{a,I}V0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0.ⓑ{I}V0 ⊢ T0 ➡ T1 → @@ -97,8 +97,8 @@ qed-. fact cpr_conf_lpr_bind_zeta: ∀L0,V0,T0. ( ∀L,T. ⦃L0,+ⓓV0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0.ⓓV0 ⊢ T0 ➡ T1 → @@ -114,8 +114,8 @@ qed-. fact cpr_conf_lpr_zeta_zeta: ∀L0,V0,T0. ( ∀L,T. ⦃L0,+ⓓV0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀T1. L0.ⓓV0 ⊢ T0 ➡ T1 → ∀X1. ⇧[O, 1] X1 ≡ T1 → @@ -133,8 +133,8 @@ qed-. fact cpr_conf_lpr_flat_flat: ∀I,L0,V0,T0. ( ∀L,T. ⦃L0,ⓕ{I}V0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0 ⊢ T0 ➡ T1 → @@ -150,8 +150,8 @@ qed-. fact cpr_conf_lpr_flat_tau: ∀L0,V0,T0. ( ∀L,T. ⦃L0,ⓝV0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀V1,T1. L0 ⊢ T0 ➡ T1 → ∀T2. L0 ⊢ T0 ➡ T2 → @@ -165,8 +165,8 @@ qed-. fact cpr_conf_lpr_tau_tau: ∀L0,V0,T0. ( ∀L,T. ⦃L0,ⓝV0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀T1. L0 ⊢ T0 ➡ T1 → ∀T2. L0 ⊢ T0 ➡ T2 → @@ -180,8 +180,8 @@ qed-. fact cpr_conf_lpr_flat_beta: ∀a,L0,V0,W0,T0. ( ∀L,T. ⦃L0,ⓐV0.ⓛ{a}W0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0 ⊢ ⓛ{a}W0.T0 ➡ T1 → @@ -205,8 +205,8 @@ qed-. fact cpr_conf_lpr_flat_theta: ∀a,L0,V0,W0,T0. ( ∀L,T. ⦃L0,ⓐV0.ⓓ{a}W0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀V1. L0 ⊢ V0 ➡ V1 → ∀T1. L0 ⊢ ⓓ{a}W0.T0 ➡ T1 → @@ -234,8 +234,8 @@ qed-. fact cpr_conf_lpr_beta_beta: ∀a,L0,V0,W0,T0. ( ∀L,T. ⦃L0,ⓐV0.ⓛ{a}W0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀V1. L0 ⊢ V0 ➡ V1 → ∀W1. L0 ⊢ W0 ➡ W1 → ∀T1. L0.ⓛW0 ⊢ T0 ➡ T1 → @@ -256,8 +256,8 @@ qed-. fact cpr_conf_lpr_theta_theta: ∀a,L0,V0,W0,T0. ( ∀L,T. ⦃L0,ⓐV0.ⓓ{a}W0.T0⦄ ⊃+ ⦃L, T⦄ → - ∀T1. L ⊢ T ➡ T1 → ∀T2. L ⊢ T ➡ T2 → - ∀L1. L ⊢ ➡ L1 → ∀L2. L ⊢ ➡ L2 → + ∀T1. ⦃G, L⦄ ⊢ T ➡ T1 → ∀T2. ⦃G, L⦄ ⊢ T ➡ T2 → + ∀L1. ⦃G, L⦄ ⊢ ➡ L1 → ∀L2. ⦃G, L⦄ ⊢ ➡ L2 → ∃∃T0. L1 ⊢ T1 ➡ T0 & L2 ⊢ T2 ➡ T0 ) → ∀V1. L0 ⊢ V0 ➡ V1 → ∀U1. ⇧[O, 1] V1 ≡ U1 → diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx.ma index d34212910..547520714 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx.ma @@ -25,50 +25,50 @@ interpretation "extended parallel reduction (local environment, sn variant)" (* Basic inversion lemmas ***************************************************) -lemma lpx_inv_atom1: ∀h,g,L2. ⦃h, ⋆⦄ ⊢ ➡[g] L2 → L2 = ⋆. +lemma lpx_inv_atom1: ∀h,g,L2. ⦃h, ⋆⦄ ⊢ ➡[h, g] L2 → L2 = ⋆. /2 width=4 by lpx_sn_inv_atom1_aux/ qed-. -lemma lpx_inv_pair1: ∀h,g,I,K1,V1,L2. ⦃h, K1.ⓑ{I}V1⦄ ⊢ ➡[g] L2 → - ∃∃K2,V2. ⦃h, K1⦄ ⊢ ➡[g] K2 & ⦃h, K1⦄ ⊢ V1 ➡[g] V2 & +lemma lpx_inv_pair1: ∀h,g,I,K1,V1,L2. ⦃h, K1.ⓑ{I}V1⦄ ⊢ ➡[h, g] L2 → + ∃∃K2,V2. ⦃h, K1⦄ ⊢ ➡[h, g] K2 & ⦃h, K1⦄ ⊢ V1 ➡[h, g] V2 & L2 = K2. ⓑ{I} V2. /2 width=3 by lpx_sn_inv_pair1_aux/ qed-. -lemma lpx_inv_atom2: ∀h,g,L1. ⦃h, L1⦄ ⊢ ➡[g] ⋆ → L1 = ⋆. +lemma lpx_inv_atom2: ∀h,g,L1. ⦃h, L1⦄ ⊢ ➡[h, g] ⋆ → L1 = ⋆. /2 width=4 by lpx_sn_inv_atom2_aux/ qed-. -lemma lpx_inv_pair2: ∀h,g,I,L1,K2,V2. ⦃h, L1⦄ ⊢ ➡[g] K2.ⓑ{I}V2 → - ∃∃K1,V1. ⦃h, K1⦄ ⊢ ➡[g] K2 & ⦃h, K1⦄ ⊢ V1 ➡[g] V2 & +lemma lpx_inv_pair2: ∀h,g,I,L1,K2,V2. ⦃h, L1⦄ ⊢ ➡[h, g] K2.ⓑ{I}V2 → + ∃∃K1,V1. ⦃h, K1⦄ ⊢ ➡[h, g] K2 & ⦃h, K1⦄ ⊢ V1 ➡[h, g] V2 & L1 = K1. ⓑ{I} V1. /2 width=3 by lpx_sn_inv_pair2_aux/ qed-. (* Basic properties *********************************************************) -lemma lpx_refl: ∀h,g,L. ⦃h, L⦄ ⊢ ➡[g] L. +lemma lpx_refl: ∀h,g,L. ⦃G, L⦄ ⊢ ➡[h, g] L. /2 width=1 by lpx_sn_refl/ qed. -lemma lpx_pair: ∀h,g,I,K1,K2,V1,V2. ⦃h, K1⦄ ⊢ ➡[g] K2 → ⦃h, K1⦄ ⊢ V1 ➡[g] V2 → - ⦃h, K1.ⓑ{I}V1⦄ ⊢ ➡[g] K2.ⓑ{I}V2. +lemma lpx_pair: ∀h,g,I,K1,K2,V1,V2. ⦃h, K1⦄ ⊢ ➡[h, g] K2 → ⦃h, K1⦄ ⊢ V1 ➡[h, g] V2 → + ⦃h, K1.ⓑ{I}V1⦄ ⊢ ➡[h, g] K2.ⓑ{I}V2. /2 width=1/ qed. -lemma lpx_append: ∀h,g,K1,K2. ⦃h, K1⦄ ⊢ ➡[g] K2 → ∀L1,L2. ⦃h, L1⦄ ⊢ ➡[g] L2 → - ⦃h, L1 @@ K1⦄ ⊢ ➡[g] L2 @@ K2. +lemma lpx_append: ∀h,g,K1,K2. ⦃h, K1⦄ ⊢ ➡[h, g] K2 → ∀L1,L2. ⦃h, L1⦄ ⊢ ➡[h, g] L2 → + ⦃h, L1 @@ K1⦄ ⊢ ➡[h, g] L2 @@ K2. /3 width=1 by lpx_sn_append, cpx_append/ qed. -lemma lpr_lpx: ∀h,g,L1,L2. L1 ⊢ ➡ L2 → ⦃h, L1⦄ ⊢ ➡[g] L2. +lemma lpr_lpx: ∀h,g,L1,L2. L1 ⊢ ➡ L2 → ⦃h, L1⦄ ⊢ ➡[h, g] L2. #h #g #L1 #L2 #H elim H -L1 -L2 // /3 width=1/ qed. (* Basic forward lemmas *****************************************************) -lemma lpx_fwd_length: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡[g] L2 → |L1| = |L2|. +lemma lpx_fwd_length: ∀h,g,L1,L2. ⦃h, L1⦄ ⊢ ➡[h, g] L2 → |L1| = |L2|. /2 width=2 by lpx_sn_fwd_length/ qed-. (* Advanced forward lemmas **************************************************) -lemma lpx_fwd_append1: ∀h,g,K1,L1,L. ⦃h, K1 @@ L1⦄ ⊢ ➡[g] L → - ∃∃K2,L2. ⦃h, K1⦄ ⊢ ➡[g] K2 & L = K2 @@ L2. +lemma lpx_fwd_append1: ∀h,g,K1,L1,L. ⦃h, K1 @@ L1⦄ ⊢ ➡[h, g] L → + ∃∃K2,L2. ⦃h, K1⦄ ⊢ ➡[h, g] K2 & L = K2 @@ L2. /2 width=2 by lpx_sn_fwd_append1/ qed-. -lemma lpx_fwd_append2: ∀h,g,L,K2,L2. ⦃h, L⦄ ⊢ ➡[g] K2 @@ L2 → - ∃∃K1,L1. ⦃h, K1⦄ ⊢ ➡[g] K2 & L = K1 @@ L1. +lemma lpx_fwd_append2: ∀h,g,L,K2,L2. ⦃G, L⦄ ⊢ ➡[h, g] K2 @@ L2 → + ∃∃K1,L1. ⦃h, K1⦄ ⊢ ➡[h, g] K2 & L = K1 @@ L1. /2 width=2 by lpx_sn_fwd_append2/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_aaa.ma index 6fa24124f..41d1386dd 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_aaa.ma @@ -21,8 +21,8 @@ include "basic_2/reduction/lpx_ldrop.ma". (* Properties on atomic arity assignment for terms **************************) (* Note: lemma 500 *) -lemma aaa_cpx_lpx_conf: ∀h,g,L1,T1,A. L1 ⊢ T1 ⁝ A → ∀T2. ⦃h, L1⦄ ⊢ T1 ➡[g] T2 → - ∀L2. ⦃h, L1⦄ ⊢ ➡[g] L2 → L2 ⊢ T2 ⁝ A. +lemma aaa_cpx_lpx_conf: ∀h,g,L1,T1,A. L1 ⊢ T1 ⁝ A → ∀T2. ⦃h, L1⦄ ⊢ T1 ➡[h, g] T2 → + ∀L2. ⦃h, L1⦄ ⊢ ➡[h, g] L2 → L2 ⊢ T2 ⁝ A. #h #g #L1 #T1 #A #H elim H -L1 -T1 -A [ #L1 #k #X #H elim (cpx_inv_sort1 … H) -H // * // @@ -67,13 +67,13 @@ lemma aaa_cpx_lpx_conf: ∀h,g,L1,T1,A. L1 ⊢ T1 ⁝ A → ∀T2. ⦃h, L1⦄ ] qed-. -lemma aaa_cpx_conf: ∀h,g,L,T1,A. L ⊢ T1 ⁝ A → ∀T2. ⦃h, L⦄ ⊢ T1 ➡[g] T2 → L ⊢ T2 ⁝ A. +lemma aaa_cpx_conf: ∀h,g,L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ⦃G, L⦄ ⊢ T2 ⁝ A. /2 width=7 by aaa_cpx_lpx_conf/ qed-. -lemma aaa_lpx_conf: ∀h,g,L1,T,A. L1 ⊢ T ⁝ A → ∀L2. ⦃h, L1⦄ ⊢ ➡[g] L2 → L2 ⊢ T ⁝ A. +lemma aaa_lpx_conf: ∀h,g,L1,T,A. L1 ⊢ T ⁝ A → ∀L2. ⦃h, L1⦄ ⊢ ➡[h, g] L2 → L2 ⊢ T ⁝ A. /2 width=7 by aaa_cpx_lpx_conf/ qed-. -lemma aaa_cpr_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → ∀T2. L ⊢ T1 ➡ T2 → L ⊢ T2 ⁝ A. +lemma aaa_cpr_conf: ∀L,T1,A. ⦃G, L⦄ ⊢ T1 ⁝ A → ∀T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ T2 ⁝ A. /3 width=5 by aaa_cpx_conf, cpr_cpx/ qed-. lemma aaa_lpr_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. L1 ⊢ ➡ L2 → L2 ⊢ T ⁝ A. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/fsup.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/fsup.ma index 53f7b5584..80ec9b0b5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/fsup.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/fsup.ma @@ -12,36 +12,37 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/supterm_4.ma". +include "basic_2/notation/relations/supterm_6.ma". include "basic_2/grammar/cl_weight.ma". include "basic_2/relocation/ldrop.ma". (* SUPCLOSURE ***************************************************************) -inductive fsup: bi_relation lenv term ≝ -| fsup_lref_O : ∀I,L,V. fsup (L.ⓑ{I}V) (#0) L V -| fsup_pair_sn : ∀I,L,V,T. fsup L (②{I}V.T) L V -| fsup_bind_dx : ∀a,I,L,V,T. fsup L (ⓑ{a,I}V.T) (L.ⓑ{I}V) T -| fsup_flat_dx : ∀I,L,V,T. fsup L (ⓕ{I}V.T) L T -| fsup_ldrop_lt: ∀L,K,T,U,d,e. - ⇩[d, e] L ≡ K → ⇧[d, e] T ≡ U → 0 < e → fsup L U K T -| fsup_ldrop : ∀L1,K1,K2,T1,T2,U1,d,e. +(* activate genv *) +inductive fsup: tri_relation genv lenv term ≝ +| fsup_lref_O : ∀I,G,L,V. fsup G (L.ⓑ{I}V) (#0) G L V +| fsup_pair_sn : ∀I,G,L,V,T. fsup G L (②{I}V.T) G L V +| fsup_bind_dx : ∀a,I,G,L,V,T. fsup G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T +| fsup_flat_dx : ∀I,G,L,V,T. fsup G L (ⓕ{I}V.T) G L T +| fsup_ldrop_lt: ∀G,L,K,T,U,d,e. + ⇩[d, e] L ≡ K → ⇧[d, e] T ≡ U → 0 < e → fsup G L U G K T +| fsup_ldrop : ∀G1,G2,L1,K1,K2,T1,T2,U1,d,e. ⇩[d, e] L1 ≡ K1 → ⇧[d, e] T1 ≡ U1 → - fsup K1 T1 K2 T2 → fsup L1 U1 K2 T2 + fsup G1 K1 T1 G2 K2 T2 → fsup G1 L1 U1 G2 K2 T2 . interpretation "structural successor (closure)" - 'SupTerm L1 T1 L2 T2 = (fsup L1 T1 L2 T2). + 'SupTerm G1 L1 T1 G2 L2 T2 = (fsup G1 L1 T1 G2 L2 T2). (* Basic properties *********************************************************) -lemma fsup_lref_S_lt: ∀I,L,K,V,T,i. 0 < i → ⦃L, #(i-1)⦄ ⊃ ⦃K, T⦄ → ⦃L.ⓑ{I}V, #i⦄ ⊃ ⦃K, T⦄. -#I #L #K #V #T #i #Hi #H /3 width=7 by fsup_ldrop, ldrop_ldrop, lift_lref_ge_minus/ (**) (* auto too slow without trace *) +lemma fsup_lref_S_lt: ∀I,G1,G2,L,K,V,T,i. 0 < i → ⦃G1, L, #(i-1)⦄ ⊃ ⦃G2, K, T⦄ → ⦃G1, L.ⓑ{I}V, #i⦄ ⊃ ⦃G2, K, T⦄. +#I #G1 #G2 #L #K #V #T #i #Hi #H /3 width=7 by fsup_ldrop, ldrop_ldrop, lift_lref_ge_minus/ (**) (* auto too slow without trace *) qed. -lemma fsup_lref: ∀I,K,V,i,L. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃L, #i⦄ ⊃ ⦃K, V⦄. -#I #K #V #i @(nat_elim1 i) -i #i #IH #L #H +lemma fsup_lref: ∀I,G,K,V,i,L. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊃ ⦃G, K, V⦄. +#I #G #K #V #i @(nat_elim1 i) -i #i #IH #L #H elim (ldrop_inv_O1_pair2 … H) -H * [ #H1 #H2 destruct // | #I1 #K1 #V1 #HK1 #H #Hi destruct @@ -51,32 +52,32 @@ qed. (* Basic forward lemmas *****************************************************) -lemma fsup_fwd_fw: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ♯{L2, T2} < ♯{L1, T1}. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 // -[ #L #K #T #U #d #e #HLK #HTU #HKL +lemma fsup_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} < ♯{G1, L1, T1}. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // +[ #G #L #K #T #U #d #e #HLK #HTU #HKL lapply (ldrop_fwd_lw_lt … HLK HKL) -HKL -HLK #HKL lapply (lift_fwd_tw … HTU) -d -e #H normalize in ⊢ (?%%); /2 width=1/ -| #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 +| #G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 lapply (ldrop_fwd_lw … HLK1) -HLK1 #HLK1 lapply (lift_fwd_tw … HTU1) -HTU1 #HTU1 @(lt_to_le_to_lt … IHT12) -IHT12 /2 width=1/ ] qed-. -fact fsup_fwd_length_lref1_aux: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → +fact fsup_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ → ∀i. T1 = #i → |L2| < |L1|. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 [1: normalize // |3: #a |5: /2 width=4 by ldrop_fwd_length_lt4/ -|6: #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #i #H destruct +|6: #G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #i #H destruct lapply (ldrop_fwd_length_le4 … HLK1) -HLK1 #HLK1 elim (lift_inv_lref2 … HTU1) -HTU1 * #Hdei #H destruct @(lt_to_le_to_lt … HLK1) /2 width=2/ -] #I #L #V #T #j #H destruct +] #I #G #L #V #T #j #H destruct qed-. -lemma fsup_fwd_length_lref1: ∀L1,L2,T2,i. ⦃L1, #i⦄ ⊃ ⦃L2, T2⦄ → |L2| < |L1|. -/2 width=5 by fsup_fwd_length_lref1_aux/ +lemma fsup_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊃ ⦃G2, L2, T2⦄ → |L2| < |L1|. +/2 width=7 by fsup_fwd_length_lref1_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/fsupq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/fsupq.ma index 4c40afe09..d45a3c83f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/fsupq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/fsupq.ma @@ -12,60 +12,62 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/suptermopt_4.ma". +include "basic_2/notation/relations/suptermopt_6.ma". include "basic_2/relocation/fsup.ma". (* OPTIONAL SUPCLOSURE ******************************************************) -inductive fsupq: bi_relation lenv term ≝ -| fsupq_refl : ∀L,T. fsupq L T L T -| fsupq_lref_O : ∀I,L,V. fsupq (L.ⓑ{I}V) (#0) L V -| fsupq_pair_sn: ∀I,L,V,T. fsupq L (②{I}V.T) L V -| fsupq_bind_dx: ∀a,I,L,V,T. fsupq L (ⓑ{a,I}V.T) (L.ⓑ{I}V) T -| fsupq_flat_dx: ∀I,L,V,T. fsupq L (ⓕ{I}V.T) L T -| fsupq_ldrop : ∀L1,K1,K2,T1,T2,U1,d,e. - ⇩[d, e] L1 ≡ K1 → ⇧[d, e] T1 ≡ U1 → - fsupq K1 T1 K2 T2 → fsupq L1 U1 K2 T2 +(* activate genv *) +inductive fsupq: tri_relation genv lenv term ≝ +| fsupq_refl : ∀G,L,T. fsupq G L T G L T +| fsupq_lref_O : ∀I,G,L,V. fsupq G (L.ⓑ{I}V) (#0) G L V +| fsupq_pair_sn: ∀I,G,L,V,T. fsupq G L (②{I}V.T) G L V +| fsupq_bind_dx: ∀a,I,G,L,V,T. fsupq G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T +| fsupq_flat_dx: ∀I,G, L,V,T. fsupq G L (ⓕ{I}V.T) G L T +| fsupq_ldrop : ∀G1,G2,L1,K1,K2,T1,T2,U1,d,e. + ⇩[d, e] L1 ≡ K1 → ⇧[d, e] T1 ≡ U1 → + fsupq G1 K1 T1 G2 K2 T2 → fsupq G1 L1 U1 G2 K2 T2 . interpretation "optional structural successor (closure)" - 'SupTermOpt L1 T1 L2 T2 = (fsupq L1 T1 L2 T2). + 'SupTermOpt G1 L1 T1 G2 L2 T2 = (fsupq G1 L1 T1 G2 L2 T2). (* Basic properties *********************************************************) -lemma fsup_fsupq: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 // /2 width=7/ qed. +lemma fsup_fsupq: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 // /2 width=7/ qed. (* Basic properties *********************************************************) -lemma fsupq_lref_S_lt: ∀I,L,K,V,T,i. 0 < i → ⦃L, #(i-1)⦄ ⊃⸮ ⦃K, T⦄ → ⦃L.ⓑ{I}V, #i⦄ ⊃⸮ ⦃K, T⦄. +lemma fsupq_lref_S_lt: ∀I,G1,G2,L,K,V,T,i. + 0 < i → ⦃G1, L, #(i-1)⦄ ⊃⸮ ⦃G2, K, T⦄ → ⦃G1, L.ⓑ{I}V, #i⦄ ⊃⸮ ⦃G2, K, T⦄. /3 width=7/ qed. -lemma fsupq_lref: ∀I,K,V,i,L. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃L, #i⦄ ⊃⸮ ⦃K, V⦄. +lemma fsupq_lref: ∀I,G,K,V,i,L. ⇩[0, i] L ≡ K.ⓑ{I}V → ⦃G, L, #i⦄ ⊃⸮ ⦃G, K, V⦄. /3 width=2/ qed. (* Basic forward lemmas *****************************************************) -lemma fsupq_fwd_fw: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → ♯{L2, T2} ≤ ♯{L1, T1}. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 // [1,2,3: /2 width=1/ ] -#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 +lemma fsupq_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // [1,2,3: /2 width=1/ ] +#G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 lapply (ldrop_fwd_lw … HLK1) -HLK1 #HLK1 lapply (lift_fwd_tw … HTU1) -HTU1 #HTU1 @(transitive_le … IHT12) -IHT12 /2 width=1/ qed-. -fact fsupq_fwd_length_lref1_aux: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → +fact fsupq_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄ → ∀i. T1 = #i → |L2| ≤ |L1|. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 // -[ #a #I #L #V #T #j #H destruct -| #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #i #H destruct +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // +[ #a #I #G #L #V #T #j #H destruct +| #G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #i #H destruct lapply (ldrop_fwd_length_le4 … HLK1) -HLK1 #HLK1 elim (lift_inv_lref2 … HTU1) -HTU1 * #Hdei #H destruct @(transitive_le … HLK1) /2 width=2/ ] qed-. -lemma fsupq_fwd_length_lref1: ∀L1,L2,T2,i. ⦃L1, #i⦄ ⊃⸮ ⦃L2, T2⦄ → |L2| ≤ |L1|. -/2 width=5 by fsupq_fwd_length_lref1_aux/ +lemma fsupq_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊃⸮ ⦃G2, L2, T2⦄ → |L2| ≤ |L1|. +/2 width=7 by fsupq_fwd_length_lref1_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/fsupq_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/fsupq_alt.ma index 360972607..e9bc6c533 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/fsupq_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/fsupq_alt.ma @@ -12,40 +12,40 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/suptermoptalt_4.ma". +include "basic_2/notation/relations/suptermoptalt_6.ma". include "basic_2/relocation/fsupq.ma". (* OPTIONAL SUPCLOSURE ******************************************************) (* alternative definition of fsupq *) -definition fsupqa: bi_relation lenv term ≝ bi_RC … fsup. +definition fsupqa: tri_relation genv lenv term ≝ tri_RC … fsup. interpretation "optional structural successor (closure) alternative" - 'SupTermOptAlt L1 T1 L2 T2 = (fsupqa L1 T1 L2 T2). + 'SupTermOptAlt G1 L1 T1 G2 L2 T2 = (fsupqa G1 L1 T1 G2 L2 T2). (* Basic properties *********************************************************) -lemma fsupqa_refl: bi_reflexive … fsupqa. +lemma fsupqa_refl: tri_reflexive … fsupqa. // qed. -lemma fsupqa_ldrop: ∀K1,K2,T1,T2. ⦃K1, T1⦄ ⊃⊃⸮ ⦃K2, T2⦄ → +lemma fsupqa_ldrop: ∀G1,G2,K1,K2,T1,T2. ⦃G1, K1, T1⦄ ⊃⊃⸮ ⦃G2, K2, T2⦄ → ∀L1,d,e. ⇩[d, e] L1 ≡ K1 → - ∀U1. ⇧[d, e] T1 ≡ U1 → ⦃L1, U1⦄ ⊃⊃⸮ ⦃K2, T2⦄. -#K1 #K2 #T1 #T2 * [ /3 width=7/ ] * #H1 #H2 destruct + ∀U1. ⇧[d, e] T1 ≡ U1 → ⦃G1, L1, U1⦄ ⊃⊃⸮ ⦃G2, K2, T2⦄. +#G1 #G2 #K1 #K2 #T1 #T2 * [ /3 width=7/ ] * #H1 #H2 #H3 destruct #L1 #d #e #HLK #U1 #HTU elim (eq_or_gt e) [2: /3 width=5/ ] #H destruct >(ldrop_inv_O2 … HLK) -L1 >(lift_inv_O2 … HTU) -T2 -d // qed. (* Main properties **********************************************************) -theorem fsupq_fsupqa: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃⊃⸮ ⦃L2, T2⦄. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 // /2 width=1/ /2 width=7/ +theorem fsupq_fsupqa: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊃⊃⸮ ⦃G2, L2, T2⦄. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // /2 width=1/ /2 width=7/ qed. (* Main inversion properties ************************************************) -theorem fsupqa_inv_fsupq: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃⊃⸮ ⦃L2, T2⦄ → ⦃L1, T1⦄ ⊃⸮ ⦃L2, T2⦄. -#L1 #L2 #T1 #T2 #H elim H -H /2 width=1/ -* #H1 #H2 destruct // +theorem fsupqa_inv_fsupq: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃⊃⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H /2 width=1/ +* #H1 #H2 #H3 destruct // qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/gdrop.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/gdrop.ma deleted file mode 100644 index 277e464b7..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/gdrop.ma +++ /dev/null @@ -1,81 +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/rdrop_3.ma". -include "basic_2/grammar/genv.ma". - -(* GLOBAL ENVIRONMENT SLICING ***********************************************) - -inductive gdrop (e:nat): relation genv ≝ -| gdrop_gt: ∀G. |G| ≤ e → gdrop e G (⋆) -| gdrop_eq: ∀G. |G| = e + 1 → gdrop e G G -| gdrop_lt: ∀I,G1,G2,V. e < |G1| → gdrop e G1 G2 → gdrop e (G1. ⓑ{I} V) G2 -. - -interpretation "global slicing" - 'RDrop e G1 G2 = (gdrop e G1 G2). - -(* basic inversion lemmas ***************************************************) - -lemma gdrop_inv_gt: ∀G1,G2,e. ⇩[e] G1 ≡ G2 → |G1| ≤ e → G2 = ⋆. -#G1 #G2 #e * -G1 -G2 // -[ #G #H >H -H >commutative_plus #H - lapply (le_plus_to_le_r … 0 H) -H #H - lapply (le_n_O_to_eq … H) -H #H destruct -| #I #G1 #G2 #V #H1 #_ #H2 - lapply (le_to_lt_to_lt … H2 H1) -H2 -H1 normalize in ⊢ (? % ? → ?); >commutative_plus #H - lapply (lt_plus_to_lt_l … 0 H) -H #H - elim (lt_zero_false … H) -] -qed-. - -lemma gdrop_inv_eq: ∀G1,G2,e. ⇩[e] G1 ≡ G2 → |G1| = e + 1 → G1 = G2. -#G1 #G2 #e * -G1 -G2 // -[ #G #H1 #H2 >H2 in H1; -H2 >commutative_plus #H - lapply (le_plus_to_le_r … 0 H) -H #H - lapply (le_n_O_to_eq … H) -H #H destruct -| #I #G1 #G2 #V #H1 #_ normalize #H2 - <(injective_plus_l … H2) in H1; -H2 #H - elim (lt_refl_false … H) -] -qed-. - -fact gdrop_inv_lt_aux: ∀I,G,G1,G2,V,e. ⇩[e] G ≡ G2 → G = G1. ⓑ{I} V → - e < |G1| → ⇩[e] G1 ≡ G2. -#I #G #G1 #G2 #V #e * -G -G2 -[ #G #H1 #H destruct #H2 - lapply (le_to_lt_to_lt … H1 H2) -H1 -H2 normalize in ⊢ (? % ? → ?); >commutative_plus #H - lapply (lt_plus_to_lt_l … 0 H) -H #H - elim (lt_zero_false … H) -| #G #H1 #H2 destruct >(injective_plus_l … H1) -H1 #H - elim (lt_refl_false … H) -| #J #G #G2 #W #_ #HG2 #H destruct // -] -qed. - -lemma gdrop_inv_lt: ∀I,G1,G2,V,e. - ⇩[e] G1. ⓑ{I} V ≡ G2 → e < |G1| → ⇩[e] G1 ≡ G2. -/2 width=5/ qed-. - -(* Basic properties *********************************************************) - -lemma gdrop_total: ∀e,G1. ∃G2. ⇩[e] G1 ≡ G2. -#e #G1 elim G1 -G1 /3 width=2/ -#I #V #G1 * #G2 #HG12 -elim (lt_or_eq_or_gt e (|G1|)) #He -[ /3 width=2/ -| destruct /3 width=2/ -| @ex_intro [2: @gdrop_gt normalize /2 width=1/ | skip ] (**) (* explicit constructor *) -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/gdrop_gdrop.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/gdrop_gdrop.ma deleted file mode 100644 index a133e29c6..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/gdrop_gdrop.ma +++ /dev/null @@ -1,40 +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/relocation/gdrop.ma". - -(* GLOBAL ENVIRONMENT SLICING ***********************************************) - -(* Main properties **********************************************************) - -theorem gdrop_mono: ∀G,G1,e. ⇩[e] G ≡ G1 → ∀G2. ⇩[e] G ≡ G2 → G1 = G2. -#G #G1 #e #H elim H -G -G1 -[ #G #He #G2 #H - >(gdrop_inv_gt … H He) -H -He // -| #G #He #G2 #H - >(gdrop_inv_eq … H He) -H -He // -| #I #G #G1 #V #He #_ #IHG1 #G2 #H - lapply (gdrop_inv_lt … H He) -H -He /2 width=1/ -] -qed-. - -lemma gdrop_dec: ∀G1,G2,e. Decidable (⇩[e] G1 ≡ G2). -#G1 #G2 #e -elim (gdrop_total e G1) #G #HG1 -elim (genv_eq_dec G G2) #HG2 -[ destruct /2 width=1/ -| @or_intror #HG12 - lapply (gdrop_mono … HG1 … HG12) -HG1 -HG12 /2 width=1/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/gget.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/gget.ma new file mode 100644 index 000000000..5ef11dc4b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/gget.ma @@ -0,0 +1,81 @@ +(**************************************************************************) +(* ___ *) +(* ||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/rdrop_3.ma". +include "basic_2/grammar/genv.ma". + +(* GLOBAL ENVIRONMENT READING ***********************************************) + +inductive gget (e:nat): relation genv ≝ +| gget_gt: ∀G. |G| ≤ e → gget e G (⋆) +| gget_eq: ∀G. |G| = e + 1 → gget e G G +| gget_lt: ∀I,G1,G2,V. e < |G1| → gget e G1 G2 → gget e (G1. ⓑ{I} V) G2 +. + +interpretation "global reading" + 'RDrop e G1 G2 = (gget e G1 G2). + +(* basic inversion lemmas ***************************************************) + +lemma gget_inv_gt: ∀G1,G2,e. ⇩[e] G1 ≡ G2 → |G1| ≤ e → G2 = ⋆. +#G1 #G2 #e * -G1 -G2 // +[ #G #H >H -H >commutative_plus #H + lapply (le_plus_to_le_r … 0 H) -H #H + lapply (le_n_O_to_eq … H) -H #H destruct +| #I #G1 #G2 #V #H1 #_ #H2 + lapply (le_to_lt_to_lt … H2 H1) -H2 -H1 normalize in ⊢ (? % ? → ?); >commutative_plus #H + lapply (lt_plus_to_lt_l … 0 H) -H #H + elim (lt_zero_false … H) +] +qed-. + +lemma gget_inv_eq: ∀G1,G2,e. ⇩[e] G1 ≡ G2 → |G1| = e + 1 → G1 = G2. +#G1 #G2 #e * -G1 -G2 // +[ #G #H1 #H2 >H2 in H1; -H2 >commutative_plus #H + lapply (le_plus_to_le_r … 0 H) -H #H + lapply (le_n_O_to_eq … H) -H #H destruct +| #I #G1 #G2 #V #H1 #_ normalize #H2 + <(injective_plus_l … H2) in H1; -H2 #H + elim (lt_refl_false … H) +] +qed-. + +fact gget_inv_lt_aux: ∀I,G,G1,G2,V,e. ⇩[e] G ≡ G2 → G = G1. ⓑ{I} V → + e < |G1| → ⇩[e] G1 ≡ G2. +#I #G #G1 #G2 #V #e * -G -G2 +[ #G #H1 #H destruct #H2 + lapply (le_to_lt_to_lt … H1 H2) -H1 -H2 normalize in ⊢ (? % ? → ?); >commutative_plus #H + lapply (lt_plus_to_lt_l … 0 H) -H #H + elim (lt_zero_false … H) +| #G #H1 #H2 destruct >(injective_plus_l … H1) -H1 #H + elim (lt_refl_false … H) +| #J #G #G2 #W #_ #HG2 #H destruct // +] +qed-. + +lemma gget_inv_lt: ∀I,G1,G2,V,e. + ⇩[e] G1. ⓑ{I} V ≡ G2 → e < |G1| → ⇩[e] G1 ≡ G2. +/2 width=5 by gget_inv_lt_aux/ qed-. + +(* Basic properties *********************************************************) + +lemma gget_total: ∀e,G1. ∃G2. ⇩[e] G1 ≡ G2. +#e #G1 elim G1 -G1 /3 width=2/ +#I #V #G1 * #G2 #HG12 +elim (lt_or_eq_or_gt e (|G1|)) #He +[ /3 width=2/ +| destruct /3 width=2/ +| @ex_intro [2: @gget_gt normalize /2 width=1/ | skip ] (**) (* explicit constructor *) +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/gget_gget.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/gget_gget.ma new file mode 100644 index 000000000..6c1001b6a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/gget_gget.ma @@ -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/relocation/gget.ma". + +(* GLOBAL ENVIRONMENT READING ***********************************************) + +(* Main properties **********************************************************) + +theorem gget_mono: ∀G,G1,e. ⇩[e] G ≡ G1 → ∀G2. ⇩[e] G ≡ G2 → G1 = G2. +#G #G1 #e #H elim H -G -G1 +[ #G #He #G2 #H + >(gget_inv_gt … H He) -H -He // +| #G #He #G2 #H + >(gget_inv_eq … H He) -H -He // +| #I #G #G1 #V #He #_ #IHG1 #G2 #H + lapply (gget_inv_lt … H He) -H -He /2 width=1/ +] +qed-. + +lemma gget_dec: ∀G1,G2,e. Decidable (⇩[e] G1 ≡ G2). +#G1 #G2 #e +elim (gget_total e G1) #G #HG1 +elim (genv_eq_dec G G2) #HG2 +[ destruct /2 width=1/ +| @or_intror #HG12 + lapply (gget_mono … HG1 … HG12) -HG1 -HG12 /2 width=1/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma index 63584de4a..4c4f1e572 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa.ma @@ -34,7 +34,7 @@ interpretation "atomic arity assignment (term)" (* Basic inversion lemmas ***************************************************) -fact aaa_inv_sort_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀k. T = ⋆k → A = ⓪. +fact aaa_inv_sort_aux: ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀k. T = ⋆k → A = ⓪. #L #T #A * -L -T -A [ // | #I #L #K #V #B #i #_ #_ #k #H destruct @@ -45,10 +45,10 @@ fact aaa_inv_sort_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀k. T = ⋆k → A = ⓪. ] qed. -lemma aaa_inv_sort: ∀L,A,k. L ⊢ ⋆k ⁝ A → A = ⓪. +lemma aaa_inv_sort: ∀L,A,k. ⦃G, L⦄ ⊢ ⋆k ⁝ A → A = ⓪. /2 width=5/ qed-. -fact aaa_inv_lref_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀i. T = #i → +fact aaa_inv_lref_aux: ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀i. T = #i → ∃∃I,K,V. ⇩[0, i] L ≡ K. ⓑ{I} V & K ⊢ V ⁝ A. #L #T #A * -L -T -A [ #L #k #i #H destruct @@ -60,11 +60,11 @@ fact aaa_inv_lref_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀i. T = #i → ] qed. -lemma aaa_inv_lref: ∀L,A,i. L ⊢ #i ⁝ A → +lemma aaa_inv_lref: ∀L,A,i. ⦃G, L⦄ ⊢ #i ⁝ A → ∃∃I,K,V. ⇩[0, i] L ≡ K. ⓑ{I} V & K ⊢ V ⁝ A. /2 width=3/ qed-. -fact aaa_inv_gref_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀p. T = §p → ⊥. +fact aaa_inv_gref_aux: ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀p. T = §p → ⊥. #L #T #A * -L -T -A [ #L #k #q #H destruct | #I #L #K #V #B #i #HLK #HB #q #H destruct @@ -75,11 +75,11 @@ fact aaa_inv_gref_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀p. T = §p → ⊥. ] qed. -lemma aaa_inv_gref: ∀L,A,p. L ⊢ §p ⁝ A → ⊥. +lemma aaa_inv_gref: ∀L,A,p. ⦃G, L⦄ ⊢ §p ⁝ A → ⊥. /2 width=6/ qed-. -fact aaa_inv_abbr_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀a,W,U. T = ⓓ{a}W. U → - ∃∃B. L ⊢ W ⁝ B & L. ⓓW ⊢ U ⁝ A. +fact aaa_inv_abbr_aux: ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀a,W,U. T = ⓓ{a}W. U → + ∃∃B. ⦃G, L⦄ ⊢ W ⁝ B & L. ⓓW ⊢ U ⁝ A. #L #T #A * -L -T -A [ #L #k #a #W #U #H destruct | #I #L #K #V #B #i #_ #_ #a #W #U #H destruct @@ -90,12 +90,12 @@ fact aaa_inv_abbr_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀a,W,U. T = ⓓ{a}W. U → ] qed. -lemma aaa_inv_abbr: ∀a,L,V,T,A. L ⊢ ⓓ{a}V. T ⁝ A → - ∃∃B. L ⊢ V ⁝ B & L. ⓓV ⊢ T ⁝ A. +lemma aaa_inv_abbr: ∀a,L,V,T,A. ⦃G, L⦄ ⊢ ⓓ{a}V. T ⁝ A → + ∃∃B. ⦃G, L⦄ ⊢ V ⁝ B & L. ⓓV ⊢ T ⁝ A. /2 width=4/ qed-. -fact aaa_inv_abst_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀a,W,U. T = ⓛ{a}W. U → - ∃∃B1,B2. L ⊢ W ⁝ B1 & L. ⓛW ⊢ U ⁝ B2 & A = ②B1. B2. +fact aaa_inv_abst_aux: ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀a,W,U. T = ⓛ{a}W. U → + ∃∃B1,B2. ⦃G, L⦄ ⊢ W ⁝ B1 & L. ⓛW ⊢ U ⁝ B2 & A = ②B1. B2. #L #T #A * -L -T -A [ #L #k #a #W #U #H destruct | #I #L #K #V #B #i #_ #_ #a #W #U #H destruct @@ -106,12 +106,12 @@ fact aaa_inv_abst_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀a,W,U. T = ⓛ{a}W. U → ] qed. -lemma aaa_inv_abst: ∀a,L,W,T,A. L ⊢ ⓛ{a}W. T ⁝ A → - ∃∃B1,B2. L ⊢ W ⁝ B1 & L. ⓛW ⊢ T ⁝ B2 & A = ②B1. B2. +lemma aaa_inv_abst: ∀a,L,W,T,A. ⦃G, L⦄ ⊢ ⓛ{a}W. T ⁝ A → + ∃∃B1,B2. ⦃G, L⦄ ⊢ W ⁝ B1 & L. ⓛW ⊢ T ⁝ B2 & A = ②B1. B2. /2 width=4/ qed-. -fact aaa_inv_appl_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀W,U. T = ⓐW. U → - ∃∃B. L ⊢ W ⁝ B & L ⊢ U ⁝ ②B. A. +fact aaa_inv_appl_aux: ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀W,U. T = ⓐW. U → + ∃∃B. ⦃G, L⦄ ⊢ W ⁝ B & ⦃G, L⦄ ⊢ U ⁝ ②B. A. #L #T #A * -L -T -A [ #L #k #W #U #H destruct | #I #L #K #V #B #i #_ #_ #W #U #H destruct @@ -122,12 +122,12 @@ fact aaa_inv_appl_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀W,U. T = ⓐW. U → ] qed. -lemma aaa_inv_appl: ∀L,V,T,A. L ⊢ ⓐV. T ⁝ A → - ∃∃B. L ⊢ V ⁝ B & L ⊢ T ⁝ ②B. A. +lemma aaa_inv_appl: ∀L,V,T,A. ⦃G, L⦄ ⊢ ⓐV. T ⁝ A → + ∃∃B. ⦃G, L⦄ ⊢ V ⁝ B & ⦃G, L⦄ ⊢ T ⁝ ②B. A. /2 width=3/ qed-. -fact aaa_inv_cast_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀W,U. T = ⓝW. U → - L ⊢ W ⁝ A ∧ L ⊢ U ⁝ A. +fact aaa_inv_cast_aux: ∀L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀W,U. T = ⓝW. U → + ⦃G, L⦄ ⊢ W ⁝ A ∧ ⦃G, L⦄ ⊢ U ⁝ A. #L #T #A * -L -T -A [ #L #k #W #U #H destruct | #I #L #K #V #B #i #_ #_ #W #U #H destruct @@ -138,6 +138,6 @@ fact aaa_inv_cast_aux: ∀L,T,A. L ⊢ T ⁝ A → ∀W,U. T = ⓝW. U → ] qed. -lemma aaa_inv_cast: ∀L,W,T,A. L ⊢ ⓝW. T ⁝ A → - L ⊢ W ⁝ A ∧ L ⊢ T ⁝ A. +lemma aaa_inv_cast: ∀L,W,T,A. ⦃G, L⦄ ⊢ ⓝW. T ⁝ A → + ⦃G, L⦄ ⊢ W ⁝ A ∧ ⦃G, L⦄ ⊢ T ⁝ A. /2 width=3/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_aaa.ma index fce1c02e4..4bf7888c0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_aaa.ma @@ -19,7 +19,7 @@ include "basic_2/static/aaa.ma". (* Main properties **********************************************************) -theorem aaa_mono: ∀L,T,A1. L ⊢ T ⁝ A1 → ∀A2. L ⊢ T ⁝ A2 → A1 = A2. +theorem aaa_mono: ∀L,T,A1. ⦃G, L⦄ ⊢ T ⁝ A1 → ∀A2. ⦃G, L⦄ ⊢ T ⁝ A2 → A1 = A2. #L #T #A1 #H elim H -L -T -A1 [ #L #k #A2 #H >(aaa_inv_sort … H) -H // diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ssta.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ssta.ma index 6f4f40872..9421d7436 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ssta.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/ssta.ma @@ -35,11 +35,11 @@ interpretation "stratified static type assignment (term)" 'StaticType h g L T U l = (ssta h g l L T U). definition ssta_step: ∀h. sd h → lenv → relation term ≝ λh,g,L,T,U. - ∃l. ⦃h, L⦄ ⊢ T •[g] ⦃l+1, U⦄. + ∃l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, U⦄. (* Basic inversion lemmas ************************************************) -fact ssta_inv_sort1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → ∀k0. T = ⋆k0 → +fact ssta_inv_sort1_aux: ∀h,g,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀k0. T = ⋆k0 → deg h g k0 l ∧ U = ⋆(next h k0). #h #g #L #T #U #l * -L -T -U -l [ #L #k #l #Hkl #k0 #H destruct /2 width=1/ @@ -51,15 +51,15 @@ fact ssta_inv_sort1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → qed. (* Basic_1: was just: sty0_gen_sort *) -lemma ssta_inv_sort1: ∀h,g,L,U,k,l. ⦃h, L⦄ ⊢ ⋆k •[g] ⦃l, U⦄ → +lemma ssta_inv_sort1: ∀h,g,L,U,k,l. ⦃G, L⦄ ⊢ ⋆k •[h, g] ⦃l, U⦄ → deg h g k l ∧ U = ⋆(next h k). /2 width=4/ qed-. -fact ssta_inv_lref1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → ∀j. T = #j → - (∃∃K,V,W. ⇩[0, j] L ≡ K. ⓓV & ⦃h, K⦄ ⊢ V •[g] ⦃l, W⦄ & +fact ssta_inv_lref1_aux: ∀h,g,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀j. T = #j → + (∃∃K,V,W. ⇩[0, j] L ≡ K. ⓓV & ⦃h, K⦄ ⊢ V •[h, g] ⦃l, W⦄ & ⇧[0, j + 1] W ≡ U ) ∨ - (∃∃K,W,V,l0. ⇩[0, j] L ≡ K. ⓛW & ⦃h, K⦄ ⊢ W •[g] ⦃l0, V⦄ & + (∃∃K,W,V,l0. ⇩[0, j] L ≡ K. ⓛW & ⦃h, K⦄ ⊢ W •[h, g] ⦃l0, V⦄ & ⇧[0, j + 1] W ≡ U & l = l0 + 1 ). #h #g #L #T #U #l * -L -T -U -l @@ -73,16 +73,16 @@ fact ssta_inv_lref1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → qed. (* Basic_1: was just: sty0_gen_lref *) -lemma ssta_inv_lref1: ∀h,g,L,U,i,l. ⦃h, L⦄ ⊢ #i •[g] ⦃l, U⦄ → - (∃∃K,V,W. ⇩[0, i] L ≡ K. ⓓV & ⦃h, K⦄ ⊢ V •[g] ⦃l, W⦄ & +lemma ssta_inv_lref1: ∀h,g,L,U,i,l. ⦃G, L⦄ ⊢ #i •[h, g] ⦃l, U⦄ → + (∃∃K,V,W. ⇩[0, i] L ≡ K. ⓓV & ⦃h, K⦄ ⊢ V •[h, g] ⦃l, W⦄ & ⇧[0, i + 1] W ≡ U ) ∨ - (∃∃K,W,V,l0. ⇩[0, i] L ≡ K. ⓛW & ⦃h, K⦄ ⊢ W •[g] ⦃l0, V⦄ & + (∃∃K,W,V,l0. ⇩[0, i] L ≡ K. ⓛW & ⦃h, K⦄ ⊢ W •[h, g] ⦃l0, V⦄ & ⇧[0, i + 1] W ≡ U & l = l0 + 1 ). /2 width=3/ qed-. -fact ssta_inv_gref1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → ∀p0. T = §p0 → ⊥. +fact ssta_inv_gref1_aux: ∀h,g,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀p0. T = §p0 → ⊥. #h #g #L #T #U #l * -L -T -U -l [ #L #k #l #_ #p0 #H destruct | #L #K #V #W #U #i #l #_ #_ #_ #p0 #H destruct @@ -92,12 +92,12 @@ fact ssta_inv_gref1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → | #L #W #T #U #l #_ #p0 #H destruct qed. -lemma ssta_inv_gref1: ∀h,g,L,U,p,l. ⦃h, L⦄ ⊢ §p •[g] ⦃l, U⦄ → ⊥. +lemma ssta_inv_gref1: ∀h,g,L,U,p,l. ⦃G, L⦄ ⊢ §p •[h, g] ⦃l, U⦄ → ⊥. /2 width=9/ qed-. -fact ssta_inv_bind1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → +fact ssta_inv_bind1_aux: ∀h,g,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀a,I,X,Y. T = ⓑ{a,I}Y.X → - ∃∃Z. ⦃h, L.ⓑ{I}Y⦄ ⊢ X •[g] ⦃l, Z⦄ & U = ⓑ{a,I}Y.Z. + ∃∃Z. ⦃h, L.ⓑ{I}Y⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓑ{a,I}Y.Z. #h #g #L #T #U #l * -L -T -U -l [ #L #k #l #_ #a #I #X #Y #H destruct | #L #K #V #W #U #i #l #_ #_ #_ #a #I #X #Y #H destruct @@ -109,12 +109,12 @@ fact ssta_inv_bind1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → qed. (* Basic_1: was just: sty0_gen_bind *) -lemma ssta_inv_bind1: ∀h,g,a,I,L,Y,X,U,l. ⦃h, L⦄ ⊢ ⓑ{a,I}Y.X •[g] ⦃l, U⦄ → - ∃∃Z. ⦃h, L.ⓑ{I}Y⦄ ⊢ X •[g] ⦃l, Z⦄ & U = ⓑ{a,I}Y.Z. +lemma ssta_inv_bind1: ∀h,g,a,I,L,Y,X,U,l. ⦃G, L⦄ ⊢ ⓑ{a,I}Y.X •[h, g] ⦃l, U⦄ → + ∃∃Z. ⦃h, L.ⓑ{I}Y⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓑ{a,I}Y.Z. /2 width=3/ qed-. -fact ssta_inv_appl1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → ∀X,Y. T = ⓐY.X → - ∃∃Z. ⦃h, L⦄ ⊢ X •[g] ⦃l, Z⦄ & U = ⓐY.Z. +fact ssta_inv_appl1_aux: ∀h,g,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ∀X,Y. T = ⓐY.X → + ∃∃Z. ⦃G, L⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓐY.Z. #h #g #L #T #U #l * -L -T -U -l [ #L #k #l #_ #X #Y #H destruct | #L #K #V #W #U #i #l #_ #_ #_ #X #Y #H destruct @@ -126,12 +126,12 @@ fact ssta_inv_appl1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → qed. (* Basic_1: was just: sty0_gen_appl *) -lemma ssta_inv_appl1: ∀h,g,L,Y,X,U,l. ⦃h, L⦄ ⊢ ⓐY.X •[g] ⦃l, U⦄ → - ∃∃Z. ⦃h, L⦄ ⊢ X •[g] ⦃l, Z⦄ & U = ⓐY.Z. +lemma ssta_inv_appl1: ∀h,g,L,Y,X,U,l. ⦃G, L⦄ ⊢ ⓐY.X •[h, g] ⦃l, U⦄ → + ∃∃Z. ⦃G, L⦄ ⊢ X •[h, g] ⦃l, Z⦄ & U = ⓐY.Z. /2 width=3/ qed-. -fact ssta_inv_cast1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → - ∀X,Y. T = ⓝY.X → ⦃h, L⦄ ⊢ X •[g] ⦃l, U⦄. +fact ssta_inv_cast1_aux: ∀h,g,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → + ∀X,Y. T = ⓝY.X → ⦃G, L⦄ ⊢ X •[h, g] ⦃l, U⦄. #h #g #L #T #U #l * -L -T -U -l [ #L #k #l #_ #X #Y #H destruct | #L #K #V #W #U #l #i #_ #_ #_ #X #Y #H destruct @@ -143,6 +143,6 @@ fact ssta_inv_cast1_aux: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → qed. (* Basic_1: was just: sty0_gen_cast *) -lemma ssta_inv_cast1: ∀h,g,L,X,Y,U,l. ⦃h, L⦄ ⊢ ⓝY.X •[g] ⦃l, U⦄ → - ⦃h, L⦄ ⊢ X •[g] ⦃l, U⦄. +lemma ssta_inv_cast1: ∀h,g,L,X,Y,U,l. ⦃G, L⦄ ⊢ ⓝY.X •[h, g] ⦃l, U⦄ → + ⦃G, L⦄ ⊢ X •[h, g] ⦃l, U⦄. /2 width=4/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_aaa.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_aaa.ma index eafec7430..0f035cf7c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_aaa.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_aaa.ma @@ -19,7 +19,7 @@ include "basic_2/static/ssta.ma". (* Properties on atomic arity assignment for terms **************************) -lemma ssta_aaa: ∀h,g,L,T,A. L ⊢ T ⁝ A → ∀U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → L ⊢ U ⁝ A. +lemma ssta_aaa: ∀h,g,L,T,A. ⦃G, L⦄ ⊢ T ⁝ A → ∀U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → ⦃G, L⦄ ⊢ U ⁝ A. #h #g #L #T #A #H elim H -L -T -A [ #L #k #U #l #H elim (ssta_inv_sort1 … H) -H #_ #H destruct // diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lift.ma index 10585d61d..00e512498 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lift.ma @@ -20,9 +20,9 @@ include "basic_2/static/ssta.ma". (* Properties on relocation *************************************************) (* Basic_1: was just: sty0_lift *) -lemma ssta_lift: ∀h,g,L1,T1,U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ → +lemma ssta_lift: ∀h,g,L1,T1,U1,l. ⦃h, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ → ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 → - ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄. + ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃h, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄. #h #g #L1 #T1 #U1 #l #H elim H -L1 -T1 -U1 -l [ #L1 #k #l #Hkl #L2 #d #e #HL21 #X1 #H1 #X2 #H2 >(lift_inv_sort1 … H1) -X1 @@ -60,9 +60,9 @@ lemma ssta_lift: ∀h,g,L1,T1,U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ → qed. (* Note: apparently this was missing in basic_1 *) -lemma ssta_inv_lift1: ∀h,g,L2,T2,U2,l. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄ → +lemma ssta_inv_lift1: ∀h,g,L2,T2,U2,l. ⦃h, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄ → ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ & ⇧[d, e] U1 ≡ U2. + ∃∃U1. ⦃h, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ & ⇧[d, e] U1 ≡ U2. #h #g #L2 #T2 #U2 #l #H elim H -L2 -T2 -U2 -l [ #L2 #k #l #Hkl #L1 #d #e #_ #X #H >(lift_inv_sort2 … H) -X /3 width=3/ @@ -105,8 +105,8 @@ qed. (* Advanced forvard lemmas **************************************************) (* Basic_1: was just: sty0_correct *) -lemma ssta_fwd_correct: ∀h,g,L,T,U,l. ⦃h, L⦄ ⊢ T •[g] ⦃l, U⦄ → - ∃T0. ⦃h, L⦄ ⊢ U •[g] ⦃l-1, T0⦄. +lemma ssta_fwd_correct: ∀h,g,L,T,U,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l, U⦄ → + ∃T0. ⦃G, L⦄ ⊢ U •[h, g] ⦃l-1, T0⦄. #h #g #L #T #U #l #H elim H -L -T -U -l [ /4 width=2/ | #L #K #V #W #W0 #i #l #HLK #_ #HW0 * #V0 #HWV0 diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_ssta.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_ssta.ma index 617ef508f..87e0d2168 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_ssta.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_ssta.ma @@ -19,8 +19,8 @@ include "basic_2/static/ssta_lift.ma". (* Main properties **********************************************************) (* Note: apparently this was missing in basic_1 *) -theorem ssta_mono: ∀h,g,L,T,U1,l1. ⦃h, L⦄ ⊢ T •[g] ⦃l1, U1⦄ → - ∀U2,l2. ⦃h, L⦄ ⊢ T •[g] ⦃l2, U2⦄ → l1 = l2 ∧ U1 = U2. +theorem ssta_mono: ∀h,g,L,T,U1,l1. ⦃G, L⦄ ⊢ T •[h, g] ⦃l1, U1⦄ → + ∀U2,l2. ⦃G, L⦄ ⊢ T •[h, g] ⦃l2, U2⦄ → l1 = l2 ∧ U1 = U2. #h #g #L #T #U1 #l1 #H elim H -L -T -U1 -l1 [ #L #k #l #Hkl #X #l2 #H elim (ssta_inv_sort1 … H) -H #Hkl2 #H destruct @@ -49,7 +49,7 @@ qed-. (* Advanced inversion lemmas ************************************************) -lemma ssta_inv_refl_pos: ∀h,g,L,T,l. ⦃h, L⦄ ⊢ T •[g] ⦃l+1, T⦄ → ⊥. +lemma ssta_inv_refl_pos: ∀h,g,L,T,l. ⦃G, L⦄ ⊢ T •[h, g] ⦃l+1, T⦄ → ⊥. #h #g #L #T #l #HTT elim (ssta_fwd_correct … HTT) (lift_mono … HX … HU12) -X // @@ -42,9 +42,9 @@ qed. (* Inversion lemmas on relocation *******************************************) -lemma sstas_inv_lift1: ∀h,g,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 •*[g] U2 → +lemma sstas_inv_lift1: ∀h,g,L2,T2,U2. ⦃h, L2⦄ ⊢ T2 •*[h, g] U2 → ∀L1,d,e. ⇩[d, e] L2 ≡ L1 → ∀T1. ⇧[d, e] T1 ≡ T2 → - ∃∃U1. ⦃h, L1⦄ ⊢ T1 •*[g] U1 & ⇧[d, e] U1 ≡ U2. + ∃∃U1. ⦃h, L1⦄ ⊢ T1 •*[h, g] U1 & ⇧[d, e] U1 ≡ U2. #h #g #L2 #T2 #U2 #H @(sstas_ind_dx … H) -T2 /2 width=3/ #T0 #U0 #l0 #HTU0 #_ #IHU01 #L1 #d #e #HL21 #U1 #HU12 elim (ssta_inv_lift1 … HTU0 … HL21 … HU12) -HTU0 -HU12 #U #HU1 #HU0 diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/sstas_sstas.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/sstas_sstas.ma index 0b0dcca42..2d7febf1d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/sstas_sstas.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/unfold/sstas_sstas.ma @@ -19,8 +19,8 @@ include "basic_2/unfold/sstas.ma". (* Advanced inversion lemmas ************************************************) -lemma sstas_inv_O: ∀h,g,L,T,U. ⦃h, L⦄ ⊢ T •*[g] U → - ∀T0. ⦃h, L⦄ ⊢ T •[g] ⦃0, T0⦄ → U = T. +lemma sstas_inv_O: ∀h,g,L,T,U. ⦃G, L⦄ ⊢ T •*[h, g] U → + ∀T0. ⦃G, L⦄ ⊢ T •[h, g] ⦃0, T0⦄ → U = T. #h #g #L #T #U #H @(sstas_ind_dx … H) -T // #T0 #U0 #l0 #HTU0 #_ #_ #T1 #HT01 elim (ssta_mono … HTU0 … HT01)