From: Ferruccio Guidi Date: Tue, 4 Feb 2014 20:16:21 +0000 (+0000) Subject: evaluation for context-sensitive extended substitution (cpye) completed! X-Git-Tag: make_still_working~980 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=e2527c6784c2593ca67af35fafaf0b3725d80a60;p=helm.git evaluation for context-sensitive extended substitution (cpye) completed! --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubstevalalt_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubstevalalt_6.ma index dc32bedc3..18cdf608f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubstevalalt_6.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubstevalalt_6.ma @@ -16,4 +16,4 @@ notation "hvbox( ⦃ term 46 G, break term 46 L ⦄ ⊢ break term 46 T1 ▶ ▶ * break [ term 46 d , break term 46 e ] break 𝐍 ⦃ term 46 T2 ⦄ )" non associative with precedence 45 - for @{ 'PRedEval $G $L $T1 $T2 $d $e }. + for @{ 'PSubstEvalAlt $G $L $T1 $T2 $d $e }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/cny_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cny_lift.ma index b1acaafaf..2d9b36856 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/cny_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/cny_lift.ma @@ -85,3 +85,25 @@ lemma cny_subst: ∀I,G,L,K,V,W,i,d,e. d ≤ yinj i → i < d + e → ⇩[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ ▶[O, ⫰(d+e-i)] 𝐍⦃V⦄ → ⇧[O, i+1] V ≡ W → ⦃G, L⦄ ⊢ ▶[d, e] 𝐍⦃W⦄. /3 width=13 by cny_subst_aux, ldrop_fwd_drop2/ qed-. + +(* Advanced inversion lemmas ************************************************) + +fact cny_inv_subst_aux: ∀G,L,K,V,W,i,d,e. d ≤ yinj i → i < d + e → + ⇩[i+1] L ≡ K → ⦃G, L⦄ ⊢ ▶[d, e] 𝐍⦃W⦄ → + ⇧[O, i+1] V ≡ W → ⦃G, K⦄ ⊢ ▶[O, ⫰(d+e-i)] 𝐍⦃V⦄. +#G #L #K #V #W #i #d #e #Hdi #Hide #HLK #HW #HVW +lapply (cny_narrow … HW (i+1) (⫰(d+e-i)) ? ?) -HW +[ >yplus_SO2 ylt_inv_O1 + [ >ymax_pre_sn_comm /2 width=2 by ylt_fwd_le/ + | lapply (monotonic_ylt_minus_dx … Hide i ?) // + ] +| /2 width=3 by yle_trans/ +| #HW lapply (cny_lift_inv_ge … HW … HLK … HVW ?) // -L -W + >yplus_inj >yminus_refl // +] +qed-. + +lemma cny_inv_subst: ∀I,G,L,K,V,W,i,d,e. d ≤ yinj i → i < d + e → + ⇩[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ▶[d, e] 𝐍⦃W⦄ → + ⇧[O, i+1] V ≡ W → ⦃G, K⦄ ⊢ ▶[O, ⫰(d+e-i)] 𝐍⦃V⦄. +/3 width=13 by cny_inv_subst_aux, ldrop_fwd_drop2/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye.ma index c93e64972..82a1e3b46 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye.ma @@ -57,4 +57,26 @@ qed. (* Basic inversion lemmas ***************************************************) -lemma \ No newline at end of file +lemma cpye_inv_sort1: ∀G,L,X,d,e,k. ⦃G, L⦄ ⊢ ⋆k ▶*[d, e] 𝐍⦃X⦄ → X = ⋆k. +#G #L #X #d #e #k * /2 width=5 by cpys_inv_sort1/ +qed-. + +lemma cpye_inv_gref1: ∀G,L,X,d,e,p. ⦃G, L⦄ ⊢ §p ▶*[d, e] 𝐍⦃X⦄ → X = §p. +#G #L #X #d #e #p * /2 width=5 by cpys_inv_gref1/ +qed-. + +lemma cpye_inv_bind1: ∀a,I,G,L,V1,T1,X,d,e. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[d, e] 𝐍⦃X⦄ → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[d, e] 𝐍⦃V2⦄ & ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] 𝐍⦃T2⦄ & + X = ⓑ{a,I}V2.T2. +#a #I #G #L #V1 #T1 #X #d #e * #H1 #H2 elim (cpys_inv_bind1 … H1) -H1 +#V2 #T2 #HV12 #HT12 #H destruct elim (cny_inv_bind … H2) -H2 +/5 width=8 by lsuby_cny_conf, lsuby_succ, ex3_2_intro, conj/ +qed-. + +lemma cpye_inv_flat1: ∀I,G,L,V1,T1,X,d,e. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[d, e] 𝐍⦃X⦄ → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[d, e] 𝐍⦃V2⦄ & ⦃G, L⦄ ⊢ T1 ▶*[d, e] 𝐍⦃T2⦄ & + X = ⓕ{I}V2.T2. +#I #G #L #V1 #T1 #X #d #e * #H1 #H2 elim (cpys_inv_flat1 … H1) -H1 +#V2 #T2 #HV12 #HT12 #H destruct elim (cny_inv_flat … H2) -H2 +/3 width=5 by ex3_2_intro, conj/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye_alt.ma new file mode 100644 index 000000000..a9dc53cf5 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye_alt.ma @@ -0,0 +1,89 @@ +(**************************************************************************) +(* ___ *) +(* ||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/psubstevalalt_6.ma". +include "basic_2/substitution/cpye_lift.ma". + +(* EVALUATION FOR CONTEXT-SENSITIVE EXTENDED SUBSTITUTION ON TERMS **********) + +(* Note: alternative definition of cpye *) +inductive cpyea: ynat → ynat → relation4 genv lenv term term ≝ +| cpyea_sort : ∀G,L,d,e,k. cpyea d e G L (⋆k) (⋆k) +| cpyea_free : ∀G,L,d,e,i. |L| ≤ i → cpyea d e G L (#i) (#i) +| cpyea_top : ∀G,L,d,e,i. d + e ≤ yinj i → cpyea d e G L (#i) (#i) +| cpyea_skip : ∀G,L,d,e,i. yinj i < d → cpyea d e G L (#i) (#i) +| cpyea_subst: ∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → yinj i < d+e → + ⇩[i] L ≡ K.ⓑ{I}V1 → cpyea (yinj 0) (⫰(d+e-yinj i)) G K V1 V2 → + ⇧[0, i+1] V2 ≡ W2 → cpyea d e G L (#i) W2 +| cpyea_gref : ∀G,L,d,e,p. cpyea d e G L (§p) (§p) +| cpyea_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e. + cpyea d e G L V1 V2 → cpyea (⫯d) e G (L.ⓑ{I}V1) T1 T2 → + cpyea d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) +| cpyea_flat : ∀I,G,L,V1,V2,T1,T2,d,e. + cpyea d e G L V1 V2 → cpyea d e G L T1 T2 → + cpyea d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) +. + +interpretation + "evaluation for context-sensitive extended substitution (term) alternative" + 'PSubstEvalAlt G L T1 T2 d e = (cpyea d e G L T1 T2). + +(* Main properties **********************************************************) + +theorem cpye_cpyea: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] 𝐍⦃T2⦄ → ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] 𝐍⦃T2⦄. +#G #L #T1 @(fqup_wf_ind_eq … G L T1) -G -L -T1 +#Z #Y #X #IH #G #L * * +[ #k #_ #_ #_ #T2 #d #e #H -X -Y -Z >(cpye_inv_sort1 … H) -H // +| #i #HG #HL #HT #T2 #d #e #H destruct + elim (cpye_inv_lref1 … H) -H * + /4 width=7 by cpyea_subst, cpyea_free, cpyea_top, cpyea_skip, fqup_lref/ +| #p #_ #_ #_ #T2 #d #e #H -X -Y -Z >(cpye_inv_gref1 … H) -H // +| #a #I #V1 #T1 #HG #HL #HT #T #d #e #H destruct + elim (cpye_inv_bind1 … H) -H /3 width=1 by cpyea_bind/ +| #I #V1 #T1 #HG #HL #HT #T #d #e #H destruct + elim (cpye_inv_flat1 … H) -H /3 width=1 by cpyea_flat/ +] +qed. + +(* Main inversion properties ************************************************) + +theorem cpyea_inv_cpye: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] 𝐍⦃T2⦄ → ⦃G, L⦄ ⊢ T1 ▶*[d, e] 𝐍⦃T2⦄. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e +/2 width=7 by cpye_subst, cpye_flat, cpye_bind, cpye_skip, cpye_top, cpye_free/ +qed-. + +(* Advanced eliminators *****************************************************) + +lemma cpye_ind_alt: ∀R:ynat→ynat→relation4 genv lenv term term. + (∀G,L,d,e,k. R d e G L (⋆k) (⋆k)) → + (∀G,L,d,e,i. |L| ≤ i → R d e G L (#i) (#i)) → + (∀G,L,d,e,i. d + e ≤ yinj i → R d e G L (#i) (#i)) → + (∀G,L,d,e,i. yinj i < d → R d e G L (#i) (#i)) → + (∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → yinj i < d + e → + ⇩[i] L ≡ K.ⓑ{I}V1 → ⦃G, K⦄ ⊢ V1 ▶*[yinj O, ⫰(d+e-yinj i)] 𝐍⦃V2⦄ → + ⇧[O, i+1] V2 ≡ W2 → R (yinj O) (⫰(d+e-yinj i)) G K V1 V2 → R d e G L (#i) W2 + ) → + (∀G,L,d,e,p. R d e G L (§p) (§p)) → + (∀a,I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] 𝐍⦃V2⦄ → + ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] 𝐍⦃T2⦄ → R d e G L V1 V2 → + R (⫯d) e G (L.ⓑ{I}V1) T1 T2 → R d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) + ) → + (∀I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] 𝐍⦃V2⦄ → + ⦃G, L⦄ ⊢ T1 ▶*[d, e] 𝐍⦃T2⦄ → R d e G L V1 V2 → + R d e G L T1 T2 → R d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) + ) → + ∀d,e,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] 𝐍⦃T2⦄ → R d e G L T1 T2. +#R #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #d #e #G #L #T1 #T2 #H elim (cpye_cpyea … H) -G -L -T1 -T2 -d -e +/3 width=8 by cpyea_inv_cpye/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye_lift.ma index a14708dfc..1d494871c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpye_lift.ma @@ -19,6 +19,8 @@ include "basic_2/substitution/cpye.ma". (* EVALUATION FOR CONTEXT-SENSITIVE EXTENDED SUBSTITUTION ON TERMS **********) +(* Advanced properties ******************************************************) + lemma cpye_subst: ∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d + e → ⇩[i] L ≡ K.ⓑ{I}V1 → ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(d+e-i)] 𝐍⦃V2⦄ → ⇧[O, i+1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ▶*[d, e] 𝐍⦃W2⦄. @@ -46,3 +48,46 @@ lemma cpys_total: ∀G,L,T1,d,e. ∃T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] 𝐍⦃T2 /3 width=2 by cpye_flat, ex_intro/ ] qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma cpye_inv_lref1: ∀G,L,T2,d,e,i. ⦃G, L⦄ ⊢ #i ▶*[d, e] 𝐍⦃T2⦄ → + ∨∨ |L| ≤ i ∧ T2 = #i + | d + e ≤ yinj i ∧ T2 = #i + | yinj i < d ∧ T2 = #i + | ∃∃I,K,V1,V2. d ≤ yinj i & yinj i < d + e & + ⇩[i] L ≡ K.ⓑ{I}V1 & + ⦃G, K⦄ ⊢ V1 ▶*[yinj 0, ⫰(d+e-yinj i)] 𝐍⦃V2⦄ & + ⇧[O, i+1] V2 ≡ T2. +#G #L #T2 #i #d #e * #H1 #H2 elim (cpys_inv_lref1 … H1) -H1 +[ #H destruct elim (cny_inv_lref … H2) -H2 + /3 width=1 by or4_intro0, or4_intro1, or4_intro2, conj/ +| * #I #K #V1 #V2 #Hdi #Hide #HLK #HV12 #HVT2 + @or4_intro3 @(ex5_4_intro … HLK … HVT2) (**) (* explicit constructor *) + /4 width=13 by cny_inv_subst_aux, ldrop_fwd_drop2, conj/ +] +qed-. + +lemma cpye_inv_lref1_free: ∀G,L,T2,d,e,i. ⦃G, L⦄ ⊢ #i ▶*[d, e] 𝐍⦃T2⦄ → + (∨∨ |L| ≤ i | d + e ≤ yinj i | yinj i < d) → T2 = #i. +#G #L #T2 #d #e #i #H * elim (cpye_inv_lref1 … H) -H * // +#I #K #V1 #V2 #Hdi #Hide #HLK #_ #_ #H +[ elim (lt_refl_false i) -d + @(lt_to_le_to_lt … H) -H /2 width=5 by ldrop_fwd_length_lt2/ (**) (* full auto slow: 19s *) +] +elim (ylt_yle_false … H) // +qed-. + +lemma cpye_inv_lref1_subst: ∀G,L,T2,d,e,i. ⦃G, L⦄ ⊢ #i ▶*[d, e] 𝐍⦃T2⦄ → + ∀I,K,V1,V2. d ≤ yinj i → yinj i < d + e → + ⇩[i] L ≡ K.ⓑ{I}V1 → ⇧[O, i+1] V2 ≡ T2 → + ⦃G, K⦄ ⊢ V1 ▶*[yinj 0, ⫰(d+e-yinj i)] 𝐍⦃V2⦄. +#G #L #T2 #d #e #i #H #I #K #V1 #V2 #Hdi #Hide #HLK #HVT2 elim (cpye_inv_lref1 … H) -H * +[ #H elim (lt_refl_false i) -V2 -T2 -d + @(lt_to_le_to_lt … H) -H /2 width=5 by ldrop_fwd_length_lt2/ +|2,3: #H elim (ylt_yle_false … H) // +| #Z #Y #X1 #X2 #_ #_ #HLY #HX12 #HXT2 + lapply (ldrop_mono … HLY … HLK) -HLY -HLK #H destruct + lapply (lift_inj … HXT2 … HVT2) -HXT2 -HVT2 #H destruct // +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma index 02f291022..6a5e2694c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma @@ -75,11 +75,13 @@ lemma cpys_cpysa: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L (* Basic inversion lemmas ***************************************************) -lemma cpysa_cpys: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. +lemma cpysa_inv_cpys: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. #G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e /2 width=7 by cpys_subst, cpys_flat, cpys_bind, cpy_cpys/ qed-. +(* Advanced eliminators *****************************************************) + lemma cpys_ind_alt: ∀R:ynat→ynat→relation4 genv lenv term term. (∀I,G,L,d,e. R d e G L (⓪{I}) (⓪{I})) → (∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d + e → @@ -96,5 +98,5 @@ lemma cpys_ind_alt: ∀R:ynat→ynat→relation4 genv lenv term term. ) → ∀d,e,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R d e G L T1 T2. #R #H1 #H2 #H3 #H4 #d #e #G #L #T1 #T2 #H elim (cpys_cpysa … H) -G -L -T1 -T2 -d -e -/3 width=8 by cpysa_cpys/ +/3 width=8 by cpysa_inv_cpys/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl index 43e2aa669..7dccdcc14 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl +++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl @@ -71,11 +71,11 @@ table { ] class "cyan" [ { "computation" * } { - [ { "context-sensitive extended evaluation" * } { + [ { "evaluation for context-sensitive extended reduction" * } { [ "cpxe ( ⦃?,?⦄ ⊢ ➡*[?,?] 𝐍⦃?⦄ )" * ] } ] - [ { "context-sensitive evaluation" * } { + [ { "evaluation for context-sensitive reduction" * } { [ "cpre ( ⦃?,?⦄ ⊢ ➡* 𝐍⦃?⦄ )" "cpre_cpre" * ] } ] @@ -212,6 +212,10 @@ table { [ "lleq ( ? ⋕[?,?] ? )" "lleq_alt ( ? ⋕⋕[?,?] ? )" "lleq_ldrop" + "lleq_fqus" + "lleq_lleq" + "lleq_ext" * ] } ] + [ { "evaluation for contxt-sensitive extended substitution" * } { + [ "cpye ( ⦃?,?⦄ ⊢ ? ▶*[?,?] 𝐍⦃?⦄ )" "cpye_alt ( ⦃?,?⦄ ⊢ ? ▶▶*[?,?] 𝐍⦃?⦄ )" "cpye_lift" * ] + } + ] [ { "contxt-sensitive extended multiple substitution" * } { [ "cpys ( ⦃?,?⦄ ⊢ ? ▶*[?,?] ? )" "cpys_alt ( ⦃?,?⦄ ⊢ ? ▶▶*[?,?] ? )" "cpys_lift" + "cpys_cpys" * ] }